From 878187d7f7e7e20b6628a971b457d5d2a95cf1c1 Mon Sep 17 00:00:00 2001 From: Steve Dunham Date: Sat, 5 Aug 2023 11:15:30 -0700 Subject: [PATCH 01/43] [ fix ] Totality checker misses indirect references to partial data --- CHANGELOG.md | 3 ++ src/Core/Termination/CallGraph.idr | 16 ++++++-- src/Idris/REPL/Opts.idr | 2 + tests/Main.idr | 2 +- tests/idris2/positivity004/expected | 54 ++++++++++++++++++++++++++- tests/idris2/total020/Issue-3030.idr | 33 ++++++++++++++++ tests/idris2/total020/Issue-3030b.idr | 18 +++++++++ tests/idris2/total020/Issue-524.idr | 14 +++++++ tests/idris2/total020/expected | 27 ++++++++++++++ tests/idris2/total020/run | 6 +++ 10 files changed, 168 insertions(+), 7 deletions(-) create mode 100644 tests/idris2/total020/Issue-3030.idr create mode 100644 tests/idris2/total020/Issue-3030b.idr create mode 100644 tests/idris2/total020/Issue-524.idr create mode 100644 tests/idris2/total020/expected create mode 100755 tests/idris2/total020/run diff --git a/CHANGELOG.md b/CHANGELOG.md index fc70675b43..dd75e0d9e1 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -101,6 +101,9 @@ * Fixed a bug that caused `f` to sometimes be replaced by `fx` after matching `fx = f x`. +* Fixed a bug in the totality checker that missed indirect references to + partial data. + ### Library changes #### Prelude diff --git a/src/Core/Termination/CallGraph.idr b/src/Core/Termination/CallGraph.idr index 083bf62d57..1f92096b8d 100644 --- a/src/Core/Termination/CallGraph.idr +++ b/src/Core/Termination/CallGraph.idr @@ -101,11 +101,19 @@ mutual do scs <- traverse (findSC defs env Unguarded pats) args pure (concat scs) (Guarded, Ref fc (DataCon _ _) cn, args) => - do scs <- traverse (findSC defs env Guarded pats) args - pure (concat scs) + do Just ty <- lookupTyExact cn (gamma defs) + | Nothing => do + log "totality" 50 $ "Lookup failed" + findSCcall defs env Guarded pats fc cn 0 args + arity <- getArity defs [] ty + findSCcall defs env Guarded pats fc cn arity args (Toplevel, Ref fc (DataCon _ _) cn, args) => - do scs <- traverse (findSC defs env Guarded pats) args - pure (concat scs) + do Just ty <- lookupTyExact cn (gamma defs) + | Nothing => do + log "totality" 50 $ "Lookup failed" + findSCcall defs env Guarded pats fc cn 0 args + arity <- getArity defs [] ty + findSCcall defs env Guarded pats fc cn arity args (_, Ref fc Func fn, args) => do logC "totality" 50 $ pure $ "Looking up type of " ++ show !(toFullNames fn) diff --git a/src/Idris/REPL/Opts.idr b/src/Idris/REPL/Opts.idr index dec692129f..e339f24595 100644 --- a/src/Idris/REPL/Opts.idr +++ b/src/Idris/REPL/Opts.idr @@ -59,6 +59,7 @@ record REPLOpts where litStyle : Maybe String -> Maybe LiterateStyle litStyle = join . map isLitFile +covering export defaultOpts : Maybe String -> OutputMode -> List (String, Codegen) -> REPLOpts defaultOpts fname outmode cgs @@ -108,6 +109,7 @@ setMainFile : {auto o : Ref ROpts REPLOpts} -> setMainFile src = update ROpts { mainfile := src, literateStyle := litStyle src } +covering export resetProofState : {auto o : Ref ROpts REPLOpts} -> Core () resetProofState = update ROpts { psResult := Nothing, diff --git a/tests/Main.idr b/tests/Main.idr index 0944e6e618..a6bf92dbd4 100644 --- a/tests/Main.idr +++ b/tests/Main.idr @@ -233,7 +233,7 @@ idrisTestsTotality = MkTestPool "Totality checking" [] Nothing "total001", "total002", "total003", "total004", "total005", "total006", "total007", "total008", "total009", "total010", "total011", "total012", "total013", "total014", "total015", - "total016", "total017", "total018", "total019" + "total016", "total017", "total018", "total019", "total020" ] -- This will only work with an Idris compiled via Chez or Racket, but at diff --git a/tests/idris2/positivity004/expected b/tests/idris2/positivity004/expected index 6448c09544..0f3eef7d24 100644 --- a/tests/idris2/positivity004/expected +++ b/tests/idris2/positivity004/expected @@ -14,6 +14,26 @@ Issue1771-1:4:3--4:8 4 | MkFix : f (Fix f) -> Fix f ^^^^^ +Error: argh is not total, possibly not terminating due to function Main.MkFix being reachable via Main.yesF -> Main.MkFix + +Issue1771-1:15:1--15:12 + 11 | + 12 | notF : Not F + 13 | notF (MkFix f) = f (yesF f) + 14 | + 15 | argh : Void + ^^^^^^^^^^^ + +Error: notF is not total, possibly not terminating due to call to Main.yesF + +Issue1771-1:12:1--12:13 + 08 | + 09 | yesF : Not F -> F + 10 | yesF nf = MkFix nf + 11 | + 12 | notF : Not F + ^^^^^^^^^^^^ + Error: yesF is not total, possibly not terminating due to call to Main.MkFix Issue1771-1:9:1--9:18 @@ -40,6 +60,26 @@ Issue1771-2:4:3--4:8 4 | MkFix : ((0 g : Type -> Type) -> g === Not -> g F) -> F ^^^^^ +Error: argh is not total, possibly not terminating due to function Main.MkFix being reachable via Main.yesF -> Main.MkFix + +Issue1771-2:14:1--14:12 + 10 | notF (MkFix f) = + 11 | let g = f Not Refl in + 12 | g (yesF g) + 13 | + 14 | argh : Void + ^^^^^^^^^^^ + +Error: notF is not total, possibly not terminating due to call to Main.yesF + +Issue1771-2:9:1--9:13 + 5 | + 6 | yesF : Not F -> F + 7 | yesF nf = MkFix (\ g, Refl => nf) + 8 | + 9 | notF : Not F + ^^^^^^^^^^^^ + Error: yesF is not total, possibly not terminating due to call to Main.MkFix Issue1771-2:6:1--6:18 @@ -67,7 +107,17 @@ Issue1771-3:10:3--10:6 10 | MkF : Wrap (Not F) -> F ^^^ -Error: notF is not total, possibly not terminating due to call to Main.F +Error: argh is not total, possibly not terminating due to function Main.MkF being reachable via Main.yesF -> Main.MkF + +Issue1771-3:18:1--18:12 + 14 | + 15 | notF : Not F + 16 | notF (MkF f) = unwrap f (yesF $ unwrap f) + 17 | + 18 | argh : Void + ^^^^^^^^^^^ + +Error: notF is not total, possibly not terminating due to calls to Main.F, Main.yesF Issue1771-3:15:1--15:13 11 | @@ -77,7 +127,7 @@ Issue1771-3:15:1--15:13 15 | notF : Not F ^^^^^^^^^^^^ -Error: yesF is not total, possibly not terminating due to calls to Main.F, Main.MkF +Error: yesF is not total, possibly not terminating due to call to Main.MkF Issue1771-3:12:1--12:18 08 | diff --git a/tests/idris2/total020/Issue-3030.idr b/tests/idris2/total020/Issue-3030.idr new file mode 100644 index 0000000000..46badcc9fd --- /dev/null +++ b/tests/idris2/total020/Issue-3030.idr @@ -0,0 +1,33 @@ + +-- %default total + +record Oops where + constructor MkOops + runOops : Not Oops + +total +notOops : Not Oops +notOops x = runOops x x + +-- total +oops : Oops +oops = MkOops notOops + +total +boom : Void +boom = runOops oops oops + +data Foo = MkFoo (Not Foo) + +runFoo : Foo -> Not Foo +runFoo (MkFoo nf) = nf + +notFoo : Not Foo +notFoo x = runFoo x x + +foo : Foo +foo = MkFoo notFoo + +total +boom2 : Void +boom2 = runFoo foo foo diff --git a/tests/idris2/total020/Issue-3030b.idr b/tests/idris2/total020/Issue-3030b.idr new file mode 100644 index 0000000000..18ba1f759d --- /dev/null +++ b/tests/idris2/total020/Issue-3030b.idr @@ -0,0 +1,18 @@ + +data Oops = MkOops (Not Oops) + +total +runOops : Oops -> Not Oops +runOops (MkOops nf) = nf + +total +notOops : Not Oops +notOops x = runOops x x + +covering +oops : Oops +oops = MkOops notOops + +total +boom : Void +boom = runOops oops oops diff --git a/tests/idris2/total020/Issue-524.idr b/tests/idris2/total020/Issue-524.idr new file mode 100644 index 0000000000..23fd25297e --- /dev/null +++ b/tests/idris2/total020/Issue-524.idr @@ -0,0 +1,14 @@ +data Bad = MkBad (Not Bad) + +hmmm : Bad -> Not Bad +hmmm (MkBad n) = n + +ok : Not Bad +ok bad = hmmm bad bad + +bad : Bad +bad = MkBad ok + +total +ohno : Void +ohno = ok bad diff --git a/tests/idris2/total020/expected b/tests/idris2/total020/expected new file mode 100644 index 0000000000..2c4c2ca7de --- /dev/null +++ b/tests/idris2/total020/expected @@ -0,0 +1,27 @@ +1/1: Building Issue-3030 (Issue-3030.idr) +Error: boom is not total, possibly not terminating due to function Main.MkOops being reachable via Main.oops -> Main.MkOops + +Issue-3030:16:1--17:12 + 16 | total + 17 | boom : Void + +Error: boom2 is not total, possibly not terminating due to function Main.MkFoo being reachable via Main.foo -> Main.MkFoo + +Issue-3030:31:1--32:13 + 31 | total + 32 | boom2 : Void + +1/1: Building Issue-3030b (Issue-3030b.idr) +Error: boom is not total, possibly not terminating due to function Main.MkOops being reachable via Main.oops -> Main.MkOops + +Issue-3030b:16:1--17:12 + 16 | total + 17 | boom : Void + +1/1: Building Issue-524 (Issue-524.idr) +Error: ohno is not total, possibly not terminating due to function Main.MkBad being reachable via Main.bad -> Main.MkBad + +Issue-524:12:1--13:12 + 12 | total + 13 | ohno : Void + diff --git a/tests/idris2/total020/run b/tests/idris2/total020/run new file mode 100755 index 0000000000..1f977a6226 --- /dev/null +++ b/tests/idris2/total020/run @@ -0,0 +1,6 @@ +rm -rf build + +$1 --no-color --console-width 0 --no-banner --check Issue-3030.idr +$1 --no-color --console-width 0 --no-banner --check Issue-3030b.idr +$1 --no-color --console-width 0 --no-banner --check Issue-524.idr + From 0b3e04aef069fdadc4905e0891a2e8dfbc8e3ae7 Mon Sep 17 00:00:00 2001 From: Steve Dunham Date: Sat, 12 Aug 2023 12:57:55 -0700 Subject: [PATCH 02/43] update tests for DCon in size change --- tests/idris2/casetree003/expected | 2 +- tests/idris2/coverage018/expected | 2 ++ tests/idris2/debug001/expected | 8 +++++--- tests/idris2/inlining001/expected | 4 ++++ tests/idris2/lazy003/expected | 2 +- 5 files changed, 13 insertions(+), 5 deletions(-) diff --git a/tests/idris2/casetree003/expected b/tests/idris2/casetree003/expected index c737896b60..7545583ccc 100644 --- a/tests/idris2/casetree003/expected +++ b/tests/idris2/casetree003/expected @@ -14,6 +14,6 @@ Compiled: \ {arg:1}, {arg:2} => case {arg:2} of Refers to: Main.view, Main.idL, Prelude.Basics.Nil, Prelude.Basics.(::) Refers to (runtime): Main.view, Main.idL, Main.Nil, Main.(::), Prelude.Basics.Nil, Prelude.Basics.(::) Flags: covering -Size change: Main.idL: [Just (0, Same), Just (2, Smaller), Nothing] +Size change: Prelude.Basics.Nil: [Just (0, Same)], Prelude.Basics.(::): [Just (0, Same), Just (2, Smaller), Nothing], Main.idL: [Just (0, Same), Just (2, Smaller), Nothing] Main> Bye for now! diff --git a/tests/idris2/coverage018/expected b/tests/idris2/coverage018/expected index 2aab32ecc8..1a5cb76d56 100644 --- a/tests/idris2/coverage018/expected +++ b/tests/idris2/coverage018/expected @@ -22,6 +22,7 @@ Compiled: \ {arg:0} => case {arg:0} of } Refers to: Prelude.Basics.True, Prelude.Basics.False Flags: allguarded, covering +Size change: Prelude.Basics.True: [], Prelude.Basics.False: [] Issue1831_1> Bye for now! 1/1: Building Issue1831_2 (Issue1831_2.idr) @@ -48,5 +49,6 @@ Compiled: \ {arg:0} => case {arg:0} of } Refers to: Prelude.Basics.True, Prelude.Basics.False Flags: allguarded, covering +Size change: Prelude.Basics.True: [], Prelude.Basics.False: [] Issue1831_2> Bye for now! diff --git a/tests/idris2/debug001/expected b/tests/idris2/debug001/expected index 1025d2aa91..865aed30c0 100644 --- a/tests/idris2/debug001/expected +++ b/tests/idris2/debug001/expected @@ -11,6 +11,7 @@ Compiled: \ {arg:0} => case {arg:0} of Refers to: Prelude.Basics.True, Prelude.Basics.False Refers to (runtime): Prelude.Types.Nat Flags: allguarded, covering +Size change: Prelude.Basics.True: [], Prelude.Basics.False: [] Main> Main.isInt32 Arguments [{arg:0}] Compile time tree: case {arg:0} of @@ -23,6 +24,7 @@ Compiled: \ {arg:0} => case {arg:0} of Refers to: Prelude.Basics.True, Prelude.Basics.False Refers to (runtime): Int32 Flags: allguarded, covering +Size change: Prelude.Basics.True: [], Prelude.Basics.False: [] Main> Prelude.Types.plus Arguments [{arg:0}, {arg:1}] Compile time tree: case {arg:0} of @@ -35,7 +37,7 @@ Compiled: \ {arg:0}, {arg:1} => case {arg:0} of } Refers to: Prelude.Types.plus, Prelude.Types.S Flags: total -Size change: Prelude.Types.plus: [Just (0, Smaller), Just (1, Same)] +Size change: Prelude.Types.S: [Nothing], Prelude.Types.plus: [Just (0, Smaller), Just (1, Same)] Main> Prelude.Types.minus Arguments [{arg:0}, {arg:1}] Compile time tree: case {arg:0} of @@ -63,7 +65,7 @@ Compiled: \ {arg:0}, {arg:1} => case {arg:0} of Refers to: Prelude.Types.minus, Prelude.Types.Z Refers to (runtime): Prelude.Types.prim__integerToNat Flags: total -Size change: Prelude.Types.minus: [Just (0, Smaller), Just (1, Smaller)] +Size change: Prelude.Types.Z: [], Prelude.Types.minus: [Just (0, Smaller), Just (1, Smaller)] Main> Prelude.Types.SnocList.filter Arguments [{arg:0}, {arg:1}, {arg:2}] Compile time tree: case {arg:2} of @@ -78,6 +80,6 @@ Compiled: \ {arg:1}, {arg:2} => case {arg:2} of Refers to: Prelude.Basics.SnocList, Prelude.Basics.Lin, Prelude.Types.SnocList.case block in filter, Prelude.Types.SnocList.filter Refers to (runtime): Prelude.Basics.Lin, Prelude.Basics.(:<), Prelude.Types.SnocList.filter Flags: total -Size change: Prelude.Types.SnocList.filter: [Just (0, Same), Just (1, Same), Just (2, Smaller)] +Size change: Prelude.Basics.Lin: [Just (0, Same)], Prelude.Types.SnocList.filter: [Just (0, Same), Just (1, Same), Just (2, Smaller)], Prelude.Basics.(:<): [Just (0, Same), Nothing, Just (2, Smaller)] Main> Bye for now! diff --git a/tests/idris2/inlining001/expected b/tests/idris2/inlining001/expected index 118c7f2b57..f38209e97f 100644 --- a/tests/idris2/inlining001/expected +++ b/tests/idris2/inlining001/expected @@ -19,22 +19,26 @@ Compile time tree: 10 Compiled: 10 Refers to: Prelude.Types.Z, Prelude.Types.S Flags: allguarded, covering, noinline +Size change: Prelude.Types.S: [Nothing], Prelude.Types.Z: [] Main> Main.heuristicPublicInline Arguments [] Compile time tree: 2 Compiled: 2 Refers to: Prelude.Types.Z, Prelude.Types.S Flags: inline, allguarded, covering +Size change: Prelude.Types.S: [Nothing], Prelude.Types.Z: [] Main> Main.exportedForced Arguments [] Compile time tree: 33 Compiled: 33 Refers to: Prelude.Types.Z, Prelude.Types.S Flags: allguarded, covering, inline +Size change: Prelude.Types.S: [Nothing], Prelude.Types.Z: [] Main> Main.exportedUnforced Arguments [] Compile time tree: 66 Compiled: 66 Refers to: Prelude.Types.Z, Prelude.Types.S Flags: allguarded, covering +Size change: Prelude.Types.S: [Nothing], Prelude.Types.Z: [] Main> Bye for now! diff --git a/tests/idris2/lazy003/expected b/tests/idris2/lazy003/expected index d6f7652144..bccbd8148a 100644 --- a/tests/idris2/lazy003/expected +++ b/tests/idris2/lazy003/expected @@ -22,5 +22,5 @@ Compiled: \ {arg:0} => let f = Force Lazy (Main.switch {arg:0}) in Refers to: Main.case block in switch3, Main.{_:985}, Main.switch, Builtin.Pair, Prelude.Types.Nat Refers to (runtime): Main.switch, Builtin.MkPair Flags: covering -Size change: Main.switch: [Just (0, Same)] +Size change: Main.switch: [Just (0, Same)], Builtin.MkPair: [Nothing, Nothing, Nothing, Nothing] Main> Bye for now! From 09466a96bb7289c7e7b36d3783b288c577262209 Mon Sep 17 00:00:00 2001 From: Robert Wright Date: Tue, 29 Aug 2023 11:47:42 +0100 Subject: [PATCH 03/43] Add testing utilities script --- tests/allbackends/basic048/run | 9 ++-- tests/allbackends/evaluator004/run | 7 ++-- tests/allbackends/evaluator005/run | 5 +-- tests/allbackends/issue2362/run | 5 +-- tests/allbackends/perf006/run | 5 +-- tests/allbackends/popen2/run | 4 +- tests/allschemes/channels001/run | 4 +- tests/allschemes/channels002/run | 4 +- tests/allschemes/channels003/run | 4 +- tests/allschemes/channels004/run | 4 +- tests/allschemes/channels005/run | 4 +- tests/allschemes/channels006/run | 4 +- tests/allschemes/scheme001/run | 4 +- tests/allschemes/scheme002/run | 8 ++-- tests/base/control_app001/run | 5 +-- tests/base/control_monad_instances/run | 9 ++-- tests/base/data_bits001/run | 5 +-- tests/base/data_bits002/run | 5 +-- tests/base/data_fin/run | 8 ++-- tests/base/data_integral/run | 4 +- tests/base/data_list001/run | 4 +- tests/base/data_list002/run | 4 +- tests/base/data_singleton/run | 4 +- tests/base/data_snoclist/run | 4 +- tests/base/data_string/run | 4 +- tests/base/data_string_lines001/run | 4 +- tests/base/data_string_parse_proof/run | 4 +- tests/base/data_string_unlines001/run | 4 +- tests/base/data_vect001/run | 4 +- tests/base/deriving_foldable/run | 4 +- tests/base/deriving_functor/run | 6 +-- tests/base/deriving_show/run | 4 +- tests/base/deriving_traversable/run | 4 +- tests/base/system_directory/run | 4 +- tests/base/system_env/run | 8 ++-- tests/base/system_errno/run | 5 +-- tests/base/system_file001/run | 13 +++--- tests/base/system_file_copyFile/run | 4 +- tests/base/system_file_fRead/run | 4 +- tests/base/system_file_popen/run | 4 +- tests/base/system_file_remove/run | 8 ++-- tests/base/system_info001/run | 5 +-- tests/base/system_info_os001/run | 5 +-- tests/base/system_run/run | 4 +- tests/base/system_signal001/run | 5 +-- tests/base/system_signal002/run | 5 +-- tests/base/system_signal003/run | 5 +-- tests/base/system_signal004/run | 5 +-- tests/base/system_system/run | 8 ++-- tests/base/system_time001/run | 13 +++--- tests/chez/barrier001/run | 4 +- tests/chez/bitops/run | 5 +-- tests/chez/buffer001/run | 6 +-- tests/chez/casts/run | 5 +-- tests/chez/chez001/run | 5 +-- tests/chez/chez002/run | 4 +- tests/chez/chez003/run | 4 +- tests/chez/chez004/run | 4 +- tests/chez/chez005/run | 5 +-- tests/chez/chez006/run | 7 ++-- tests/chez/chez007/run | 5 +-- tests/chez/chez008/run | 5 +-- tests/chez/chez009/run | 5 +-- tests/chez/chez010/run | 5 ++- tests/chez/chez011/run | 5 +-- tests/chez/chez012/run | 4 +- tests/chez/chez013/run | 5 ++- tests/chez/chez014/run | 5 +-- tests/chez/chez015/run | 5 +-- tests/chez/chez016/run | 5 ++- tests/chez/chez017/run | 6 ++- tests/chez/chez018/run | 4 +- tests/chez/chez019/run | 4 +- tests/chez/chez020/run | 5 +-- tests/chez/chez021/run | 5 +-- tests/chez/chez022/run | 6 ++- tests/chez/chez023/run | 4 +- tests/chez/chez024/run | 5 +-- tests/chez/chez025/run | 5 +-- tests/chez/chez026/run | 4 +- tests/chez/chez027/run | 5 +-- tests/chez/chez028/run | 5 +-- tests/chez/chez029/run | 5 +-- tests/chez/chez030/run | 5 +-- tests/chez/chez031/run | 9 ++-- tests/chez/chez032/run | 5 +-- tests/chez/chez033/run | 4 +- tests/chez/chez034/run | 5 +-- tests/chez/chez035/run | 8 ++-- tests/chez/chez036/run | 8 ++-- tests/chez/constfold/run | 7 ++-- tests/chez/constfold2/run | 7 ++-- tests/chez/constfold3/run | 7 ++-- tests/chez/forkjoin001/run | 4 +- tests/chez/futures001/run | 10 +++-- tests/chez/inlineiobind/run | 6 +-- tests/chez/integers/run | 5 +-- tests/chez/memo/run | 5 +-- tests/chez/nat2fin/run | 7 ++-- tests/chez/newints/run | 5 +-- tests/chez/perf001/run | 5 +-- tests/chez/reg001/run | 5 +-- tests/chez/semaphores001/run | 4 +- tests/chez/semaphores002/run | 4 +- tests/codegen/builtin001/run | 4 +- tests/codegen/con001/run | 4 +- tests/codegen/enum/run | 7 ++-- tests/contrib/json_001/run | 5 +-- tests/contrib/json_002/run | 4 +- tests/contrib/json_003/run | 4 +- tests/contrib/lexer/run | 4 +- tests/contrib/list_alternating/run | 4 +- tests/contrib/perf001/run | 4 +- tests/contrib/sortedmap_001/run | 4 +- .../contrib/system_directory_tree_copyDir/run | 4 +- tests/gambit/bitops001/run | 5 ++- tests/ideMode/ideMode001/run | 5 +-- tests/ideMode/ideMode002/run | 7 ++-- tests/ideMode/ideMode003/run | 5 +-- tests/ideMode/ideMode004/run | 5 +-- tests/ideMode/ideMode005/regenerate | 41 +++++++++---------- tests/ideMode/ideMode005/run | 40 +++++++++--------- tests/ideMode/ideMode006/run | 4 +- tests/idris2/api001/run | 7 ++-- tests/idris2/basic001/run | 5 +-- tests/idris2/basic002/run | 5 +-- tests/idris2/basic003/run | 7 ++-- tests/idris2/basic004/run | 5 +-- tests/idris2/basic005/run | 5 +-- tests/idris2/basic006/run | 5 +-- tests/idris2/basic007/run | 5 +-- tests/idris2/basic008/run | 5 +-- tests/idris2/basic009/run | 5 +-- tests/idris2/basic010/run | 5 +-- tests/idris2/basic011/run | 9 ++-- tests/idris2/basic012/run | 5 +-- tests/idris2/basic013/run | 5 +-- tests/idris2/basic014/run | 5 +-- tests/idris2/basic015/run | 5 +-- tests/idris2/basic016/run | 7 ++-- tests/idris2/basic017/run | 5 +-- tests/idris2/basic018/run | 5 +-- tests/idris2/basic019/run | 5 +-- tests/idris2/basic020/run | 5 +-- tests/idris2/basic021/run | 5 +-- tests/idris2/basic022/run | 5 +-- tests/idris2/basic023/run | 5 +-- tests/idris2/basic024/run | 5 +-- tests/idris2/basic025/run | 5 +-- tests/idris2/basic026/run | 5 +-- tests/idris2/basic027/run | 5 +-- tests/idris2/basic028/run | 5 +-- tests/idris2/basic029/run | 5 +-- tests/idris2/basic030/run | 5 +-- tests/idris2/basic031/run | 5 +-- tests/idris2/basic032/run | 7 ++-- tests/idris2/basic033/run | 5 +-- tests/idris2/basic034/run | 5 +-- tests/idris2/basic035/run | 5 +-- tests/idris2/basic036/run | 5 +-- tests/idris2/basic037/run | 7 ++-- tests/idris2/basic038/run | 5 +-- tests/idris2/basic039/run | 5 +-- tests/idris2/basic040/run | 5 +-- tests/idris2/basic041/run | 5 +-- tests/idris2/basic042/run | 7 ++-- tests/idris2/basic043/run | 5 +-- tests/idris2/basic044/run | 7 ++-- tests/idris2/basic045/run | 5 +-- tests/idris2/basic046/run | 5 +-- tests/idris2/basic047/run | 5 +-- tests/idris2/basic049/run | 5 +-- tests/idris2/basic050/run | 5 +-- tests/idris2/basic051/run | 5 +-- tests/idris2/basic052/run | 5 +-- tests/idris2/basic053/run | 5 +-- tests/idris2/basic054/run | 5 +-- tests/idris2/basic055/run | 5 +-- tests/idris2/basic056/run | 5 +-- tests/idris2/basic057/run | 5 +-- tests/idris2/basic058/run | 7 ++-- tests/idris2/basic059/run | 5 +-- tests/idris2/basic060/run | 5 +-- tests/idris2/basic061/run | 4 +- tests/idris2/basic062/run | 4 +- tests/idris2/basic063/run | 4 +- tests/idris2/basic064/run | 4 +- tests/idris2/basic065/run | 4 +- tests/idris2/basic066/run | 4 +- tests/idris2/basic067/run | 8 ++-- tests/idris2/basic068/run | 4 +- tests/idris2/basic069/run | 4 +- tests/idris2/basic070/run | 10 ++--- tests/idris2/builtin001/run | 5 +-- tests/idris2/builtin002/run | 5 +-- tests/idris2/builtin003/run | 5 +-- tests/idris2/builtin004/run | 5 +-- tests/idris2/builtin005/run | 5 +-- tests/idris2/builtin006/run | 5 +-- tests/idris2/builtin007/run | 5 +-- tests/idris2/builtin008/run | 5 +-- tests/idris2/builtin009/run | 5 +-- tests/idris2/builtin010/run | 5 +-- tests/idris2/builtin011/run | 5 +-- tests/idris2/builtin012/run | 5 +-- tests/idris2/casetree001/run | 7 ++-- tests/idris2/casetree002/run | 7 ++-- tests/idris2/casetree003/run | 4 +- tests/idris2/casetree004/run | 6 +-- tests/idris2/coverage001/run | 9 ++-- tests/idris2/coverage002/run | 5 +-- tests/idris2/coverage003/run | 5 +-- tests/idris2/coverage004/run | 5 +-- tests/idris2/coverage005/run | 5 +-- tests/idris2/coverage006/run | 5 +-- tests/idris2/coverage007/run | 5 +-- tests/idris2/coverage008/run | 5 +-- tests/idris2/coverage009/run | 5 +-- tests/idris2/coverage010/run | 5 +-- tests/idris2/coverage011/run | 5 +-- tests/idris2/coverage012/run | 7 ++-- tests/idris2/coverage013/run | 7 ++-- tests/idris2/coverage014/run | 5 +-- tests/idris2/coverage015/run | 7 ++-- tests/idris2/coverage016/run | 7 ++-- tests/idris2/coverage017/run | 5 +-- tests/idris2/coverage018/run | 6 +-- tests/idris2/coverage019/run | 4 +- tests/idris2/data001/run | 5 +-- tests/idris2/data002/run | 4 +- tests/idris2/debug001/run | 4 +- tests/idris2/docs001/run | 5 +-- tests/idris2/docs002/run | 5 +-- tests/idris2/docs003/run | 5 +-- tests/idris2/docs004/run | 6 +-- tests/idris2/docs005/run | 4 +- tests/idris2/dotted001/run | 4 +- tests/idris2/error001/run | 5 +-- tests/idris2/error002/run | 5 +-- tests/idris2/error003/run | 5 +-- tests/idris2/error004/run | 7 ++-- tests/idris2/error005/run | 5 +-- tests/idris2/error006/run | 5 +-- tests/idris2/error007/run | 5 +-- tests/idris2/error008/run | 5 +-- tests/idris2/error009/run | 5 +-- tests/idris2/error010/run | 5 +-- tests/idris2/error011/run | 5 +-- tests/idris2/error012/run | 4 +- tests/idris2/error013/run | 5 +-- tests/idris2/error014/run | 5 +-- tests/idris2/error015/run | 5 +-- tests/idris2/error016/run | 7 ++-- tests/idris2/error017/run | 7 ++-- tests/idris2/error018/run | 11 +++-- tests/idris2/error019/run | 5 +-- tests/idris2/error020/run | 5 +-- tests/idris2/error021/run | 5 +-- tests/idris2/error022/run | 4 +- tests/idris2/error023/run | 6 +-- tests/idris2/error024/run | 4 +- tests/idris2/error025/run | 4 +- tests/idris2/error026/run | 4 +- tests/idris2/error027/run | 4 +- tests/idris2/eta001/run | 5 +-- tests/idris2/evaluator001/run | 5 +-- tests/idris2/evaluator002/run | 5 +-- tests/idris2/evaluator003/run | 5 +-- tests/idris2/evaluator004/run | 4 +- tests/idris2/failing001/run | 4 +- tests/idris2/failing002/run | 12 +++--- tests/idris2/failing003/run | 4 +- tests/idris2/failing004/run | 4 +- tests/idris2/golden001/000-hello/run | 0 tests/idris2/golden001/run | 8 ++-- tests/idris2/idiom001/run | 5 +-- tests/idris2/import001/run | 9 ++-- tests/idris2/import002/run | 4 +- tests/idris2/import003/run | 7 ++-- tests/idris2/import004/run | 7 ++-- tests/idris2/import005/run | 5 +-- tests/idris2/import006/run | 4 +- tests/idris2/import007/run | 4 +- tests/idris2/import008/run | 8 ++-- tests/idris2/import009/run | 8 ++-- tests/idris2/inlining001/run | 5 +-- tests/idris2/interactive001/run | 5 +-- tests/idris2/interactive002/run | 5 +-- tests/idris2/interactive003/run | 7 ++-- tests/idris2/interactive004/run | 5 +-- tests/idris2/interactive005/run | 5 +-- tests/idris2/interactive006/run | 5 +-- tests/idris2/interactive007/run | 5 +-- tests/idris2/interactive008/run | 5 +-- tests/idris2/interactive009/run | 5 +-- tests/idris2/interactive010/run | 5 +-- tests/idris2/interactive011/run | 5 +-- tests/idris2/interactive012/run | 5 +-- tests/idris2/interactive013/run | 5 +-- tests/idris2/interactive014/run | 5 +-- tests/idris2/interactive015/run | 5 +-- tests/idris2/interactive016/run | 5 +-- tests/idris2/interactive017/run | 5 +-- tests/idris2/interactive018/run | 5 +-- tests/idris2/interactive019/run | 5 +-- tests/idris2/interactive020/run | 5 +-- tests/idris2/interactive021/run | 5 +-- tests/idris2/interactive022/run | 5 +-- tests/idris2/interactive023/run | 5 +-- tests/idris2/interactive024/run | 5 +-- tests/idris2/interactive025/run | 5 +-- tests/idris2/interactive026/run | 5 +-- tests/idris2/interactive027/run | 5 +-- tests/idris2/interactive028/run | 5 +-- tests/idris2/interactive029/run | 5 +-- tests/idris2/interactive030/run | 5 +-- tests/idris2/interactive031/run | 5 +-- tests/idris2/interactive032/run | 5 +-- tests/idris2/interactive033/run | 5 +-- tests/idris2/interactive034/run | 5 +-- tests/idris2/interactive035/run | 5 +-- tests/idris2/interactive036/run | 5 +-- tests/idris2/interactive037/run | 4 +- tests/idris2/interactive038/run | 5 +-- tests/idris2/interactive039/run | 4 +- tests/idris2/interactive040/run | 4 +- tests/idris2/interactive041/run | 4 +- tests/idris2/interactive042/run | 10 ++--- tests/idris2/interactive043/run | 6 +-- tests/idris2/interactive044/run | 4 +- tests/idris2/interactive045/run | 4 +- tests/idris2/interactive046/run | 5 +-- tests/idris2/interface001/run | 7 ++-- tests/idris2/interface002/run | 5 +-- tests/idris2/interface003/run | 5 +-- tests/idris2/interface004/run | 5 +-- tests/idris2/interface005/run | 5 +-- tests/idris2/interface006/run | 5 +-- tests/idris2/interface007/run | 5 +-- tests/idris2/interface008/run | 5 +-- tests/idris2/interface009/run | 5 +-- tests/idris2/interface010/run | 5 +-- tests/idris2/interface011/run | 5 +-- tests/idris2/interface012/run | 5 +-- tests/idris2/interface013/run | 5 +-- tests/idris2/interface014/run | 5 +-- tests/idris2/interface015/run | 5 +-- tests/idris2/interface016/run | 5 +-- tests/idris2/interface017/run | 5 +-- tests/idris2/interface018/run | 9 ++-- tests/idris2/interface019/run | 5 +-- tests/idris2/interface020/run | 5 +-- tests/idris2/interface021/run | 5 +-- tests/idris2/interface022/run | 5 +-- tests/idris2/interface023/run | 5 +-- tests/idris2/interface024/run | 5 +-- tests/idris2/interface025/run | 5 +-- tests/idris2/interface026/run | 5 +-- tests/idris2/interface027/run | 5 +-- tests/idris2/interface028/run | 4 +- tests/idris2/interface029/run | 8 ++-- tests/idris2/interpolation001/run | 5 +-- tests/idris2/interpolation002/run | 5 +-- tests/idris2/interpolation003/run | 5 +-- tests/idris2/interpolation004/run | 4 +- tests/idris2/interpreter001/run | 4 +- tests/idris2/interpreter002/run | 5 +-- tests/idris2/interpreter003/run | 4 +- tests/idris2/interpreter004/run | 4 +- tests/idris2/interpreter005/run | 5 +-- tests/idris2/interpreter006/run | 4 +- tests/idris2/interpreter007/run | 4 +- tests/idris2/interpreter008/run | 4 +- tests/idris2/lazy001/run | 5 +-- tests/idris2/lazy002/run | 5 +-- tests/idris2/lazy003/run | 5 +-- tests/idris2/linear001/run | 5 +-- tests/idris2/linear002/run | 5 +-- tests/idris2/linear003/run | 5 +-- tests/idris2/linear004/run | 5 +-- tests/idris2/linear005/run | 5 +-- tests/idris2/linear006/run | 5 +-- tests/idris2/linear007/run | 5 +-- tests/idris2/linear008/run | 5 +-- tests/idris2/linear009/run | 5 +-- tests/idris2/linear010/run | 4 +- tests/idris2/linear011/run | 4 +- tests/idris2/linear012/run | 5 +-- tests/idris2/linear013/run | 5 +-- tests/idris2/linear014/run | 4 +- tests/idris2/linear015/run | 4 +- tests/idris2/linear016/run | 6 +-- tests/idris2/linear017/run | 4 +- tests/idris2/literate001/run | 5 +-- tests/idris2/literate002/run | 7 ++-- tests/idris2/literate003/run | 5 +-- tests/idris2/literate004/run | 5 +-- tests/idris2/literate005/run | 5 +-- tests/idris2/literate006/run | 5 +-- tests/idris2/literate007/run | 7 ++-- tests/idris2/literate008/run | 5 +-- tests/idris2/literate009/run | 5 +-- tests/idris2/literate010/run | 5 +-- tests/idris2/literate011/run | 5 +-- tests/idris2/literate012/run | 5 +-- tests/idris2/literate013/run | 6 +-- tests/idris2/literate014/run | 5 +-- tests/idris2/literate015/run | 5 +-- tests/idris2/literate016/run | 5 +-- tests/idris2/literate017/run | 8 ++-- tests/idris2/namespace001/run | 7 ++-- tests/idris2/namespace002/run | 4 +- tests/idris2/namespace003/run | 4 +- tests/idris2/namespace004/run | 4 +- tests/idris2/namespace005/run | 28 ++++++------- tests/idris2/params001/run | 7 ++-- tests/idris2/params002/run | 5 +-- tests/idris2/params003/run | 5 +-- tests/idris2/perf001/run | 5 +-- tests/idris2/perf002/run | 5 +-- tests/idris2/perf003/run | 5 +-- tests/idris2/perf004/run | 5 +-- tests/idris2/perf005/run | 11 +++-- tests/idris2/perf007/run | 5 +-- tests/idris2/perf008/run | 5 +-- tests/idris2/perf009/run | 4 +- tests/idris2/perf010/run | 4 +- tests/idris2/perf011/run | 4 +- tests/idris2/perf012/run | 4 +- tests/idris2/perf2202/run | 6 +-- tests/idris2/perror001/run | 5 +-- tests/idris2/perror002/run | 5 +-- tests/idris2/perror003/run | 7 ++-- tests/idris2/perror004/run | 5 +-- tests/idris2/perror005/run | 5 +-- tests/idris2/perror006/run | 5 +-- tests/idris2/perror007/run | 27 ++++++------ tests/idris2/perror008/run | 19 ++++----- tests/idris2/perror009/run | 4 +- tests/idris2/perror010/run | 12 +++--- tests/idris2/perror011/run | 10 ++--- tests/idris2/perror012/run | 6 +-- tests/idris2/perror013/run | 10 ++--- tests/idris2/perror014/run | 4 +- tests/idris2/perror015/run | 4 +- tests/idris2/perror016/run | 10 ++--- tests/idris2/perror017/run | 4 +- tests/idris2/perror018/run | 8 ++-- tests/idris2/perror019/run | 4 +- tests/idris2/perror020/run | 6 +-- tests/idris2/perror021/run | 5 +-- tests/idris2/perror022/run | 5 +-- tests/idris2/perror023/run | 4 +- tests/idris2/perror024/run | 4 +- tests/idris2/perror025/run | 4 +- tests/idris2/perror026/run | 4 +- tests/idris2/perror027/run | 4 +- tests/idris2/perror028/run | 4 +- tests/idris2/perror029/run | 4 +- tests/idris2/pkg001/run | 5 +-- tests/idris2/pkg002/run | 5 +-- tests/idris2/pkg003/run | 16 ++++---- tests/idris2/pkg004/run | 5 +-- tests/idris2/pkg005/run | 9 ++-- tests/idris2/pkg006/run | 13 +++--- tests/idris2/pkg007/run | 10 +++-- tests/idris2/pkg008/run | 5 ++- tests/idris2/pkg009/run | 4 +- tests/idris2/pkg010/run | 4 +- tests/idris2/pkg011/run | 12 +++--- tests/idris2/pkg012/run | 4 +- tests/idris2/pkg013/run | 4 +- tests/idris2/pkg014/run | 4 +- tests/idris2/pkg015/run | 4 +- tests/idris2/pkg016/run | 11 ++--- tests/idris2/pkg017/run | 10 +++-- tests/idris2/positivity001/run | 5 +-- tests/idris2/positivity002/run | 5 +-- tests/idris2/positivity003/run | 5 +-- tests/idris2/positivity004/run | 8 ++-- tests/idris2/pretty001/run | 4 +- tests/idris2/pretty002/run | 4 +- tests/idris2/primloop/run | 4 +- tests/idris2/quantifiers001/run | 4 +- tests/idris2/real001/run | 7 ++-- tests/idris2/real002/run | 7 ++-- tests/idris2/record001/run | 5 +-- tests/idris2/record002/run | 5 +-- tests/idris2/record003/run | 5 +-- tests/idris2/record004/run | 5 +-- tests/idris2/record005/run | 5 +-- tests/idris2/record006/run | 5 +-- tests/idris2/record007/run | 5 +-- tests/idris2/record008/run | 5 +-- tests/idris2/record009/run | 5 +-- tests/idris2/record010/run | 5 +-- tests/idris2/record011/run | 4 +- tests/idris2/record012/run | 4 +- tests/idris2/record013/run | 4 +- tests/idris2/record014/run | 4 +- tests/idris2/record015/run | 4 +- tests/idris2/record016/run | 4 +- tests/idris2/record017/run | 4 +- tests/idris2/record018/run | 5 +-- tests/idris2/record019/run | 4 +- tests/idris2/reflection001/run | 5 +-- tests/idris2/reflection002/run | 5 +-- tests/idris2/reflection003/run | 5 +-- tests/idris2/reflection004/run | 5 +-- tests/idris2/reflection005/run | 5 +-- tests/idris2/reflection006/run | 5 +-- tests/idris2/reflection007/run | 5 +-- tests/idris2/reflection008/run | 5 +-- tests/idris2/reflection009/run | 5 +-- tests/idris2/reflection010/run | 4 +- tests/idris2/reflection011/run | 5 +-- tests/idris2/reflection012/run | 5 +-- tests/idris2/reflection013/run | 5 +-- tests/idris2/reflection014/run | 5 +-- tests/idris2/reflection015/run | 5 +-- tests/idris2/reflection016/run | 7 ++-- tests/idris2/reflection017/run | 6 +-- tests/idris2/reflection018/run | 4 +- tests/idris2/reflection019/run | 4 +- tests/idris2/reflection020/run | 8 ++-- tests/idris2/reflection021/run | 4 +- tests/idris2/reg001/run | 5 +-- tests/idris2/reg002/run | 5 +-- tests/idris2/reg003/run | 4 +- tests/idris2/reg004/run | 5 +-- tests/idris2/reg005/run | 5 +-- tests/idris2/reg006/run | 5 +-- tests/idris2/reg007/run | 5 +-- tests/idris2/reg008/run | 5 +-- tests/idris2/reg009/run | 5 +-- tests/idris2/reg010/run | 5 +-- tests/idris2/reg011/run | 5 +-- tests/idris2/reg012/run | 5 +-- tests/idris2/reg013/run | 5 +-- tests/idris2/reg014/run | 5 +-- tests/idris2/reg015/run | 5 +-- tests/idris2/reg016/run | 5 +-- tests/idris2/reg017/run | 5 +-- tests/idris2/reg018/run | 5 +-- tests/idris2/reg019/run | 5 +-- tests/idris2/reg020/run | 5 +-- tests/idris2/reg021/run | 5 +-- tests/idris2/reg022/run | 5 +-- tests/idris2/reg023/run | 5 +-- tests/idris2/reg024/run | 5 +-- tests/idris2/reg025/run | 5 +-- tests/idris2/reg026/run | 5 +-- tests/idris2/reg027/run | 5 +-- tests/idris2/reg028/run | 5 +-- tests/idris2/reg029/run | 5 +-- tests/idris2/reg030/run | 5 +-- tests/idris2/reg031/run | 5 +-- tests/idris2/reg032/run | 5 +-- tests/idris2/reg033/run | 5 +-- tests/idris2/reg034/run | 5 +-- tests/idris2/reg035/run | 5 +-- tests/idris2/reg036/run | 5 +-- tests/idris2/reg037/run | 5 +-- tests/idris2/reg038/run | 7 ++-- tests/idris2/reg039/run | 5 +-- tests/idris2/reg040/run | 5 +-- tests/idris2/reg041/run | 5 +-- tests/idris2/reg042/run | 5 +-- tests/idris2/reg043/run | 6 +-- tests/idris2/reg044/run | 5 +-- tests/idris2/reg045/run | 4 +- tests/idris2/reg046/run | 4 +- tests/idris2/reg047/run | 5 ++- tests/idris2/reg048/run | 4 +- tests/idris2/reg049/run | 4 +- tests/idris2/reg050/run | 4 +- tests/idris2/reg051/run | 4 +- tests/idris2/reg052/run | 4 +- tests/idris2/repl001/run | 4 +- tests/idris2/repl002/run | 4 +- tests/idris2/repl003/run | 4 +- tests/idris2/repl004/run | 4 +- tests/idris2/repl005/run | 4 +- tests/idris2/repl006/run | 4 +- tests/idris2/rewrite001/run | 4 +- tests/idris2/schemeeval001/run | 5 +-- tests/idris2/schemeeval002/run | 5 +-- tests/idris2/schemeeval003/run | 5 +-- tests/idris2/schemeeval004/run | 5 +-- tests/idris2/schemeeval005/run | 5 +-- tests/idris2/schemeeval006/run | 4 +- tests/idris2/spec001/run | 12 +++--- tests/idris2/termination001/run | 4 +- tests/idris2/total001/run | 5 +-- tests/idris2/total002/run | 5 +-- tests/idris2/total003/run | 5 +-- tests/idris2/total004/run | 5 +-- tests/idris2/total005/run | 5 +-- tests/idris2/total006/run | 5 +-- tests/idris2/total007/run | 5 +-- tests/idris2/total008/run | 5 +-- tests/idris2/total009/run | 5 +-- tests/idris2/total010/run | 5 +-- tests/idris2/total011/run | 12 +++--- tests/idris2/total012/run | 16 ++++---- tests/idris2/total013/run | 4 +- tests/idris2/total014/run | 4 +- tests/idris2/total015/run | 4 +- tests/idris2/total016/run | 6 +-- tests/idris2/total017/run | 4 +- tests/idris2/total018/run | 4 +- tests/idris2/total019/run | 4 +- tests/idris2/total020/run | 4 +- tests/idris2/unification001/run | 4 +- tests/idris2/warning001/run | 13 +++--- tests/idris2/warning002/run | 7 ++-- tests/idris2/warning003/run | 5 +-- tests/idris2/warning004/run | 6 +-- tests/idris2/with001/run | 5 +-- tests/idris2/with002/run | 5 +-- tests/idris2/with003/run | 5 +-- tests/idris2/with004/run | 9 ++-- tests/idris2/with005/run | 7 ++-- tests/idris2/with006/run | 4 +- tests/idris2/with007/run | 4 +- tests/idris2/with008/run | 4 +- tests/idris2/with009/run | 4 +- tests/idris2/with010/run | 4 +- tests/idris2/with011/run | 4 +- tests/node/args/run | 5 +-- tests/node/bitops/run | 5 +-- tests/node/casts/run | 5 +-- tests/node/doubles/run | 5 +-- tests/node/fastConcat/run | 5 +-- tests/node/fix1839/run | 6 +-- tests/node/idiom001/run | 5 +-- tests/node/integer_array/run | 4 +- tests/node/integers/run | 5 +-- tests/node/memo/run | 5 +-- tests/node/newints/run | 5 +-- tests/node/node001/run | 5 +-- tests/node/node002/run | 4 +- tests/node/node003/run | 4 +- tests/node/node004/run | 7 +++- tests/node/node005/run | 5 +-- tests/node/node006/run | 7 ++-- tests/node/node007/run | 5 +-- tests/node/node008/run | 5 +-- tests/node/node009/run | 5 +-- tests/node/node011/run | 5 +-- tests/node/node012/run | 4 +- tests/node/node014/run | 5 +-- tests/node/node015/run | 5 +-- tests/node/node017/run | 7 +++- tests/node/node018/run | 6 ++- tests/node/node019/run | 5 +-- tests/node/node020/run | 5 +-- tests/node/node021/run | 5 +-- tests/node/node022/run | 5 +-- tests/node/node023/run | 5 +-- tests/node/node024/run | 5 +-- tests/node/node025/run | 5 +-- tests/node/node026/run | 5 +-- tests/node/node027/run | 4 +- tests/node/nomangle001/run | 6 +-- tests/node/nomangle002/run | 6 +-- tests/node/perf001/run | 5 +-- tests/node/reg001/run | 5 +-- tests/node/reg002/run | 5 +-- tests/node/stringcast/run | 5 +-- tests/node/syntax001/run | 5 +-- tests/node/tailrec001/run | 5 +-- tests/node/tailrec002/run | 6 +-- tests/node/tailrec_libs/run | 4 +- tests/prelude/bind001/run | 4 +- tests/prelude/char001/run | 4 +- tests/prelude/double001/run | 4 +- tests/prelude/nat001/run | 4 +- tests/prelude/reg001/run | 5 +-- tests/racket/barrier001/run | 4 +- tests/racket/conditions001/run | 4 +- tests/racket/conditions002/run | 4 +- tests/racket/conditions003/run | 4 +- tests/racket/conditions004/run | 4 +- tests/racket/conditions005/run | 4 +- tests/racket/conditions006/run | 4 +- tests/racket/conditions007/run | 4 +- tests/racket/ffi001/run | 4 +- tests/racket/forkjoin001/run | 4 +- tests/racket/futures001/run | 6 +-- tests/racket/mutex001/run | 4 +- tests/racket/mutex002/run | 4 +- tests/racket/mutex003/run | 4 +- tests/racket/mutex004/run | 4 +- tests/racket/mutex005/run | 4 +- tests/racket/semaphores001/run | 4 +- tests/racket/semaphores002/run | 4 +- tests/refc/args/run | 5 +-- tests/refc/buffer/run | 4 +- tests/refc/clock/run | 5 +-- tests/refc/doubles/run | 5 +-- tests/refc/garbageCollect/run | 4 +- tests/refc/integers/run | 5 +-- tests/refc/issue1778/run | 5 +-- tests/refc/issue2424/run | 4 +- tests/refc/issue2452/run | 4 +- tests/refc/refc001/run | 5 +-- tests/refc/refc002/run | 5 +-- tests/refc/refc003/run | 4 +- tests/refc/reg001/run | 4 +- tests/refc/strings/run | 5 +-- tests/templates/simple-test/notes.md | 6 +-- tests/templates/simple-test/run | 5 +-- tests/templates/ttimp/notes.md | 5 ++- tests/templates/ttimp/run | 6 +-- tests/templates/with-ipkg/notes.md | 7 ++-- tests/templates/with-ipkg/run | 5 +-- tests/testutils.sh | 23 +++++++++++ tests/ttimp/basic001/run | 6 +-- tests/ttimp/basic002/run | 5 +-- tests/ttimp/basic003/run | 6 +-- tests/ttimp/basic004/run | 5 +-- tests/ttimp/basic005/run | 5 +-- tests/ttimp/basic006/run | 5 +-- tests/ttimp/coverage002/run | 5 +-- tests/ttimp/dot001/run | 11 +++-- tests/ttimp/eta001/run | 5 +-- tests/ttimp/lazy001/run | 8 ++-- tests/ttimp/nest001/run | 5 +-- tests/ttimp/nest002/run | 5 +-- tests/ttimp/perf001/run | 5 +-- tests/ttimp/perf002/run | 5 +-- tests/ttimp/perf003/run | 5 +-- tests/ttimp/qtt001/run | 5 +-- tests/ttimp/qtt003/run | 5 +-- tests/ttimp/record001/run | 6 +-- tests/ttimp/record002/run | 5 +-- tests/ttimp/record003/run | 6 +-- tests/ttimp/record004/run | 6 +-- tests/ttimp/total001/run | 5 +-- tests/ttimp/total002/run | 5 +-- tests/ttimp/total003/run | 5 +-- tests/typedd-book/chapter01/run | 7 ++-- tests/typedd-book/chapter02/run | 5 +-- tests/typedd-book/chapter03/run | 5 +-- tests/typedd-book/chapter04/run | 5 +-- tests/typedd-book/chapter05/run | 5 +-- tests/typedd-book/chapter06/run | 5 +-- tests/typedd-book/chapter07/run | 5 +-- tests/typedd-book/chapter08/run | 5 +-- tests/typedd-book/chapter09/run | 5 +-- tests/typedd-book/chapter10/run | 7 ++-- tests/typedd-book/chapter11/run | 5 +-- tests/typedd-book/chapter12/run | 5 +-- tests/typedd-book/chapter13/run | 5 +-- tests/typedd-book/chapter14/run | 5 +-- tests/vmcode/basic001/run | 5 +-- 757 files changed, 1878 insertions(+), 2214 deletions(-) mode change 100755 => 100644 tests/idris2/golden001/000-hello/run create mode 100755 tests/testutils.sh diff --git a/tests/allbackends/basic048/run b/tests/allbackends/basic048/run index 0073df211b..032c8cb727 100644 --- a/tests/allbackends/basic048/run +++ b/tests/allbackends/basic048/run @@ -1,6 +1,5 @@ -rm -rf build - -$1 --no-banner --no-color --console-width 0 "Module'.idr" < input -$1 --exec main "Module'.idr" -EDITOR=true $1 --no-banner --no-color --console-width 0 "Module'.idr" < input-ed +. ../../testutils.sh +idris2 "Module'.idr" < input +run "Module'.idr" +EDITOR=true idris2 "Module'.idr" < input-ed diff --git a/tests/allbackends/evaluator004/run b/tests/allbackends/evaluator004/run index 7aab6792b4..1ca13565a9 100644 --- a/tests/allbackends/evaluator004/run +++ b/tests/allbackends/evaluator004/run @@ -1,5 +1,4 @@ -rm -rf build - -$1 --no-banner --no-color --console-width 0 Issue735.idr < input -$1 --exec main Issue735.idr +. ../../testutils.sh +idris2 Issue735.idr < input +run Issue735.idr diff --git a/tests/allbackends/evaluator005/run b/tests/allbackends/evaluator005/run index 37e38ad0ec..a138c1c6b2 100644 --- a/tests/allbackends/evaluator005/run +++ b/tests/allbackends/evaluator005/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --exec main Issue1200.idr +. ../../testutils.sh +run Issue1200.idr diff --git a/tests/allbackends/issue2362/run b/tests/allbackends/issue2362/run index 35bef96f9b..574a2fe3a7 100644 --- a/tests/allbackends/issue2362/run +++ b/tests/allbackends/issue2362/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --exec main Issue2362.idr +. ../../testutils.sh +run Issue2362.idr diff --git a/tests/allbackends/perf006/run b/tests/allbackends/perf006/run index bcf4cbfc0a..1bc540cab8 100644 --- a/tests/allbackends/perf006/run +++ b/tests/allbackends/perf006/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --exec main Issue775.idr +. ../../testutils.sh +run Issue775.idr diff --git a/tests/allbackends/popen2/run b/tests/allbackends/popen2/run index 7fa6822316..ceb556ba04 100644 --- a/tests/allbackends/popen2/run +++ b/tests/allbackends/popen2/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --exec main Test.idr +run Test.idr diff --git a/tests/allschemes/channels001/run b/tests/allschemes/channels001/run index 82545ce087..c6e6f2ab79 100644 --- a/tests/allschemes/channels001/run +++ b/tests/allschemes/channels001/run @@ -1 +1,3 @@ -$1 --no-banner --no-color --console-width 0 Main.idr --exec main +. ../../testutils.sh + +run Main.idr diff --git a/tests/allschemes/channels002/run b/tests/allschemes/channels002/run index bd38ceecab..c6e6f2ab79 100644 --- a/tests/allschemes/channels002/run +++ b/tests/allschemes/channels002/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-banner --no-color --console-width 0 Main.idr --exec main +run Main.idr diff --git a/tests/allschemes/channels003/run b/tests/allschemes/channels003/run index bd38ceecab..c6e6f2ab79 100644 --- a/tests/allschemes/channels003/run +++ b/tests/allschemes/channels003/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-banner --no-color --console-width 0 Main.idr --exec main +run Main.idr diff --git a/tests/allschemes/channels004/run b/tests/allschemes/channels004/run index bd38ceecab..c6e6f2ab79 100644 --- a/tests/allschemes/channels004/run +++ b/tests/allschemes/channels004/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-banner --no-color --console-width 0 Main.idr --exec main +run Main.idr diff --git a/tests/allschemes/channels005/run b/tests/allschemes/channels005/run index bd38ceecab..c6e6f2ab79 100644 --- a/tests/allschemes/channels005/run +++ b/tests/allschemes/channels005/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-banner --no-color --console-width 0 Main.idr --exec main +run Main.idr diff --git a/tests/allschemes/channels006/run b/tests/allschemes/channels006/run index bd38ceecab..c6e6f2ab79 100644 --- a/tests/allschemes/channels006/run +++ b/tests/allschemes/channels006/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-banner --no-color --console-width 0 Main.idr --exec main +run Main.idr diff --git a/tests/allschemes/scheme001/run b/tests/allschemes/scheme001/run index 3415b5fad9..f000c530e4 100644 --- a/tests/allschemes/scheme001/run +++ b/tests/allschemes/scheme001/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --exec main CastStringDouble.idr +run CastStringDouble.idr diff --git a/tests/allschemes/scheme002/run b/tests/allschemes/scheme002/run index ef7704711d..59b8b3ab66 100755 --- a/tests/allschemes/scheme002/run +++ b/tests/allschemes/scheme002/run @@ -1,13 +1,13 @@ -rm -rf build +. ../../testutils.sh # observe that errors are correctly reported as zero. -$1 --exec main TermSize.idr expected -$1 --no-banner --no-color --console-width 0 NumProcessors.idr --exec main - +run NumProcessors.idr diff --git a/tests/base/system_info_os001/run b/tests/base/system_info_os001/run index a143be7878..15034e1601 100755 --- a/tests/base/system_info_os001/run +++ b/tests/base/system_info_os001/run @@ -1,4 +1,4 @@ -rm -rf build +. ../../testutils.sh case "$(uname)" in Linux*) EXPECTED_OS=unix ;; @@ -12,5 +12,4 @@ case "$(uname)" in esac export EXPECTED_OS -$1 --no-banner --no-color --console-width 0 Os.idr --exec main - +run Os.idr diff --git a/tests/base/system_run/run b/tests/base/system_run/run index b79096cc46..125d488562 100644 --- a/tests/base/system_run/run +++ b/tests/base/system_run/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --no-banner Run.idr < input +idris2 Run.idr < input diff --git a/tests/base/system_signal001/run b/tests/base/system_signal001/run index 98200cbbf6..8a317bcdcb 100755 --- a/tests/base/system_signal001/run +++ b/tests/base/system_signal001/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner IgnoreSignal.idr < input +. ../../testutils.sh +idris2 IgnoreSignal.idr < input diff --git a/tests/base/system_signal002/run b/tests/base/system_signal002/run index 1fffa402da..8ea51c0059 100755 --- a/tests/base/system_signal002/run +++ b/tests/base/system_signal002/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner HandleSignal.idr < input +. ../../testutils.sh +idris2 HandleSignal.idr < input diff --git a/tests/base/system_signal003/run b/tests/base/system_signal003/run index d13f912949..e527a9082e 100755 --- a/tests/base/system_signal003/run +++ b/tests/base/system_signal003/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner DefaultSignal.idr < input 2> /dev/null +. ../../testutils.sh +idris2 DefaultSignal.idr < input 2> /dev/null diff --git a/tests/base/system_signal004/run b/tests/base/system_signal004/run index 0c713bf2df..7f4d97c77d 100755 --- a/tests/base/system_signal004/run +++ b/tests/base/system_signal004/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner HandleManySignals.idr < input +. ../../testutils.sh +idris2 HandleManySignals.idr < input diff --git a/tests/base/system_system/run b/tests/base/system_system/run index a50e660a28..6a07c03b17 100755 --- a/tests/base/system_system/run +++ b/tests/base/system_system/run @@ -1,11 +1,11 @@ -rm -rf build +. ../../testutils.sh # @system@ uses primitive functions with definitions for both # C (supported by most backends) and Node. -$1 --cg chez -o test Test.idr +idris2 --cg chez -o test Test.idr ./build/exec/test -rm -rf ./build +. ../../testutils.sh -$1 --cg node -o test.js Test.idr +idris2 --cg node -o test.js Test.idr node ./build/exec/test.js diff --git a/tests/base/system_time001/run b/tests/base/system_time001/run index 5481657ddc..34d15f1121 100755 --- a/tests/base/system_time001/run +++ b/tests/base/system_time001/run @@ -1,16 +1,15 @@ -rm -rf build +. ../../testutils.sh # @time@ uses a primitive function with definitions for both # C (supported by most backends) and Javascript (Node & Browsers). -$1 --cg chez --no-color --console-width 0 --no-banner Time.idr < input -$1 --cg node --no-color --console-width 0 --no-banner Time.idr < input +idris2 --cg chez Time.idr < input +idris2 --cg node Time.idr < input # The following backends failed for reasons unrelated to this test. They can be # uncommented at a future date. # RACKET failed to find the builtin idris support library -# $1 --cg racket --no-color --console-width 0 --no-banner ReadFilePage.idr < input +# idris2 --cg racket ReadFilePage.idr < input # REFC doesn't support :exec yet -# $1 --cg refc --no-color --console-width 0 --no-banner ReadFilePage.idr < input +# idris2 --cg refc ReadFilePage.idr < input # GAMBIT hung seemingly indefinitely -# $1 --cg gambit --no-color --console-width 0 --no-banner ReadFilePage.idr < input - +# idris2 --cg gambit ReadFilePage.idr < input diff --git a/tests/chez/barrier001/run b/tests/chez/barrier001/run index 93b2c4cac8..108c440fda 100644 --- a/tests/chez/barrier001/run +++ b/tests/chez/barrier001/run @@ -1 +1,3 @@ -$1 --no-banner --no-color --console-width 0 --cg chez Main.idr --exec main +. ../../testutils.sh + +run --cg chez Main.idr diff --git a/tests/chez/bitops/run b/tests/chez/bitops/run index 78360db1cd..14e998613d 100644 --- a/tests/chez/bitops/run +++ b/tests/chez/bitops/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-banner --no-color --console-width 0 BitOps.idr < input +. ../../testutils.sh +idris2 BitOps.idr < input diff --git a/tests/chez/buffer001/run b/tests/chez/buffer001/run index c8b4448627..17a81880d7 100755 --- a/tests/chez/buffer001/run +++ b/tests/chez/buffer001/run @@ -1,4 +1,4 @@ -rm -rf build +. ../../testutils.sh -$1 --no-banner --no-color --console-width 0 --exec main Main.idr -rm tmp.bin \ No newline at end of file +run Main.idr +rm tmp.bin diff --git a/tests/chez/casts/run b/tests/chez/casts/run index 48cd2c2f8c..5265730eaa 100644 --- a/tests/chez/casts/run +++ b/tests/chez/casts/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-banner --no-color --console-width 0 Casts.idr < input +. ../../testutils.sh +idris2 Casts.idr < input diff --git a/tests/chez/chez001/run b/tests/chez/chez001/run index 4f97a2d2d7..c1204f255d 100755 --- a/tests/chez/chez001/run +++ b/tests/chez/chez001/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner Total.idr < input +. ../../testutils.sh +idris2 Total.idr < input diff --git a/tests/chez/chez002/run b/tests/chez/chez002/run index 248f4c1244..d24032639a 100755 --- a/tests/chez/chez002/run +++ b/tests/chez/chez002/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --no-banner Pythag.idr < input +idris2 Pythag.idr < input diff --git a/tests/chez/chez003/run b/tests/chez/chez003/run index 802151e040..58796816d0 100755 --- a/tests/chez/chez003/run +++ b/tests/chez/chez003/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --no-banner IORef.idr < input +idris2 IORef.idr < input diff --git a/tests/chez/chez004/run b/tests/chez/chez004/run index 7550be3a4d..87a9948fb4 100755 --- a/tests/chez/chez004/run +++ b/tests/chez/chez004/run @@ -1,2 +1,4 @@ -$1 --no-color --console-width 0 -p contrib --no-banner Buffer.idr < input +. ../../testutils.sh + +idris2 -p contrib Buffer.idr < input rm -rf build test.buf diff --git a/tests/chez/chez005/run b/tests/chez/chez005/run index b2eb87dc0e..653275eef2 100755 --- a/tests/chez/chez005/run +++ b/tests/chez/chez005/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner Filter.idr < input +. ../../testutils.sh +idris2 Filter.idr < input diff --git a/tests/chez/chez006/run b/tests/chez/chez006/run index 5066a82edb..47e40c235d 100755 --- a/tests/chez/chez006/run +++ b/tests/chez/chez006/run @@ -1,5 +1,4 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner TypeCase.idr < input -$1 --no-color --console-width 0 --no-banner TypeCase2.idr --check +. ../../testutils.sh +idris2 TypeCase.idr < input +check TypeCase2.idr diff --git a/tests/chez/chez007/run b/tests/chez/chez007/run index 92e3bd79ba..786567d611 100755 --- a/tests/chez/chez007/run +++ b/tests/chez/chez007/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner TypeCase.idr < input +. ../../testutils.sh +idris2 TypeCase.idr < input diff --git a/tests/chez/chez008/run b/tests/chez/chez008/run index 157908b316..3c4ab0b826 100755 --- a/tests/chez/chez008/run +++ b/tests/chez/chez008/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner Nat.idr < input +. ../../testutils.sh +idris2 Nat.idr < input diff --git a/tests/chez/chez009/run b/tests/chez/chez009/run index 4558792b36..36867b3057 100755 --- a/tests/chez/chez009/run +++ b/tests/chez/chez009/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner uni.idr < input +. ../../testutils.sh +idris2 uni.idr < input diff --git a/tests/chez/chez010/run b/tests/chez/chez010/run index 95cb5d36f2..edca4eb608 100755 --- a/tests/chez/chez010/run +++ b/tests/chez/chez010/run @@ -1,3 +1,5 @@ +. ../../testutils.sh + case $(uname -s) in OpenBSD | FreeBSD | NetBSD) MAKE=gmake @@ -9,6 +11,5 @@ case $(uname -s) in esac ${MAKE} all > /dev/null -$1 --no-color --console-width 0 --no-banner CB.idr < input -rm -rf build +idris2 CB.idr < input ${MAKE} clean > /dev/null diff --git a/tests/chez/chez011/run b/tests/chez/chez011/run index c46f52a988..2f7e60c427 100644 --- a/tests/chez/chez011/run +++ b/tests/chez/chez011/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner bangs.idr < input +. ../../testutils.sh +idris2 bangs.idr < input diff --git a/tests/chez/chez012/run b/tests/chez/chez012/run index f052de2943..7f12edfdfe 100755 --- a/tests/chez/chez012/run +++ b/tests/chez/chez012/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --no-banner array.idr < input +idris2 array.idr < input diff --git a/tests/chez/chez013/run b/tests/chez/chez013/run index 3c756a9c40..2199a3b818 100755 --- a/tests/chez/chez013/run +++ b/tests/chez/chez013/run @@ -1,3 +1,5 @@ +. ../../testutils.sh + case $(uname -s) in OpenBSD | FreeBSD | NetBSD) MAKE=gmake @@ -9,6 +11,5 @@ case $(uname -s) in esac ${MAKE} all > /dev/null -$1 --no-color --console-width 0 --no-banner Struct.idr < input -rm -rf build +idris2 Struct.idr < input ${MAKE} clean > /dev/null diff --git a/tests/chez/chez014/run b/tests/chez/chez014/run index e954dd99e5..c2e5aa740d 100755 --- a/tests/chez/chez014/run +++ b/tests/chez/chez014/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner -p network Echo.idr < input +. ../../testutils.sh +idris2 -p network Echo.idr < input diff --git a/tests/chez/chez015/run b/tests/chez/chez015/run index 9566248a5b..3f99939085 100755 --- a/tests/chez/chez015/run +++ b/tests/chez/chez015/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner Numbers.idr < input +. ../../testutils.sh +idris2 Numbers.idr < input diff --git a/tests/chez/chez016/run b/tests/chez/chez016/run index 8e9ab1e4d0..003a15d3dd 100755 --- a/tests/chez/chez016/run +++ b/tests/chez/chez016/run @@ -1,4 +1,5 @@ +. ../../testutils.sh + cd "folder with spaces" || exit -rm -rf build -$1 --no-color --console-width 0 --no-banner Main.idr < ../input +idris2 Main.idr < ../input diff --git a/tests/chez/chez017/run b/tests/chez/chez017/run index 9202a5e097..a657ed7f39 100755 --- a/tests/chez/chez017/run +++ b/tests/chez/chez017/run @@ -1,4 +1,6 @@ +. ../../testutils.sh + ./gen_expected.sh -$1 --no-color --console-width 0 --no-banner dir.idr < input +idris2 dir.idr < input cat testdir/test.txt -rm -rf build testdir +rm -rf testdir diff --git a/tests/chez/chez018/run b/tests/chez/chez018/run index a9c7b54e0a..2017333d63 100755 --- a/tests/chez/chez018/run +++ b/tests/chez/chez018/run @@ -1,3 +1,5 @@ -$1 --no-color --console-width 0 --no-banner File.idr < input +. ../../testutils.sh + +idris2 File.idr < input rm -rf build testout.txt diff --git a/tests/chez/chez019/run b/tests/chez/chez019/run index 800a5a8fae..0cf3f075d2 100755 --- a/tests/chez/chez019/run +++ b/tests/chez/chez019/run @@ -1,3 +1,5 @@ -$1 --no-color --console-width 0 --no-banner partial.idr < input +. ../../testutils.sh + +idris2 partial.idr < input rm -rf build testout.txt diff --git a/tests/chez/chez020/run b/tests/chez/chez020/run index 8825141b64..0f522a9265 100644 --- a/tests/chez/chez020/run +++ b/tests/chez/chez020/run @@ -1,4 +1,3 @@ -rm -rf build - -POPEN_CMD="$1 --version" $1 --no-color --console-width 0 --no-banner Popen.idr < input +. ../../testutils.sh +POPEN_CMD="$idris2 --version" idris2 Popen.idr < input diff --git a/tests/chez/chez021/run b/tests/chez/chez021/run index f5edd446a0..6d2f3057d2 100644 --- a/tests/chez/chez021/run +++ b/tests/chez/chez021/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner Bits.idr < input +. ../../testutils.sh +idris2 Bits.idr < input diff --git a/tests/chez/chez022/run b/tests/chez/chez022/run index 563e10ce60..45ad78bf70 100755 --- a/tests/chez/chez022/run +++ b/tests/chez/chez022/run @@ -1,3 +1,5 @@ +. ../../testutils.sh + case $(uname -s) in OpenBSD | FreeBSD | NetBSD) MAKE=gmake @@ -9,6 +11,6 @@ case $(uname -s) in esac ${MAKE} all > /dev/null -$1 --no-color --console-width 0 --no-banner usealloc.idr < input -rm -rf build +idris2 usealloc.idr < input +. ../../testutils.sh ${MAKE} clean > /dev/null diff --git a/tests/chez/chez023/run b/tests/chez/chez023/run index 7af45eecb4..c70689a7a7 100644 --- a/tests/chez/chez023/run +++ b/tests/chez/chez023/run @@ -1,3 +1,5 @@ -$1 --no-color --console-width 0 --no-banner File.idr < input +. ../../testutils.sh + +idris2 File.idr < input rm -rf build test.txt diff --git a/tests/chez/chez024/run b/tests/chez/chez024/run index 888dc04d6a..d718d5caee 100644 --- a/tests/chez/chez024/run +++ b/tests/chez/chez024/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner Envy.idr < input +. ../../testutils.sh +idris2 Envy.idr < input diff --git a/tests/chez/chez025/run b/tests/chez/chez025/run index 73297b4c95..a9f3304435 100755 --- a/tests/chez/chez025/run +++ b/tests/chez/chez025/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner runst.idr < input +. ../../testutils.sh +idris2 runst.idr < input diff --git a/tests/chez/chez026/run b/tests/chez/chez026/run index 936548b9ea..d5cd39a195 100755 --- a/tests/chez/chez026/run +++ b/tests/chez/chez026/run @@ -1,4 +1,6 @@ -$1 --build dummy.ipkg +. ../../testutils.sh + +idris2 --build dummy.ipkg ./custom_output/check_dir diff --git a/tests/chez/chez027/run b/tests/chez/chez027/run index b381c8abca..99b2ea84f7 100644 --- a/tests/chez/chez027/run +++ b/tests/chez/chez027/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner -p contrib StringParser.idr < input +. ../../testutils.sh +idris2 -p contrib StringParser.idr < input diff --git a/tests/chez/chez028/run b/tests/chez/chez028/run index 10a691da01..7adaecccc2 100644 --- a/tests/chez/chez028/run +++ b/tests/chez/chez028/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-banner --no-color --console-width 0 -p contrib ExpressionParser.idr < input +. ../../testutils.sh +idris2 -p contrib ExpressionParser.idr < input diff --git a/tests/chez/chez029/run b/tests/chez/chez029/run index a31c07379e..e4a0b380e2 100644 --- a/tests/chez/chez029/run +++ b/tests/chez/chez029/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-banner --no-color --console-width 0 BitCasts.idr < input +. ../../testutils.sh +idris2 BitCasts.idr < input diff --git a/tests/chez/chez030/run b/tests/chez/chez030/run index b6b18a1a93..a919c2ac02 100644 --- a/tests/chez/chez030/run +++ b/tests/chez/chez030/run @@ -1,5 +1,4 @@ -rm -rf build +. ../../testutils.sh -$1 --no-banner --no-color --console-width 0 --directive extraRuntime=extensions.scm -o chez030 ExtraRuntime.idr +idris2 --directive extraRuntime=extensions.scm -o chez030 ExtraRuntime.idr ./build/exec/chez030 - diff --git a/tests/chez/chez031/run b/tests/chez/chez031/run index c67ed5c435..46af2cb242 100644 --- a/tests/chez/chez031/run +++ b/tests/chez/chez031/run @@ -1,6 +1,5 @@ -rm -rf build - -$1 --no-banner --no-color --console-width 0 --cg chez Specifiers.idr --exec main -$1 --no-banner --no-color --console-width 0 --cg chez Specifiers.idr -o build/foo -$1 --no-banner --no-color --console-width 0 --cg chez < input +. ../../testutils.sh +run --cg chez Specifiers.idr +idris2 --cg chez Specifiers.idr -o build/foo +idris2 --cg chez < input diff --git a/tests/chez/chez032/run b/tests/chez/chez032/run index 78360db1cd..14e998613d 100644 --- a/tests/chez/chez032/run +++ b/tests/chez/chez032/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-banner --no-color --console-width 0 BitOps.idr < input +. ../../testutils.sh +idris2 BitOps.idr < input diff --git a/tests/chez/chez033/run b/tests/chez/chez033/run index 94b6417f12..e6679e9548 100644 --- a/tests/chez/chez033/run +++ b/tests/chez/chez033/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-banner --no-color --console-width 0 Main.idr --inc chez < input +idris2 Main.idr --inc chez < input diff --git a/tests/chez/chez034/run b/tests/chez/chez034/run index a586a02e67..990c8486c0 100644 --- a/tests/chez/chez034/run +++ b/tests/chez/chez034/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-banner --no-color --console-width 0 ThreadData.idr < input +. ../../testutils.sh +idris2 ThreadData.idr < input diff --git a/tests/chez/chez035/run b/tests/chez/chez035/run index 895d856e65..d6afc23df6 100755 --- a/tests/chez/chez035/run +++ b/tests/chez/chez035/run @@ -1,10 +1,10 @@ +. ../../testutils.sh + export IDRIS2_INC_CGS= -$1 --no-banner --no-color --console-width 0 --check Mod1.idr < input +check Mod1.idr < input export IDRIS2_INC_CGS=chez -$1 --no-banner --no-color --console-width 0 -o test Mod2.idr < input +idris2 -o test Mod2.idr < input ./build/exec/test ls build/ttc/*/Mod1.so | sed -r "s/.([0-9]){10}//g" - -rm -rf build diff --git a/tests/chez/chez036/run b/tests/chez/chez036/run index 9dd6356087..e7ce11787a 100755 --- a/tests/chez/chez036/run +++ b/tests/chez/chez036/run @@ -1,5 +1,5 @@ -$1 --no-banner --no-color --console-width 0 --exec main Crash.idr -$1 --no-banner --no-color --console-width 0 --exec main Crash2.idr -$1 --no-banner --no-color --console-width 0 --exec main Crash3.idr +. ../../testutils.sh -rm -rf build +run Crash.idr +run Crash2.idr +run Crash3.idr diff --git a/tests/chez/constfold/run b/tests/chez/constfold/run index 302d8691df..8c605a6262 100644 --- a/tests/chez/constfold/run +++ b/tests/chez/constfold/run @@ -1,5 +1,4 @@ -rm -rf build - -$1 --no-banner --no-color --quiet -o fold Fold.idr -$1 --no-banner --no-color --console-width 0 Check.idr < input +. ../../testutils.sh +idris2 --quiet -o fold Fold.idr +idris2 Check.idr < input diff --git a/tests/chez/constfold2/run b/tests/chez/constfold2/run index 302d8691df..8c605a6262 100644 --- a/tests/chez/constfold2/run +++ b/tests/chez/constfold2/run @@ -1,5 +1,4 @@ -rm -rf build - -$1 --no-banner --no-color --quiet -o fold Fold.idr -$1 --no-banner --no-color --console-width 0 Check.idr < input +. ../../testutils.sh +idris2 --quiet -o fold Fold.idr +idris2 Check.idr < input diff --git a/tests/chez/constfold3/run b/tests/chez/constfold3/run index 302d8691df..8c605a6262 100644 --- a/tests/chez/constfold3/run +++ b/tests/chez/constfold3/run @@ -1,5 +1,4 @@ -rm -rf build - -$1 --no-banner --no-color --quiet -o fold Fold.idr -$1 --no-banner --no-color --console-width 0 Check.idr < input +. ../../testutils.sh +idris2 --quiet -o fold Fold.idr +idris2 Check.idr < input diff --git a/tests/chez/forkjoin001/run b/tests/chez/forkjoin001/run index 93b2c4cac8..108c440fda 100644 --- a/tests/chez/forkjoin001/run +++ b/tests/chez/forkjoin001/run @@ -1 +1,3 @@ -$1 --no-banner --no-color --console-width 0 --cg chez Main.idr --exec main +. ../../testutils.sh + +run --cg chez Main.idr diff --git a/tests/chez/futures001/run b/tests/chez/futures001/run index 56e42f7dd1..a1a267de56 100644 --- a/tests/chez/futures001/run +++ b/tests/chez/futures001/run @@ -1,6 +1,8 @@ -$1 --no-banner --no-color --console-width 0 --cg chez Futures.idr -p contrib --exec constant -$1 --no-banner --no-color --console-width 0 --cg chez Futures.idr -p contrib --exec simpleIO -$1 --no-banner --no-color --console-width 0 --cg chez Futures.idr -p contrib --exec erasureAndNonbindTest -$1 --no-banner --no-color --console-width 0 --cg chez Futures.idr -p contrib --exec map +. ../../testutils.sh + +idris2 --cg chez Futures.idr -p contrib --exec constant +idris2 --cg chez Futures.idr -p contrib --exec simpleIO +idris2 --cg chez Futures.idr -p contrib --exec erasureAndNonbindTest +idris2 --cg chez Futures.idr -p contrib --exec map rm -r build diff --git a/tests/chez/inlineiobind/run b/tests/chez/inlineiobind/run index 8158975c91..7939c44787 100644 --- a/tests/chez/inlineiobind/run +++ b/tests/chez/inlineiobind/run @@ -1,4 +1,4 @@ -rm -rf build +. ../../testutils.sh -$1 --no-banner --no-color --quiet -o main Main.idr -$1 --no-banner --no-color --console-width 0 --exec main Check.idr +idris2 --quiet -o main Main.idr +run Check.idr diff --git a/tests/chez/integers/run b/tests/chez/integers/run index c6121d5bc7..d55d4ef58a 100644 --- a/tests/chez/integers/run +++ b/tests/chez/integers/run @@ -1,5 +1,4 @@ -rm -rf build +. ../../testutils.sh -$1 --no-banner --no-color --console-width 0 -o integers TestIntegers.idr > /dev/null +idris2 -o integers TestIntegers.idr > /dev/null ./build/exec/integers - diff --git a/tests/chez/memo/run b/tests/chez/memo/run index 76b5b96e76..2a10ada75f 100644 --- a/tests/chez/memo/run +++ b/tests/chez/memo/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-banner --no-color --console-width 0 Memo.idr < input +. ../../testutils.sh +idris2 Memo.idr < input diff --git a/tests/chez/nat2fin/run b/tests/chez/nat2fin/run index 5ef8006f28..5c619e2494 100644 --- a/tests/chez/nat2fin/run +++ b/tests/chez/nat2fin/run @@ -1,5 +1,4 @@ -rm -rf build - -$1 --no-banner --no-color --quiet -o test Test.idr -$1 --no-banner --no-color --console-width 0 Check.idr < input +. ../../testutils.sh +idris2 --quiet -o test Test.idr +idris2 Check.idr < input diff --git a/tests/chez/newints/run b/tests/chez/newints/run index 02841aef25..de239c9f52 100644 --- a/tests/chez/newints/run +++ b/tests/chez/newints/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-banner --no-color --console-width 0 IntOps.idr < input +. ../../testutils.sh +idris2 IntOps.idr < input diff --git a/tests/chez/perf001/run b/tests/chez/perf001/run index c8214006e6..781211b4b8 100755 --- a/tests/chez/perf001/run +++ b/tests/chez/perf001/run @@ -1,5 +1,4 @@ -rm -rf build +. ../../testutils.sh -$1 -c Fact.idr -o fact +idris2 -c Fact.idr -o fact ./build/exec/fact - diff --git a/tests/chez/reg001/run b/tests/chez/reg001/run index 791bdc20b5..05c56b2e35 100755 --- a/tests/chez/reg001/run +++ b/tests/chez/reg001/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 numbers.idr -x main +. ../../testutils.sh +run numbers.idr diff --git a/tests/chez/semaphores001/run b/tests/chez/semaphores001/run index 93b2c4cac8..108c440fda 100644 --- a/tests/chez/semaphores001/run +++ b/tests/chez/semaphores001/run @@ -1 +1,3 @@ -$1 --no-banner --no-color --console-width 0 --cg chez Main.idr --exec main +. ../../testutils.sh + +run --cg chez Main.idr diff --git a/tests/chez/semaphores002/run b/tests/chez/semaphores002/run index 93b2c4cac8..108c440fda 100644 --- a/tests/chez/semaphores002/run +++ b/tests/chez/semaphores002/run @@ -1 +1,3 @@ -$1 --no-banner --no-color --console-width 0 --cg chez Main.idr --exec main +. ../../testutils.sh + +run --cg chez Main.idr diff --git a/tests/codegen/builtin001/run b/tests/codegen/builtin001/run index 1b2f462719..a971e6b98e 100755 --- a/tests/codegen/builtin001/run +++ b/tests/codegen/builtin001/run @@ -1,6 +1,6 @@ -rm -rf build +. ../../testutils.sh rm Main.cases -$1 --no-color --console-width 0 --no-banner --dumpcases Main.cases -o Main Main.idr +idris2 --dumpcases Main.cases -o Main Main.idr cat Main.cases | sed -E "s/[0-9]+}/N}/g" | sed -E "s/[0-9]+:[0-9]+/L:C/g" | grep 'Main.plus\|Main.bah' diff --git a/tests/codegen/con001/run b/tests/codegen/con001/run index 1ec9f53359..49ed047288 100755 --- a/tests/codegen/con001/run +++ b/tests/codegen/con001/run @@ -1,6 +1,6 @@ -rm -rf build +. ../../testutils.sh rm Main.cases -$1 --no-color --console-width 0 --no-banner --dumpcases Main.cases -o Main Main.idr +idris2 --dumpcases Main.cases -o Main Main.idr cat Main.cases | sed -E "s/[0-9]+}/N}/g" | sed -E "s/[0-9]+:[0-9]+/L:C/g" | sed '/Constructor/!d' diff --git a/tests/codegen/enum/run b/tests/codegen/enum/run index e82c50749f..2cdf4dd5c3 100644 --- a/tests/codegen/enum/run +++ b/tests/codegen/enum/run @@ -1,5 +1,4 @@ -rm -rf build - -$1 --no-banner --no-color --quiet -o enum Enum.idr -$1 --no-banner --no-color --console-width 0 Check.idr < input +. ../../testutils.sh +idris2 --quiet -o enum Enum.idr +idris2 Check.idr < input diff --git a/tests/contrib/json_001/run b/tests/contrib/json_001/run index d90a45cfa3..592466380e 100755 --- a/tests/contrib/json_001/run +++ b/tests/contrib/json_001/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-banner --no-color --console-width 0 -p contrib CharEncoding.idr < input +. ../../testutils.sh +idris2 -p contrib CharEncoding.idr < input diff --git a/tests/contrib/json_002/run b/tests/contrib/json_002/run index 23c2b4b219..49902be073 100644 --- a/tests/contrib/json_002/run +++ b/tests/contrib/json_002/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-banner --no-color --console-width 0 -p contrib ShowJSON.idr < input +idris2 -p contrib ShowJSON.idr < input diff --git a/tests/contrib/json_003/run b/tests/contrib/json_003/run index af1a2457ac..fe94f3e04d 100644 --- a/tests/contrib/json_003/run +++ b/tests/contrib/json_003/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-banner --no-color --console-width 0 -p contrib CastJSON.idr < input +idris2 -p contrib CastJSON.idr < input diff --git a/tests/contrib/lexer/run b/tests/contrib/lexer/run index e45756d6b9..3cdd834a92 100644 --- a/tests/contrib/lexer/run +++ b/tests/contrib/lexer/run @@ -1,3 +1,3 @@ -$1 --no-banner --no-color --console-width 0 -p contrib -x main Test.idr +. ../../testutils.sh -rm -rf build +run -p contrib Test.idr diff --git a/tests/contrib/list_alternating/run b/tests/contrib/list_alternating/run index 7e88e7af1a..ed570b6195 100644 --- a/tests/contrib/list_alternating/run +++ b/tests/contrib/list_alternating/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-banner --no-color --console-width 0 -p contrib AlternatingList.idr < input +idris2 -p contrib AlternatingList.idr < input diff --git a/tests/contrib/perf001/run b/tests/contrib/perf001/run index 2e5e4397f1..436de5bfa6 100755 --- a/tests/contrib/perf001/run +++ b/tests/contrib/perf001/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-banner --no-color --console-width 0 -p contrib --exec main GCDPerf.idr +run -p contrib GCDPerf.idr diff --git a/tests/contrib/sortedmap_001/run b/tests/contrib/sortedmap_001/run index 06b1db43c6..cc69869397 100644 --- a/tests/contrib/sortedmap_001/run +++ b/tests/contrib/sortedmap_001/run @@ -1,3 +1,3 @@ -$1 --no-banner --no-color --console-width 0 -p contrib SortedMapTest.idr < input +. ../../testutils.sh -rm -rf build +idris2 -p contrib SortedMapTest.idr < input diff --git a/tests/contrib/system_directory_tree_copyDir/run b/tests/contrib/system_directory_tree_copyDir/run index ca577828d8..557e7f40fa 100644 --- a/tests/contrib/system_directory_tree_copyDir/run +++ b/tests/contrib/system_directory_tree_copyDir/run @@ -1,6 +1,6 @@ -rm -rf build +. ../../testutils.sh -$1 --no-banner --no-color --console-width 0 -p contrib CopyDir.idr < input +idris2 -p contrib CopyDir.idr < input ls -R resultDir | sed '/^resultDir:$/d' base64 -i resultDir/source.bin cat resultDir/nestedDir/anotherFile.txt diff --git a/tests/gambit/bitops001/run b/tests/gambit/bitops001/run index 821e6b297c..807edbb75c 100755 --- a/tests/gambit/bitops001/run +++ b/tests/gambit/bitops001/run @@ -1,5 +1,6 @@ -$1 --cg gambit --exec main BitOps.idr > output -rm -rf build +. ../../testutils.sh + +run --cg gambit BitOps.idr > output # We finish with cmp so that the test fails if the comparison fails diff expected output diff --git a/tests/ideMode/ideMode001/run b/tests/ideMode/ideMode001/run index 32c5798eae..3d64bf912d 100755 --- a/tests/ideMode/ideMode001/run +++ b/tests/ideMode/ideMode001/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --ide-mode < input +. ../../testutils.sh +idris2 --ide-mode < input diff --git a/tests/ideMode/ideMode002/run b/tests/ideMode/ideMode002/run index 8cc736a5b9..efdff56660 100755 --- a/tests/ideMode/ideMode002/run +++ b/tests/ideMode/ideMode002/run @@ -1,5 +1,4 @@ -rm -rf build - -$1 --no-color --console-width 0 --version | awk '{print $4}' | ./gen_expected.sh -$1 --no-color --console-width 0 --ide-mode < input +. ../../testutils.sh +idris2 --version | awk '{print $4}' | ./gen_expected.sh +idris2 --ide-mode < input diff --git a/tests/ideMode/ideMode003/run b/tests/ideMode/ideMode003/run index 32c5798eae..3d64bf912d 100755 --- a/tests/ideMode/ideMode003/run +++ b/tests/ideMode/ideMode003/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --ide-mode < input +. ../../testutils.sh +idris2 --ide-mode < input diff --git a/tests/ideMode/ideMode004/run b/tests/ideMode/ideMode004/run index 32c5798eae..3d64bf912d 100755 --- a/tests/ideMode/ideMode004/run +++ b/tests/ideMode/ideMode004/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --ide-mode < input +. ../../testutils.sh +idris2 --ide-mode < input diff --git a/tests/ideMode/ideMode005/regenerate b/tests/ideMode/ideMode005/regenerate index c0fd112132..a1dac444cf 100755 --- a/tests/ideMode/ideMode005/regenerate +++ b/tests/ideMode/ideMode005/regenerate @@ -1,25 +1,24 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --ide-mode < input1 > expected1 -$1 --no-color --console-width 0 --ide-mode < input2 > expected2 -$1 --no-color --console-width 0 --ide-mode < input3 > expected3 -$1 --no-color --console-width 0 --ide-mode < input4 > expected4 -$1 --no-color --console-width 0 --ide-mode < input5 > expected5 -$1 --no-color --console-width 0 --ide-mode < input6 > expected6 -$1 --no-color --console-width 0 --ide-mode < input7 > expected7 -$1 --no-color --console-width 0 --ide-mode < input8 > expected8 -$1 --no-color --console-width 0 --ide-mode < input9 > expected9 -$1 --no-color --console-width 0 --ide-mode < inputA > expectedA -$1 --no-color --console-width 0 --ide-mode < inputB > expectedB -$1 --no-color --console-width 0 --ide-mode < inputC > expectedC -$1 --no-color --console-width 0 --ide-mode < inputD > expectedD -$1 --no-color --console-width 0 --ide-mode < inputE > expectedE -$1 --no-color --console-width 0 --ide-mode < inputF > expectedF -$1 --no-color --console-width 0 --ide-mode < inputG > expectedG -$1 --no-color --console-width 0 --ide-mode < inputH > expectedH -$1 --no-color --console-width 0 --ide-mode < inputI > expectedI -$1 --no-color --console-width 0 --ide-mode < inputJ > expectedJ +idris2 --ide-mode < input1 > expected1 +idris2 --ide-mode < input2 > expected2 +idris2 --ide-mode < input3 > expected3 +idris2 --ide-mode < input4 > expected4 +idris2 --ide-mode < input5 > expected5 +idris2 --ide-mode < input6 > expected6 +idris2 --ide-mode < input7 > expected7 +idris2 --ide-mode < input8 > expected8 +idris2 --ide-mode < input9 > expected9 +idris2 --ide-mode < inputA > expectedA +idris2 --ide-mode < inputB > expectedB +idris2 --ide-mode < inputC > expectedC +idris2 --ide-mode < inputD > expectedD +idris2 --ide-mode < inputE > expectedE +idris2 --ide-mode < inputF > expectedF +idris2 --ide-mode < inputG > expectedG +idris2 --ide-mode < inputH > expectedH +idris2 --ide-mode < inputI > expectedI +idris2 --ide-mode < inputJ > expectedJ rm -f expected output* touch expected -rm -rf build diff --git a/tests/ideMode/ideMode005/run b/tests/ideMode/ideMode005/run index 31c1e08a31..3cb5147f28 100755 --- a/tests/ideMode/ideMode005/run +++ b/tests/ideMode/ideMode005/run @@ -1,60 +1,60 @@ -rm -rf build +. ../../testutils.sh rm -f output* -$1 --no-color --console-width 0 --ide-mode < input1 > output1 +idris2 --ide-mode < input1 > output1 diff expected1 output1 >> output -$1 --no-color --console-width 0 --ide-mode < input2 > output2 +idris2 --ide-mode < input2 > output2 diff expected2 output2 >> output -$1 --no-color --console-width 0 --ide-mode < input3 > output3 +idris2 --ide-mode < input3 > output3 diff expected3 output3 >> output -$1 --no-color --console-width 0 --ide-mode < input4 > output4 +idris2 --ide-mode < input4 > output4 diff expected4 output4 >> output -$1 --no-color --console-width 0 --ide-mode < input5 > output5 +idris2 --ide-mode < input5 > output5 diff expected5 output5 >> output -$1 --no-color --console-width 0 --ide-mode < input6 > output6 +idris2 --ide-mode < input6 > output6 diff expected6 output6 >> output -$1 --no-color --console-width 0 --ide-mode < input7 > output7 +idris2 --ide-mode < input7 > output7 diff expected7 output7 >> output -$1 --no-color --console-width 0 --ide-mode < input8 > output8 +idris2 --ide-mode < input8 > output8 diff expected8 output8 >> output -$1 --no-color --console-width 0 --ide-mode < input9 > output9 +idris2 --ide-mode < input9 > output9 diff expected9 output9 >> output -$1 --no-color --console-width 0 --ide-mode < inputA > outputA +idris2 --ide-mode < inputA > outputA diff expectedA outputA >> output -$1 --no-color --console-width 0 --ide-mode < inputB > outputB +idris2 --ide-mode < inputB > outputB diff expectedB outputB >> output -$1 --no-color --console-width 0 --ide-mode < inputC > outputC +idris2 --ide-mode < inputC > outputC diff expectedC outputC >> output -$1 --no-color --console-width 0 --ide-mode < inputD > outputD +idris2 --ide-mode < inputD > outputD diff expectedD outputD >> output -$1 --no-color --console-width 0 --ide-mode < inputE > outputE +idris2 --ide-mode < inputE > outputE diff expectedE outputE >> output -$1 --no-color --console-width 0 --ide-mode < inputF > outputF +idris2 --ide-mode < inputF > outputF diff expectedF outputF >> output -$1 --no-color --console-width 0 --ide-mode < inputG > outputG +idris2 --ide-mode < inputG > outputG diff expectedG outputG >> output -$1 --no-color --console-width 0 --ide-mode < inputH > outputH +idris2 --ide-mode < inputH > outputH diff expectedH outputH >> output -$1 --no-color --console-width 0 --ide-mode < inputI > outputI +idris2 --ide-mode < inputI > outputI diff expectedI outputI >> output -$1 --no-color --console-width 0 --ide-mode < inputJ > outputJ +idris2 --ide-mode < inputJ > outputJ diff expectedJ outputJ >> output diff --git a/tests/ideMode/ideMode006/run b/tests/ideMode/ideMode006/run index 2b56741e58..3d64bf912d 100755 --- a/tests/ideMode/ideMode006/run +++ b/tests/ideMode/ideMode006/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --ide-mode < input \ No newline at end of file +idris2 --ide-mode < input diff --git a/tests/idris2/api001/run b/tests/idris2/api001/run index 20083cfbc6..22a6d70e80 100755 --- a/tests/idris2/api001/run +++ b/tests/idris2/api001/run @@ -1,5 +1,6 @@ -$1 --no-color --console-width 0 --no-banner -p idris2 -p contrib -p network LazyCodegen.idr -o lazy-idris2 -./build/exec/lazy-idris2 --no-color --console-width 0 --no-banner --cg lazy Hello.idr -o hello > output -rm -rf build +. ../../testutils.sh + +idris2 -p idris2 -p contrib -p network LazyCodegen.idr -o lazy-idris2 +./build/exec/lazy-idris2 --cg lazy Hello.idr -o hello > output diff expected output cmp -n 15 expected output diff --git a/tests/idris2/basic001/run b/tests/idris2/basic001/run index fcbabe9f57..a469d81530 100755 --- a/tests/idris2/basic001/run +++ b/tests/idris2/basic001/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner --no-prelude Vect.idr < input +. ../../testutils.sh +idris2 --no-prelude Vect.idr < input diff --git a/tests/idris2/basic002/run b/tests/idris2/basic002/run index faf9ac7d30..95fcc62fa0 100755 --- a/tests/idris2/basic002/run +++ b/tests/idris2/basic002/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner --no-prelude Do.idr < input +. ../../testutils.sh +idris2 --no-prelude Do.idr < input diff --git a/tests/idris2/basic003/run b/tests/idris2/basic003/run index 98dc2dd2a7..e525674551 100755 --- a/tests/idris2/basic003/run +++ b/tests/idris2/basic003/run @@ -1,5 +1,4 @@ -rm -rf build - -echo ':q' | $1 --no-color --console-width 0 --no-banner --no-prelude Ambig1.idr -echo ':q' | $1 --no-color --console-width 0 --no-banner --no-prelude Ambig2.idr +. ../../testutils.sh +echo ':q' | idris2 --no-prelude Ambig1.idr +echo ':q' | idris2 --no-prelude Ambig2.idr diff --git a/tests/idris2/basic004/run b/tests/idris2/basic004/run index 0fa269aa4f..48fbd46850 100755 --- a/tests/idris2/basic004/run +++ b/tests/idris2/basic004/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner --no-prelude Wheres.idr < input +. ../../testutils.sh +idris2 --no-prelude Wheres.idr < input diff --git a/tests/idris2/basic005/run b/tests/idris2/basic005/run index f4d42e472e..6b6218fd5b 100755 --- a/tests/idris2/basic005/run +++ b/tests/idris2/basic005/run @@ -1,4 +1,3 @@ -rm -rf build - -echo ':q' | $1 --no-color --console-width 0 --no-banner --no-prelude NoInfer.idr +. ../../testutils.sh +echo ':q' | idris2 --no-prelude NoInfer.idr diff --git a/tests/idris2/basic006/run b/tests/idris2/basic006/run index 5130a1f6f2..45af0f84c4 100755 --- a/tests/idris2/basic006/run +++ b/tests/idris2/basic006/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner --no-prelude PMLet.idr < input +. ../../testutils.sh +idris2 --no-prelude PMLet.idr < input diff --git a/tests/idris2/basic007/run b/tests/idris2/basic007/run index 043ee81b0e..b1a0170ead 100755 --- a/tests/idris2/basic007/run +++ b/tests/idris2/basic007/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner --no-prelude DoLocal.idr < input +. ../../testutils.sh +idris2 --no-prelude DoLocal.idr < input diff --git a/tests/idris2/basic008/run b/tests/idris2/basic008/run index 2a1c6aa536..82f762e699 100755 --- a/tests/idris2/basic008/run +++ b/tests/idris2/basic008/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner --no-prelude If.idr < input +. ../../testutils.sh +idris2 --no-prelude If.idr < input diff --git a/tests/idris2/basic009/run b/tests/idris2/basic009/run index f88d2016a4..380ff1625d 100755 --- a/tests/idris2/basic009/run +++ b/tests/idris2/basic009/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner --no-prelude LetCase.idr < input +. ../../testutils.sh +idris2 --no-prelude LetCase.idr < input diff --git a/tests/idris2/basic010/run b/tests/idris2/basic010/run index 7bdf320829..be231ce846 100755 --- a/tests/idris2/basic010/run +++ b/tests/idris2/basic010/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner --no-prelude Comp.idr < input +. ../../testutils.sh +idris2 --no-prelude Comp.idr < input diff --git a/tests/idris2/basic011/run b/tests/idris2/basic011/run index 0c94e92eac..fd4d43813b 100755 --- a/tests/idris2/basic011/run +++ b/tests/idris2/basic011/run @@ -1,6 +1,5 @@ -rm -rf build - -$1 --no-color --console-width 0 --check Dots1.idr -$1 --no-color --console-width 0 --check Dots2.idr -$1 --no-color --console-width 0 --check Dots3.idr +. ../../testutils.sh +check Dots1.idr +check Dots2.idr +check Dots3.idr diff --git a/tests/idris2/basic012/run b/tests/idris2/basic012/run index faae785765..dd6ccaa67b 100755 --- a/tests/idris2/basic012/run +++ b/tests/idris2/basic012/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --check VIndex.idr +. ../../testutils.sh +check VIndex.idr diff --git a/tests/idris2/basic013/run b/tests/idris2/basic013/run index 329bd02eb4..34847877d7 100755 --- a/tests/idris2/basic013/run +++ b/tests/idris2/basic013/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --check Implicits.idr +. ../../testutils.sh +check Implicits.idr diff --git a/tests/idris2/basic014/run b/tests/idris2/basic014/run index a960badec3..90447fd424 100755 --- a/tests/idris2/basic014/run +++ b/tests/idris2/basic014/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --check Rewrite.idr +. ../../testutils.sh +check Rewrite.idr diff --git a/tests/idris2/basic015/run b/tests/idris2/basic015/run index 02b470d628..d9f9e59d4e 100755 --- a/tests/idris2/basic015/run +++ b/tests/idris2/basic015/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --check George.idr +. ../../testutils.sh +check George.idr diff --git a/tests/idris2/basic016/run b/tests/idris2/basic016/run index 5fc11ffaf1..59d0d7f294 100755 --- a/tests/idris2/basic016/run +++ b/tests/idris2/basic016/run @@ -1,5 +1,4 @@ -rm -rf build - -$1 --no-color --console-width 0 --check Eta.idr -$1 --no-color --console-width 0 --check Eta2.idr +. ../../testutils.sh +check Eta.idr +check Eta2.idr diff --git a/tests/idris2/basic017/run b/tests/idris2/basic017/run index 4fc774f173..bd22b25656 100755 --- a/tests/idris2/basic017/run +++ b/tests/idris2/basic017/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner CaseInf.idr < input +. ../../testutils.sh +idris2 CaseInf.idr < input diff --git a/tests/idris2/basic018/run b/tests/idris2/basic018/run index 5f47c9b4ac..8b2ca2341f 100755 --- a/tests/idris2/basic018/run +++ b/tests/idris2/basic018/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --check Fin.idr +. ../../testutils.sh +check Fin.idr diff --git a/tests/idris2/basic019/run b/tests/idris2/basic019/run index 76844d05ce..4b9c047980 100755 --- a/tests/idris2/basic019/run +++ b/tests/idris2/basic019/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner CaseBlock.idr < input +. ../../testutils.sh +idris2 CaseBlock.idr < input diff --git a/tests/idris2/basic020/run b/tests/idris2/basic020/run index 8fb16f92d8..990791a5e1 100755 --- a/tests/idris2/basic020/run +++ b/tests/idris2/basic020/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner Mut.idr < input +. ../../testutils.sh +idris2 Mut.idr < input diff --git a/tests/idris2/basic021/run b/tests/idris2/basic021/run index 640b8a4ece..02a25baea9 100755 --- a/tests/idris2/basic021/run +++ b/tests/idris2/basic021/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner CaseDep.idr < input +. ../../testutils.sh +idris2 CaseDep.idr < input diff --git a/tests/idris2/basic022/run b/tests/idris2/basic022/run index 700fb28c1f..2660c7e33c 100755 --- a/tests/idris2/basic022/run +++ b/tests/idris2/basic022/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner Erase.idr < input +. ../../testutils.sh +idris2 Erase.idr < input diff --git a/tests/idris2/basic023/run b/tests/idris2/basic023/run index 1467918f87..d1e7b78c95 100755 --- a/tests/idris2/basic023/run +++ b/tests/idris2/basic023/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner Params.idr < input +. ../../testutils.sh +idris2 Params.idr < input diff --git a/tests/idris2/basic024/run b/tests/idris2/basic024/run index 4f359ddab8..b716b4af86 100755 --- a/tests/idris2/basic024/run +++ b/tests/idris2/basic024/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner PatLam.idr < input +. ../../testutils.sh +idris2 PatLam.idr < input diff --git a/tests/idris2/basic025/run b/tests/idris2/basic025/run index 47008ffee7..fecfb4b107 100755 --- a/tests/idris2/basic025/run +++ b/tests/idris2/basic025/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner < input +. ../../testutils.sh +idris2 < input diff --git a/tests/idris2/basic026/run b/tests/idris2/basic026/run index 6410a0e1a8..6f656217d9 100755 --- a/tests/idris2/basic026/run +++ b/tests/idris2/basic026/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --check Erl.idr +. ../../testutils.sh +check Erl.idr diff --git a/tests/idris2/basic027/run b/tests/idris2/basic027/run index fd57b8850a..823be3e2f8 100755 --- a/tests/idris2/basic027/run +++ b/tests/idris2/basic027/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --check Temp.idr +. ../../testutils.sh +check Temp.idr diff --git a/tests/idris2/basic028/run b/tests/idris2/basic028/run index aa632f8dd5..5122537366 100755 --- a/tests/idris2/basic028/run +++ b/tests/idris2/basic028/run @@ -1,6 +1,5 @@ -rm -rf build +. ../../testutils.sh unset IDRIS2_PATH -$1 --no-color --console-width 0 --no-banner --no-prelude < input - +idris2 --no-prelude < input diff --git a/tests/idris2/basic029/run b/tests/idris2/basic029/run index 1467918f87..d1e7b78c95 100644 --- a/tests/idris2/basic029/run +++ b/tests/idris2/basic029/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner Params.idr < input +. ../../testutils.sh +idris2 Params.idr < input diff --git a/tests/idris2/basic030/run b/tests/idris2/basic030/run index 3e8d8fc4de..2ce64ef3af 100644 --- a/tests/idris2/basic030/run +++ b/tests/idris2/basic030/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --check arity.idr +. ../../testutils.sh +check arity.idr diff --git a/tests/idris2/basic031/run b/tests/idris2/basic031/run index 4650562504..c5cd987fcd 100644 --- a/tests/idris2/basic031/run +++ b/tests/idris2/basic031/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --check erased.idr +. ../../testutils.sh +check erased.idr diff --git a/tests/idris2/basic032/run b/tests/idris2/basic032/run index 383d9af23d..ffe7a5e4a8 100755 --- a/tests/idris2/basic032/run +++ b/tests/idris2/basic032/run @@ -1,5 +1,4 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner Idiom.idr < input -$1 --no-color --console-width 0 --no-banner Idiom2.idr --check +. ../../testutils.sh +idris2 Idiom.idr < input +check Idiom2.idr diff --git a/tests/idris2/basic033/run b/tests/idris2/basic033/run index 996941b23c..9480bdeaa8 100644 --- a/tests/idris2/basic033/run +++ b/tests/idris2/basic033/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --check unboundimps.idr +. ../../testutils.sh +check unboundimps.idr diff --git a/tests/idris2/basic034/run b/tests/idris2/basic034/run index 683cf5e5b8..8fc7532fdb 100644 --- a/tests/idris2/basic034/run +++ b/tests/idris2/basic034/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --check lets.idr +. ../../testutils.sh +check lets.idr diff --git a/tests/idris2/basic035/run b/tests/idris2/basic035/run index fa9b5d43da..63bae6413c 100755 --- a/tests/idris2/basic035/run +++ b/tests/idris2/basic035/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner using.idr < input +. ../../testutils.sh +idris2 using.idr < input diff --git a/tests/idris2/basic036/run b/tests/idris2/basic036/run index 5a43c983a0..b094756dc1 100755 --- a/tests/idris2/basic036/run +++ b/tests/idris2/basic036/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner defimp.idr < input +. ../../testutils.sh +idris2 defimp.idr < input diff --git a/tests/idris2/basic037/run b/tests/idris2/basic037/run index cc9ea54809..9675168217 100644 --- a/tests/idris2/basic037/run +++ b/tests/idris2/basic037/run @@ -1,5 +1,4 @@ -rm -rf build - -echo ':q' | $1 --no-color --console-width 0 --no-banner --no-prelude Comments.idr -echo ':q' | $1 --no-color --console-width 0 --no-banner Issue279.idr +. ../../testutils.sh +echo ':q' | idris2 --no-prelude Comments.idr +echo ':q' | idris2 Issue279.idr diff --git a/tests/idris2/basic038/run b/tests/idris2/basic038/run index d18f562f54..695b0e59b2 100644 --- a/tests/idris2/basic038/run +++ b/tests/idris2/basic038/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner Resugar.idr < input +. ../../testutils.sh +idris2 Resugar.idr < input diff --git a/tests/idris2/basic039/run b/tests/idris2/basic039/run index 6002294133..37beba4789 100755 --- a/tests/idris2/basic039/run +++ b/tests/idris2/basic039/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner Main.idr < input +. ../../testutils.sh +idris2 Main.idr < input diff --git a/tests/idris2/basic040/run b/tests/idris2/basic040/run index c4672b4183..68be330b69 100755 --- a/tests/idris2/basic040/run +++ b/tests/idris2/basic040/run @@ -1,4 +1,3 @@ -rm -rf build - -echo ":q" | $1 --no-color --console-width 0 --no-banner Default.idr +. ../../testutils.sh +echo ":q" | idris2 Default.idr diff --git a/tests/idris2/basic041/run b/tests/idris2/basic041/run index c11100cd4b..c8d35dd935 100755 --- a/tests/idris2/basic041/run +++ b/tests/idris2/basic041/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --check QDo.idr +. ../../testutils.sh +check QDo.idr diff --git a/tests/idris2/basic042/run b/tests/idris2/basic042/run index d02782782e..547c22b06e 100755 --- a/tests/idris2/basic042/run +++ b/tests/idris2/basic042/run @@ -1,5 +1,4 @@ -rm -rf build - -$1 --no-banner --no-color --console-width 0 LiteralsString.idr < input -$1 --no-banner --no-color --console-width 0 LiteralsInteger.idr < input2 +. ../../testutils.sh +idris2 LiteralsString.idr < input +idris2 LiteralsInteger.idr < input2 diff --git a/tests/idris2/basic043/run b/tests/idris2/basic043/run index a31c07379e..e4a0b380e2 100644 --- a/tests/idris2/basic043/run +++ b/tests/idris2/basic043/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-banner --no-color --console-width 0 BitCasts.idr < input +. ../../testutils.sh +idris2 BitCasts.idr < input diff --git a/tests/idris2/basic044/run b/tests/idris2/basic044/run index 62f10f7b3e..0ce8521859 100644 --- a/tests/idris2/basic044/run +++ b/tests/idris2/basic044/run @@ -1,9 +1,8 @@ -rm -rf build +. ../../testutils.sh -echo ":q" | $1 --no-banner --no-color --console-width 0 --log unify.equal:10 --log unify:5 Term.idr \ +echo ":q" | idris2 --log unify.equal:10 --log unify:5 Term.idr \ | sed -E "s/[0-9]+}/N}/g" | sed -E "s/resolved([0-9]+)/resolvedN/g" \ | sed -E "s/case in ([0-9]+)/case in N/g" | sed -E "s/[0-9]+:[0-9]+/L:C/g" -echo ":q" | $1 --no-banner --no-color --console-width 0 --log unify:3 --log elab.ambiguous:5 Vec.idr \ +echo ":q" | idris2 --log unify:3 --log elab.ambiguous:5 Vec.idr \ | sed -E "s/[0-9]+}/N}/g" | sed -E "s/resolved([0-9]+)/resolvedN/g" \ | sed -E "s/case in ([0-9]+)/case in N/g" | sed -E "s/[0-9]+:[0-9]+/L:C/g" - diff --git a/tests/idris2/basic045/run b/tests/idris2/basic045/run index 8875b538f8..64c80a5b21 100644 --- a/tests/idris2/basic045/run +++ b/tests/idris2/basic045/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-banner --no-color --console-width 0 Main.idr --check +. ../../testutils.sh +check Main.idr diff --git a/tests/idris2/basic046/run b/tests/idris2/basic046/run index b4aedb8d1e..f6b12aa546 100644 --- a/tests/idris2/basic046/run +++ b/tests/idris2/basic046/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-banner --no-color --console-width 0 TupleSections.idr < input +. ../../testutils.sh +idris2 TupleSections.idr < input diff --git a/tests/idris2/basic047/run b/tests/idris2/basic047/run index ed93a2939b..2c8b2324b7 100644 --- a/tests/idris2/basic047/run +++ b/tests/idris2/basic047/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-banner --no-color --console-width 0 InterleavingLets.idr < input +. ../../testutils.sh +idris2 InterleavingLets.idr < input diff --git a/tests/idris2/basic049/run b/tests/idris2/basic049/run index 339b3ef5a2..da8d57ecaf 100755 --- a/tests/idris2/basic049/run +++ b/tests/idris2/basic049/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner Fld.idr < input +. ../../testutils.sh +idris2 Fld.idr < input diff --git a/tests/idris2/basic050/run b/tests/idris2/basic050/run index f91f415aa0..15eb96c095 100755 --- a/tests/idris2/basic050/run +++ b/tests/idris2/basic050/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner Ilc.idr < input +. ../../testutils.sh +idris2 Ilc.idr < input diff --git a/tests/idris2/basic051/run b/tests/idris2/basic051/run index 920430324b..83c4419b5d 100755 --- a/tests/idris2/basic051/run +++ b/tests/idris2/basic051/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner --check Issue833.idr +. ../../testutils.sh +check Issue833.idr diff --git a/tests/idris2/basic052/run b/tests/idris2/basic052/run index 531ef4fd79..c127a186b1 100755 --- a/tests/idris2/basic052/run +++ b/tests/idris2/basic052/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner DoubleClBrace.idr < input +. ../../testutils.sh +idris2 DoubleClBrace.idr < input diff --git a/tests/idris2/basic053/run b/tests/idris2/basic053/run index 50eb2b0d4f..235e06f77b 100755 --- a/tests/idris2/basic053/run +++ b/tests/idris2/basic053/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner UnderscoredIntegerLiterals.idr < input +. ../../testutils.sh +idris2 UnderscoredIntegerLiterals.idr < input diff --git a/tests/idris2/basic054/run b/tests/idris2/basic054/run index 1aa981c95a..5c11709ceb 100755 --- a/tests/idris2/basic054/run +++ b/tests/idris2/basic054/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner Issue1023.idr --exec main +. ../../testutils.sh +run Issue1023.idr diff --git a/tests/idris2/basic055/run b/tests/idris2/basic055/run index 78360db1cd..14e998613d 100644 --- a/tests/idris2/basic055/run +++ b/tests/idris2/basic055/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-banner --no-color --console-width 0 BitOps.idr < input +. ../../testutils.sh +idris2 BitOps.idr < input diff --git a/tests/idris2/basic056/run b/tests/idris2/basic056/run index 01694cdd82..1b277263ad 100644 --- a/tests/idris2/basic056/run +++ b/tests/idris2/basic056/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-banner --no-color --console-width 0 DoubleLit.idr < input +. ../../testutils.sh +idris2 DoubleLit.idr < input diff --git a/tests/idris2/basic057/run b/tests/idris2/basic057/run index 3de7e56cb2..eeceab8184 100644 --- a/tests/idris2/basic057/run +++ b/tests/idris2/basic057/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-banner --no-color --console-width 0 --check LetIn.idr +. ../../testutils.sh +check LetIn.idr diff --git a/tests/idris2/basic058/run b/tests/idris2/basic058/run index 1effa895d5..ca076d8403 100644 --- a/tests/idris2/basic058/run +++ b/tests/idris2/basic058/run @@ -1,5 +1,4 @@ -rm -rf build - -$1 --no-banner --no-color --console-width 0 --check DataTypeOp.idr -$1 --no-banner --no-color --console-width 0 --check DataTypeProj.idr +. ../../testutils.sh +check DataTypeOp.idr +check DataTypeProj.idr diff --git a/tests/idris2/basic059/run b/tests/idris2/basic059/run index 16bf2dca31..b2bc2105c0 100644 --- a/tests/idris2/basic059/run +++ b/tests/idris2/basic059/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-banner --no-color --console-width 0 --check MultiClaim.idr +. ../../testutils.sh +check MultiClaim.idr diff --git a/tests/idris2/basic060/run b/tests/idris2/basic060/run index fe73c01fab..ead0a5e691 100755 --- a/tests/idris2/basic060/run +++ b/tests/idris2/basic060/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner Snoc.idr < input +. ../../testutils.sh +idris2 Snoc.idr < input diff --git a/tests/idris2/basic061/run b/tests/idris2/basic061/run index 718b7d8aa6..fccd22fb91 100755 --- a/tests/idris2/basic061/run +++ b/tests/idris2/basic061/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --no-banner IgnoreDo.idr < input +idris2 IgnoreDo.idr < input diff --git a/tests/idris2/basic062/run b/tests/idris2/basic062/run index e00fe1c046..f54b91bf3a 100755 --- a/tests/idris2/basic062/run +++ b/tests/idris2/basic062/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --no-banner --check Issue1943.idr +check Issue1943.idr diff --git a/tests/idris2/basic063/run b/tests/idris2/basic063/run index 49fdf1404c..d7e55ad492 100755 --- a/tests/idris2/basic063/run +++ b/tests/idris2/basic063/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --no-banner NoDeclaration.idr < input +idris2 NoDeclaration.idr < input diff --git a/tests/idris2/basic064/run b/tests/idris2/basic064/run index 7208288c64..8dd2c571f7 100755 --- a/tests/idris2/basic064/run +++ b/tests/idris2/basic064/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --no-banner Issue2072.idr < input +idris2 Issue2072.idr < input diff --git a/tests/idris2/basic065/run b/tests/idris2/basic065/run index 2a767098e0..b8a7cd495b 100755 --- a/tests/idris2/basic065/run +++ b/tests/idris2/basic065/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --no-banner --check Issue215.idr +check Issue215.idr diff --git a/tests/idris2/basic066/run b/tests/idris2/basic066/run index 5eb48c2e1c..478fd5723f 100755 --- a/tests/idris2/basic066/run +++ b/tests/idris2/basic066/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --no-banner --check comment.idr +check comment.idr diff --git a/tests/idris2/basic067/run b/tests/idris2/basic067/run index ad57daacef..64c026cfa2 100755 --- a/tests/idris2/basic067/run +++ b/tests/idris2/basic067/run @@ -1,8 +1,8 @@ -rm -rf build +. ../../testutils.sh echo "unclosed1.idr" -$1 --no-color --console-width 0 --no-banner unclosed1.idr < input +idris2 unclosed1.idr < input echo "unclosed2.idr" -$1 --no-color --console-width 0 --no-banner unclosed2.idr < input +idris2 unclosed2.idr < input echo "unclosed3.idr" -$1 --no-color --console-width 0 --no-banner unclosed3.idr < input +idris2 unclosed3.idr < input diff --git a/tests/idris2/basic068/run b/tests/idris2/basic068/run index 3151d2f433..dc975e3814 100755 --- a/tests/idris2/basic068/run +++ b/tests/idris2/basic068/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --no-banner --exec main Issue2138.idr +run Issue2138.idr diff --git a/tests/idris2/basic069/run b/tests/idris2/basic069/run index 0fa34c289e..d23326c00f 100755 --- a/tests/idris2/basic069/run +++ b/tests/idris2/basic069/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --no-banner --exec main DebugInfo.idr +run DebugInfo.idr diff --git a/tests/idris2/basic070/run b/tests/idris2/basic070/run index 59c5abe156..b3a8f50554 100755 --- a/tests/idris2/basic070/run +++ b/tests/idris2/basic070/run @@ -1,6 +1,6 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --no-banner --check Issue3016.idr -$1 --no-color --console-width 0 --no-banner --check Issue2782.idr -$1 --no-color --console-width 0 --no-banner --check Issue2592.idr -$1 --no-color --console-width 0 --no-banner --check Issue2593.idr +check Issue3016.idr +check Issue2782.idr +check Issue2592.idr +check Issue2593.idr diff --git a/tests/idris2/builtin001/run b/tests/idris2/builtin001/run index d2c00b30c9..441b64ba30 100755 --- a/tests/idris2/builtin001/run +++ b/tests/idris2/builtin001/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner Test.idr < input +. ../../testutils.sh +idris2 Test.idr < input diff --git a/tests/idris2/builtin002/run b/tests/idris2/builtin002/run index d2c00b30c9..441b64ba30 100755 --- a/tests/idris2/builtin002/run +++ b/tests/idris2/builtin002/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner Test.idr < input +. ../../testutils.sh +idris2 Test.idr < input diff --git a/tests/idris2/builtin003/run b/tests/idris2/builtin003/run index d2c00b30c9..441b64ba30 100755 --- a/tests/idris2/builtin003/run +++ b/tests/idris2/builtin003/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner Test.idr < input +. ../../testutils.sh +idris2 Test.idr < input diff --git a/tests/idris2/builtin004/run b/tests/idris2/builtin004/run index d2c00b30c9..441b64ba30 100755 --- a/tests/idris2/builtin004/run +++ b/tests/idris2/builtin004/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner Test.idr < input +. ../../testutils.sh +idris2 Test.idr < input diff --git a/tests/idris2/builtin005/run b/tests/idris2/builtin005/run index d2c00b30c9..441b64ba30 100755 --- a/tests/idris2/builtin005/run +++ b/tests/idris2/builtin005/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner Test.idr < input +. ../../testutils.sh +idris2 Test.idr < input diff --git a/tests/idris2/builtin006/run b/tests/idris2/builtin006/run index d2c00b30c9..441b64ba30 100755 --- a/tests/idris2/builtin006/run +++ b/tests/idris2/builtin006/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner Test.idr < input +. ../../testutils.sh +idris2 Test.idr < input diff --git a/tests/idris2/builtin007/run b/tests/idris2/builtin007/run index d2c00b30c9..441b64ba30 100755 --- a/tests/idris2/builtin007/run +++ b/tests/idris2/builtin007/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner Test.idr < input +. ../../testutils.sh +idris2 Test.idr < input diff --git a/tests/idris2/builtin008/run b/tests/idris2/builtin008/run index d2c00b30c9..441b64ba30 100755 --- a/tests/idris2/builtin008/run +++ b/tests/idris2/builtin008/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner Test.idr < input +. ../../testutils.sh +idris2 Test.idr < input diff --git a/tests/idris2/builtin009/run b/tests/idris2/builtin009/run index d90e293aab..ceb556ba04 100755 --- a/tests/idris2/builtin009/run +++ b/tests/idris2/builtin009/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner --exec main Test.idr +. ../../testutils.sh +run Test.idr diff --git a/tests/idris2/builtin010/run b/tests/idris2/builtin010/run index d2c00b30c9..441b64ba30 100755 --- a/tests/idris2/builtin010/run +++ b/tests/idris2/builtin010/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner Test.idr < input +. ../../testutils.sh +idris2 Test.idr < input diff --git a/tests/idris2/builtin011/run b/tests/idris2/builtin011/run index d90e293aab..ceb556ba04 100755 --- a/tests/idris2/builtin011/run +++ b/tests/idris2/builtin011/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner --exec main Test.idr +. ../../testutils.sh +run Test.idr diff --git a/tests/idris2/builtin012/run b/tests/idris2/builtin012/run index 81c2a84bb1..192255940d 100755 --- a/tests/idris2/builtin012/run +++ b/tests/idris2/builtin012/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner --exec main Issue1799.idr +. ../../testutils.sh +run Issue1799.idr diff --git a/tests/idris2/casetree001/run b/tests/idris2/casetree001/run index 3bf8f02ab5..4d98e0f914 100755 --- a/tests/idris2/casetree001/run +++ b/tests/idris2/casetree001/run @@ -1,5 +1,4 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner --check Issue762.idr -$1 --no-color --console-width 0 --no-banner --check IsS.idr +. ../../testutils.sh +check Issue762.idr +check IsS.idr diff --git a/tests/idris2/casetree002/run b/tests/idris2/casetree002/run index 9e47f80a14..db84ac67fb 100755 --- a/tests/idris2/casetree002/run +++ b/tests/idris2/casetree002/run @@ -1,5 +1,4 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner --check DefaultCases.idr -$1 --no-color --console-width 0 --no-banner --check Issue1079.idr +. ../../testutils.sh +check DefaultCases.idr +check Issue1079.idr diff --git a/tests/idris2/casetree003/run b/tests/idris2/casetree003/run index 3942066fb7..75b9db51d2 100755 --- a/tests/idris2/casetree003/run +++ b/tests/idris2/casetree003/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --no-banner ForcedPats.idr < input +idris2 ForcedPats.idr < input diff --git a/tests/idris2/casetree004/run b/tests/idris2/casetree004/run index bb13c37f35..1c7e49e8ef 100755 --- a/tests/idris2/casetree004/run +++ b/tests/idris2/casetree004/run @@ -1,4 +1,4 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --no-banner LocalArgs.idr --check -$1 --no-color --console-width 0 --no-banner PiMatch.idr --check +check LocalArgs.idr +check PiMatch.idr diff --git a/tests/idris2/coverage001/run b/tests/idris2/coverage001/run index a386b0a6d3..57aa3d817f 100755 --- a/tests/idris2/coverage001/run +++ b/tests/idris2/coverage001/run @@ -1,6 +1,5 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner Vect.idr < input -$1 --no-color --console-width 0 --no-banner -c Vect2.idr -$1 --no-color --console-width 0 --no-banner -c Vect3.idr +. ../../testutils.sh +idris2 Vect.idr < input +idris2 -c Vect2.idr +idris2 -c Vect3.idr diff --git a/tests/idris2/coverage002/run b/tests/idris2/coverage002/run index 68a63b06f0..fcf96b59e0 100755 --- a/tests/idris2/coverage002/run +++ b/tests/idris2/coverage002/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner Vect.idr < input +. ../../testutils.sh +idris2 Vect.idr < input diff --git a/tests/idris2/coverage003/run b/tests/idris2/coverage003/run index b91eb339ab..b64bc24733 100755 --- a/tests/idris2/coverage003/run +++ b/tests/idris2/coverage003/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner Cover.idr < input +. ../../testutils.sh +idris2 Cover.idr < input diff --git a/tests/idris2/coverage004/run b/tests/idris2/coverage004/run index b91eb339ab..b64bc24733 100755 --- a/tests/idris2/coverage004/run +++ b/tests/idris2/coverage004/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner Cover.idr < input +. ../../testutils.sh +idris2 Cover.idr < input diff --git a/tests/idris2/coverage005/run b/tests/idris2/coverage005/run index b91eb339ab..b64bc24733 100755 --- a/tests/idris2/coverage005/run +++ b/tests/idris2/coverage005/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner Cover.idr < input +. ../../testutils.sh +idris2 Cover.idr < input diff --git a/tests/idris2/coverage006/run b/tests/idris2/coverage006/run index 913033700b..e18dcdb4dd 100755 --- a/tests/idris2/coverage006/run +++ b/tests/idris2/coverage006/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner foobar.idr < input +. ../../testutils.sh +idris2 foobar.idr < input diff --git a/tests/idris2/coverage007/run b/tests/idris2/coverage007/run index 163f1d7d45..a27943b96b 100755 --- a/tests/idris2/coverage007/run +++ b/tests/idris2/coverage007/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner eq.idr --check +. ../../testutils.sh +check eq.idr diff --git a/tests/idris2/coverage008/run b/tests/idris2/coverage008/run index 3b9afabd40..94ae606197 100755 --- a/tests/idris2/coverage008/run +++ b/tests/idris2/coverage008/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner wcov.idr < input +. ../../testutils.sh +idris2 wcov.idr < input diff --git a/tests/idris2/coverage009/run b/tests/idris2/coverage009/run index f32f00902c..6628258da8 100755 --- a/tests/idris2/coverage009/run +++ b/tests/idris2/coverage009/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --check unreachable.idr +. ../../testutils.sh +check unreachable.idr diff --git a/tests/idris2/coverage010/run b/tests/idris2/coverage010/run index 141b56a429..4072962b12 100755 --- a/tests/idris2/coverage010/run +++ b/tests/idris2/coverage010/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --check casetot.idr +. ../../testutils.sh +check casetot.idr diff --git a/tests/idris2/coverage011/run b/tests/idris2/coverage011/run index cebaa1c8e6..6558cbca37 100755 --- a/tests/idris2/coverage011/run +++ b/tests/idris2/coverage011/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --check Sing.idr +. ../../testutils.sh +check Sing.idr diff --git a/tests/idris2/coverage012/run b/tests/idris2/coverage012/run index 281d97b618..b73b3fcb12 100755 --- a/tests/idris2/coverage012/run +++ b/tests/idris2/coverage012/run @@ -1,5 +1,4 @@ -rm -rf build - -$1 --no-color --console-width 0 --check Issue899.idr -$1 --no-color --console-width 0 --check Issue484.idr +. ../../testutils.sh +check Issue899.idr +check Issue484.idr diff --git a/tests/idris2/coverage013/run b/tests/idris2/coverage013/run index 32ba223928..c922218e6c 100755 --- a/tests/idris2/coverage013/run +++ b/tests/idris2/coverage013/run @@ -1,5 +1,4 @@ -rm -rf build - -$1 --no-color --console-width 0 --check Issue1022.idr -$1 --no-color --console-width 0 --check Issue1022-Refl.idr +. ../../testutils.sh +check Issue1022.idr +check Issue1022-Refl.idr diff --git a/tests/idris2/coverage014/run b/tests/idris2/coverage014/run index 3d6392d1d5..8bf79ef26b 100755 --- a/tests/idris2/coverage014/run +++ b/tests/idris2/coverage014/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --check Issue794.idr +. ../../testutils.sh +check Issue794.idr diff --git a/tests/idris2/coverage015/run b/tests/idris2/coverage015/run index e7aac6ccb0..f9bf92bdc1 100755 --- a/tests/idris2/coverage015/run +++ b/tests/idris2/coverage015/run @@ -1,5 +1,4 @@ -rm -rf build - -$1 --no-color --console-width 0 --check Issue1169.idr -$1 --no-color --console-width 0 --check Issue1366.idr +. ../../testutils.sh +check Issue1169.idr +check Issue1366.idr diff --git a/tests/idris2/coverage016/run b/tests/idris2/coverage016/run index 8f71348f66..8298d0c4b0 100755 --- a/tests/idris2/coverage016/run +++ b/tests/idris2/coverage016/run @@ -1,5 +1,4 @@ -rm -rf build - -$1 --no-color --console-width 0 --check Issue633.idr -$1 --no-color --console-width 0 --check Issue633-2.idr +. ../../testutils.sh +check Issue633.idr +check Issue633-2.idr diff --git a/tests/idris2/coverage017/run b/tests/idris2/coverage017/run index 2973b7418d..7ce70e8b4c 100755 --- a/tests/idris2/coverage017/run +++ b/tests/idris2/coverage017/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --check Issue1421.idr +. ../../testutils.sh +check Issue1421.idr diff --git a/tests/idris2/coverage018/run b/tests/idris2/coverage018/run index da2293bfad..01c5ae7cb9 100755 --- a/tests/idris2/coverage018/run +++ b/tests/idris2/coverage018/run @@ -1,4 +1,4 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --no-banner --console-width 0 Issue1831_1.idr < input -$1 --no-color --no-banner --console-width 0 Issue1831_2.idr < input +idris2 Issue1831_1.idr < input +idris2 Issue1831_2.idr < input diff --git a/tests/idris2/coverage019/run b/tests/idris2/coverage019/run index 837444f926..4be2ab7098 100755 --- a/tests/idris2/coverage019/run +++ b/tests/idris2/coverage019/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --no-banner --console-width 0 --check Issue1632.idr +check Issue1632.idr diff --git a/tests/idris2/data001/run b/tests/idris2/data001/run index 254b1a67c8..bc737deb90 100755 --- a/tests/idris2/data001/run +++ b/tests/idris2/data001/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner --check TestImpl.idr +. ../../testutils.sh +check TestImpl.idr diff --git a/tests/idris2/data002/run b/tests/idris2/data002/run index b220b2656a..441b64ba30 100755 --- a/tests/idris2/data002/run +++ b/tests/idris2/data002/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --no-banner Test.idr < input +idris2 Test.idr < input diff --git a/tests/idris2/debug001/run b/tests/idris2/debug001/run index 85fc0e384e..177aeac57a 100755 --- a/tests/idris2/debug001/run +++ b/tests/idris2/debug001/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --no-banner TypePat.idr < input +idris2 TypePat.idr < input diff --git a/tests/idris2/docs001/run b/tests/idris2/docs001/run index 47008ffee7..fecfb4b107 100755 --- a/tests/idris2/docs001/run +++ b/tests/idris2/docs001/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner < input +. ../../testutils.sh +idris2 < input diff --git a/tests/idris2/docs002/run b/tests/idris2/docs002/run index 1d831834bb..75f6c27108 100755 --- a/tests/idris2/docs002/run +++ b/tests/idris2/docs002/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner Doc.idr < input +. ../../testutils.sh +idris2 Doc.idr < input diff --git a/tests/idris2/docs003/run b/tests/idris2/docs003/run index 8db60e4fa5..0dfbac62c4 100755 --- a/tests/idris2/docs003/run +++ b/tests/idris2/docs003/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner RecordDoc.idr < input +. ../../testutils.sh +idris2 RecordDoc.idr < input diff --git a/tests/idris2/docs004/run b/tests/idris2/docs004/run index d5ea4da796..5a4328a2f6 100755 --- a/tests/idris2/docs004/run +++ b/tests/idris2/docs004/run @@ -1,4 +1,4 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --no-banner DocImpl.idr < input -$1 --no-color --console-width 0 --no-banner --no-prelude List.idr < input2 +idris2 DocImpl.idr < input +idris2 --no-prelude List.idr < input2 diff --git a/tests/idris2/docs005/run b/tests/idris2/docs005/run index 30854d92ee..fecfb4b107 100755 --- a/tests/idris2/docs005/run +++ b/tests/idris2/docs005/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --no-banner < input +idris2 < input diff --git a/tests/idris2/dotted001/run b/tests/idris2/dotted001/run index b2a7e58902..ee8de68057 100755 --- a/tests/idris2/dotted001/run +++ b/tests/idris2/dotted001/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --no-banner -c Issue2726.idr +idris2 -c Issue2726.idr diff --git a/tests/idris2/error001/run b/tests/idris2/error001/run index f267c81d6d..ee1151b23a 100755 --- a/tests/idris2/error001/run +++ b/tests/idris2/error001/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --check Error.idr +. ../../testutils.sh +check Error.idr diff --git a/tests/idris2/error002/run b/tests/idris2/error002/run index f267c81d6d..ee1151b23a 100755 --- a/tests/idris2/error002/run +++ b/tests/idris2/error002/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --check Error.idr +. ../../testutils.sh +check Error.idr diff --git a/tests/idris2/error003/run b/tests/idris2/error003/run index f267c81d6d..ee1151b23a 100755 --- a/tests/idris2/error003/run +++ b/tests/idris2/error003/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --check Error.idr +. ../../testutils.sh +check Error.idr diff --git a/tests/idris2/error004/run b/tests/idris2/error004/run index 60196a4e18..7f7dc5cfe6 100755 --- a/tests/idris2/error004/run +++ b/tests/idris2/error004/run @@ -1,5 +1,4 @@ -rm -rf build - -$1 --no-color --console-width 0 --check Error1.idr -$1 --no-color --console-width 0 --check Error2.idr +. ../../testutils.sh +check Error1.idr +check Error2.idr diff --git a/tests/idris2/error005/run b/tests/idris2/error005/run index 944f42f7af..c788397c11 100755 --- a/tests/idris2/error005/run +++ b/tests/idris2/error005/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --check IfErr.idr +. ../../testutils.sh +check IfErr.idr diff --git a/tests/idris2/error006/run b/tests/idris2/error006/run index 944f42f7af..c788397c11 100755 --- a/tests/idris2/error006/run +++ b/tests/idris2/error006/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --check IfErr.idr +. ../../testutils.sh +check IfErr.idr diff --git a/tests/idris2/error007/run b/tests/idris2/error007/run index 3c883a3a8a..c1a31e8348 100755 --- a/tests/idris2/error007/run +++ b/tests/idris2/error007/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --check CongErr.idr +. ../../testutils.sh +check CongErr.idr diff --git a/tests/idris2/error008/run b/tests/idris2/error008/run index 9af77ed8a9..1f05596214 100755 --- a/tests/idris2/error008/run +++ b/tests/idris2/error008/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner DoesntExist.idr < input +. ../../testutils.sh +idris2 DoesntExist.idr < input diff --git a/tests/idris2/error009/run b/tests/idris2/error009/run index 7cdd60b994..190587fa85 100755 --- a/tests/idris2/error009/run +++ b/tests/idris2/error009/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner Exists.idr < input +. ../../testutils.sh +idris2 Exists.idr < input diff --git a/tests/idris2/error010/run b/tests/idris2/error010/run index 2eb120f7ae..62d1d3a6b3 100755 --- a/tests/idris2/error010/run +++ b/tests/idris2/error010/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 Loop.idr --check +. ../../testutils.sh +check Loop.idr diff --git a/tests/idris2/error011/run b/tests/idris2/error011/run index 13e9190a91..20825f8720 100755 --- a/tests/idris2/error011/run +++ b/tests/idris2/error011/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 ConstructorDuplicate.idr --check +. ../../testutils.sh +check ConstructorDuplicate.idr diff --git a/tests/idris2/error012/run b/tests/idris2/error012/run index 4a08d43afa..36898f13bf 100755 --- a/tests/idris2/error012/run +++ b/tests/idris2/error012/run @@ -1 +1,3 @@ -$1 --no-color --console-width 0 nothere.idr -o argh +. ../../../testutils.sh + +idris2 nothere.idr -o argh diff --git a/tests/idris2/error013/run b/tests/idris2/error013/run index d712f9313d..a7f294e9f2 100755 --- a/tests/idris2/error013/run +++ b/tests/idris2/error013/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 Issue361.idr --check +. ../../testutils.sh +check Issue361.idr diff --git a/tests/idris2/error014/run b/tests/idris2/error014/run index 69c119d285..4d51f545fc 100755 --- a/tests/idris2/error014/run +++ b/tests/idris2/error014/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 Issue735.idr --check +. ../../testutils.sh +check Issue735.idr diff --git a/tests/idris2/error015/run b/tests/idris2/error015/run index 7265d73ec0..e659beb822 100755 --- a/tests/idris2/error015/run +++ b/tests/idris2/error015/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 Issue110.idr --check +. ../../testutils.sh +check Issue110.idr diff --git a/tests/idris2/error016/run b/tests/idris2/error016/run index 7b96e2924e..b1b3a6b727 100755 --- a/tests/idris2/error016/run +++ b/tests/idris2/error016/run @@ -1,5 +1,4 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner --check Issue1230.idr -$1 --no-color --console-width 0 --no-banner < input +. ../../testutils.sh +check Issue1230.idr +idris2 < input diff --git a/tests/idris2/error017/run b/tests/idris2/error017/run index 1cb409bd1e..f5422b8d46 100755 --- a/tests/idris2/error017/run +++ b/tests/idris2/error017/run @@ -1,5 +1,4 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner --check Issue962.idr -$1 --no-color --console-width 0 --no-banner --check Issue962-case.idr +. ../../testutils.sh +check Issue962.idr +check Issue962-case.idr diff --git a/tests/idris2/error018/run b/tests/idris2/error018/run index 5779342140..19747197aa 100755 --- a/tests/idris2/error018/run +++ b/tests/idris2/error018/run @@ -1,7 +1,6 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner --check Issue1031.idr -$1 --no-color --console-width 0 --no-banner --check Issue1031-2.idr -$1 --no-color --console-width 0 --no-banner --check Issue1031-3.idr -$1 --no-color --console-width 0 --no-banner --check Issue1031-4.idr +. ../../testutils.sh +check Issue1031.idr +check Issue1031-2.idr +check Issue1031-3.idr +check Issue1031-4.idr diff --git a/tests/idris2/error019/run b/tests/idris2/error019/run index ba2affa453..043b1a25e1 100644 --- a/tests/idris2/error019/run +++ b/tests/idris2/error019/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner -Werror --check Error.idr +. ../../testutils.sh +check -Werror Error.idr diff --git a/tests/idris2/error020/run b/tests/idris2/error020/run index ba2affa453..043b1a25e1 100644 --- a/tests/idris2/error020/run +++ b/tests/idris2/error020/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner -Werror --check Error.idr +. ../../testutils.sh +check -Werror Error.idr diff --git a/tests/idris2/error021/run b/tests/idris2/error021/run index bed0bf023c..25a8e8c361 100644 --- a/tests/idris2/error021/run +++ b/tests/idris2/error021/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner -Werror --check DeepAmbig.idr +. ../../testutils.sh +check -Werror DeepAmbig.idr diff --git a/tests/idris2/error022/run b/tests/idris2/error022/run index 793ea369c3..6a9cbb2793 100644 --- a/tests/idris2/error022/run +++ b/tests/idris2/error022/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --no-banner --check UpdateLoc.idr +check UpdateLoc.idr diff --git a/tests/idris2/error023/run b/tests/idris2/error023/run index e81b6a8c2a..7f7dc5cfe6 100755 --- a/tests/idris2/error023/run +++ b/tests/idris2/error023/run @@ -1,4 +1,4 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --check Error1.idr -$1 --no-color --console-width 0 --check Error2.idr +check Error1.idr +check Error2.idr diff --git a/tests/idris2/error024/run b/tests/idris2/error024/run index 7c0c4c3b09..2f0cc21131 100755 --- a/tests/idris2/error024/run +++ b/tests/idris2/error024/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --check Error1.idr +check Error1.idr diff --git a/tests/idris2/error025/run b/tests/idris2/error025/run index 53f240257e..99d726ae76 100755 --- a/tests/idris2/error025/run +++ b/tests/idris2/error025/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --check IAlternativePrints.idr +check IAlternativePrints.idr diff --git a/tests/idris2/error026/run b/tests/idris2/error026/run index d4ac815fd5..b2ce794b2f 100755 --- a/tests/idris2/error026/run +++ b/tests/idris2/error026/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --check DoBlockFC.idr +check DoBlockFC.idr diff --git a/tests/idris2/error027/run b/tests/idris2/error027/run index b8754c9c65..ce22b719e5 100755 --- a/tests/idris2/error027/run +++ b/tests/idris2/error027/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --exec main Issue2950.idr +run Issue2950.idr diff --git a/tests/idris2/eta001/run b/tests/idris2/eta001/run index 66d7a54003..48614412aa 100755 --- a/tests/idris2/eta001/run +++ b/tests/idris2/eta001/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner --check Issue1370.idr +. ../../testutils.sh +check Issue1370.idr diff --git a/tests/idris2/evaluator001/run b/tests/idris2/evaluator001/run index db7ddef62d..f04a7b6e84 100755 --- a/tests/idris2/evaluator001/run +++ b/tests/idris2/evaluator001/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner --check Issue650.idr +. ../../testutils.sh +check Issue650.idr diff --git a/tests/idris2/evaluator002/run b/tests/idris2/evaluator002/run index 26a44c7850..2ab3dca6ed 100644 --- a/tests/idris2/evaluator002/run +++ b/tests/idris2/evaluator002/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-banner --no-color --console-width 0 --log eval.stuck:5 Main.idr < input +. ../../testutils.sh +idris2 --log eval.stuck:5 Main.idr < input diff --git a/tests/idris2/evaluator003/run b/tests/idris2/evaluator003/run index 3be99a916d..8c9eec4418 100644 --- a/tests/idris2/evaluator003/run +++ b/tests/idris2/evaluator003/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-banner --no-color --console-width 0 Issue705.idr < input +. ../../testutils.sh +idris2 Issue705.idr < input diff --git a/tests/idris2/evaluator004/run b/tests/idris2/evaluator004/run index 390b66089d..a5bc53cb97 100644 --- a/tests/idris2/evaluator004/run +++ b/tests/idris2/evaluator004/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-banner --no-color --console-width 0 Issue1282.idr < input +idris2 Issue1282.idr < input diff --git a/tests/idris2/failing001/run b/tests/idris2/failing001/run index d2bf1566ad..c296850235 100755 --- a/tests/idris2/failing001/run +++ b/tests/idris2/failing001/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --check Fail.idr +check Fail.idr diff --git a/tests/idris2/failing002/run b/tests/idris2/failing002/run index 841aa567b8..78810e46a9 100755 --- a/tests/idris2/failing002/run +++ b/tests/idris2/failing002/run @@ -1,8 +1,8 @@ -rm -rf build -$1 --no-color --console-width 0 --check FailingBug.idr +. ../../testutils.sh +check FailingBug.idr -rm -rf build -$1 --no-color --console-width 20 --check FailingBug.idr +. ../../testutils.sh +check --console-width 20 FailingBug.idr -rm -rf build -$1 --no-color --console-width 80 --check FailingBug.idr +. ../../testutils.sh +check --console-width 80 FailingBug.idr diff --git a/tests/idris2/failing003/run b/tests/idris2/failing003/run index 15f263eb15..4c16dd5097 100755 --- a/tests/idris2/failing003/run +++ b/tests/idris2/failing003/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --check FailingTotality.idr \ No newline at end of file +check FailingTotality.idr diff --git a/tests/idris2/failing004/run b/tests/idris2/failing004/run index 2161d97f23..fe450082c0 100755 --- a/tests/idris2/failing004/run +++ b/tests/idris2/failing004/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --check Issue2821.idr \ No newline at end of file +check Issue2821.idr diff --git a/tests/idris2/golden001/000-hello/run b/tests/idris2/golden001/000-hello/run old mode 100755 new mode 100644 diff --git a/tests/idris2/golden001/run b/tests/idris2/golden001/run index 527402fc1a..0b64af4e20 100755 --- a/tests/idris2/golden001/run +++ b/tests/idris2/golden001/run @@ -1,8 +1,6 @@ -rm -rf build +. ../../testutils.sh -$1 --build hello.ipkg -$1 --build test.ipkg +idris2 --build hello.ipkg +idris2 --build test.ipkg ./build/exec/runtests ../build/exec/hello --no-colour - - diff --git a/tests/idris2/idiom001/run b/tests/idris2/idiom001/run index 6002294133..37beba4789 100644 --- a/tests/idris2/idiom001/run +++ b/tests/idris2/idiom001/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner Main.idr < input +. ../../testutils.sh +idris2 Main.idr < input diff --git a/tests/idris2/import001/run b/tests/idris2/import001/run index 8a9a2baf54..5a383000ed 100755 --- a/tests/idris2/import001/run +++ b/tests/idris2/import001/run @@ -1,12 +1,11 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --no-banner --no-prelude Test.idr < input +idris2 --no-prelude Test.idr < input sleep 1 touch Mult.idr -$1 --no-color --console-width 0 --no-banner --no-prelude Test.idr < input +idris2 --no-prelude Test.idr < input # Disabled until we have a portable way to do this #sleep 1 #touch Mult.idr -#$1 --no-color --console-width 0 --no-banner --no-prelude -Xcheck-hashes Test.idr < input - +#idris2 --no-prelude -Xcheck-hashes Test.idr < input diff --git a/tests/idris2/import002/run b/tests/idris2/import002/run index 47d7ddd9a8..7ce7a33b8f 100755 --- a/tests/idris2/import002/run +++ b/tests/idris2/import002/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --no-banner --no-prelude --check Test.idr +check --no-prelude Test.idr diff --git a/tests/idris2/import003/run b/tests/idris2/import003/run index 614c7ba86c..dcb7fa69b6 100755 --- a/tests/idris2/import003/run +++ b/tests/idris2/import003/run @@ -1,5 +1,4 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner B.idr < input -$1 --no-color --console-width 0 --no-banner C.idr < input +. ../../testutils.sh +idris2 B.idr < input +idris2 C.idr < input diff --git a/tests/idris2/import004/run b/tests/idris2/import004/run index 88536f8c18..6cfa986dca 100755 --- a/tests/idris2/import004/run +++ b/tests/idris2/import004/run @@ -1,5 +1,4 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner Cycle1.idr -$1 --no-color --console-width 0 --no-banner Loop.idr +. ../../testutils.sh +idris2 Cycle1.idr +idris2 Loop.idr diff --git a/tests/idris2/import005/run b/tests/idris2/import005/run index e90e564cb9..867d00c171 100644 --- a/tests/idris2/import005/run +++ b/tests/idris2/import005/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner As.idr < input +. ../../testutils.sh +idris2 As.idr < input diff --git a/tests/idris2/import006/run b/tests/idris2/import006/run index 1f61ee8ca5..b79def1a3d 100644 --- a/tests/idris2/import006/run +++ b/tests/idris2/import006/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --no-banner --build cyclic.ipkg +idris2 --build cyclic.ipkg diff --git a/tests/idris2/import007/run b/tests/idris2/import007/run index 62fe65c073..fc9dabdf6a 100644 --- a/tests/idris2/import007/run +++ b/tests/idris2/import007/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --no-banner --check Mod.idr +check Mod.idr diff --git a/tests/idris2/import008/run b/tests/idris2/import008/run index 5b11a405e9..cd798e25c4 100644 --- a/tests/idris2/import008/run +++ b/tests/idris2/import008/run @@ -1,9 +1,11 @@ +. ../../testutils.sh + rm -rf Exe/build rm -rf Lib/build cd Lib -$1 --build +idris2 --build cd .. cd Exe -$1 --no-color --console-width 0 --no-banner --check --find-ipkg Mod.idr -$1 --no-color --console-width 0 --no-banner --exec main --find-ipkg Mod.idr +check --find-ipkg Mod.idr +run --find-ipkg Mod.idr diff --git a/tests/idris2/import009/run b/tests/idris2/import009/run index 277a2e07ca..c3e8173aff 100755 --- a/tests/idris2/import009/run +++ b/tests/idris2/import009/run @@ -1,6 +1,6 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --no-banner --check Test.idr -$1 --no-color --console-width 0 --no-banner --check Import.idr +check Test.idr +check Import.idr -$1 --no-color --console-width 0 --no-banner --check Resugar.idr +check Resugar.idr diff --git a/tests/idris2/inlining001/run b/tests/idris2/inlining001/run index 51cfa7983c..7cef0d3038 100644 --- a/tests/idris2/inlining001/run +++ b/tests/idris2/inlining001/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner Inlining.idr < input +. ../../testutils.sh +idris2 Inlining.idr < input diff --git a/tests/idris2/interactive001/run b/tests/idris2/interactive001/run index 4de3af7ea8..6d4602837f 100755 --- a/tests/idris2/interactive001/run +++ b/tests/idris2/interactive001/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner LocType.idr < input +. ../../testutils.sh +idris2 LocType.idr < input diff --git a/tests/idris2/interactive002/run b/tests/idris2/interactive002/run index a5983e604a..ec050c77b1 100755 --- a/tests/idris2/interactive002/run +++ b/tests/idris2/interactive002/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner IEdit.idr < input +. ../../testutils.sh +idris2 IEdit.idr < input diff --git a/tests/idris2/interactive003/run b/tests/idris2/interactive003/run index b826861556..45d7c58b5b 100755 --- a/tests/idris2/interactive003/run +++ b/tests/idris2/interactive003/run @@ -1,5 +1,4 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner IEdit.idr < input -$1 --no-color --console-width 0 --no-banner IEdit2.idr < input2 +. ../../testutils.sh +idris2 IEdit.idr < input +idris2 IEdit2.idr < input2 diff --git a/tests/idris2/interactive004/run b/tests/idris2/interactive004/run index a5983e604a..ec050c77b1 100755 --- a/tests/idris2/interactive004/run +++ b/tests/idris2/interactive004/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner IEdit.idr < input +. ../../testutils.sh +idris2 IEdit.idr < input diff --git a/tests/idris2/interactive005/run b/tests/idris2/interactive005/run index a5983e604a..ec050c77b1 100755 --- a/tests/idris2/interactive005/run +++ b/tests/idris2/interactive005/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner IEdit.idr < input +. ../../testutils.sh +idris2 IEdit.idr < input diff --git a/tests/idris2/interactive006/run b/tests/idris2/interactive006/run index a5983e604a..ec050c77b1 100755 --- a/tests/idris2/interactive006/run +++ b/tests/idris2/interactive006/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner IEdit.idr < input +. ../../testutils.sh +idris2 IEdit.idr < input diff --git a/tests/idris2/interactive007/run b/tests/idris2/interactive007/run index a5983e604a..ec050c77b1 100755 --- a/tests/idris2/interactive007/run +++ b/tests/idris2/interactive007/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner IEdit.idr < input +. ../../testutils.sh +idris2 IEdit.idr < input diff --git a/tests/idris2/interactive008/run b/tests/idris2/interactive008/run index a5983e604a..ec050c77b1 100755 --- a/tests/idris2/interactive008/run +++ b/tests/idris2/interactive008/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner IEdit.idr < input +. ../../testutils.sh +idris2 IEdit.idr < input diff --git a/tests/idris2/interactive009/run b/tests/idris2/interactive009/run index 770797118d..e6c63e9986 100755 --- a/tests/idris2/interactive009/run +++ b/tests/idris2/interactive009/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner Door.idr < input +. ../../testutils.sh +idris2 Door.idr < input diff --git a/tests/idris2/interactive010/run b/tests/idris2/interactive010/run index a5983e604a..ec050c77b1 100755 --- a/tests/idris2/interactive010/run +++ b/tests/idris2/interactive010/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner IEdit.idr < input +. ../../testutils.sh +idris2 IEdit.idr < input diff --git a/tests/idris2/interactive011/run b/tests/idris2/interactive011/run index a5983e604a..ec050c77b1 100755 --- a/tests/idris2/interactive011/run +++ b/tests/idris2/interactive011/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner IEdit.idr < input +. ../../testutils.sh +idris2 IEdit.idr < input diff --git a/tests/idris2/interactive012/run b/tests/idris2/interactive012/run index fbbe7668fd..bf4884fe0e 100755 --- a/tests/idris2/interactive012/run +++ b/tests/idris2/interactive012/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner WithLift.idr < input +. ../../testutils.sh +idris2 WithLift.idr < input diff --git a/tests/idris2/interactive013/run b/tests/idris2/interactive013/run index ec01540fff..8f9217658b 100755 --- a/tests/idris2/interactive013/run +++ b/tests/idris2/interactive013/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner Spacing.idr < input +. ../../testutils.sh +idris2 Spacing.idr < input diff --git a/tests/idris2/interactive014/run b/tests/idris2/interactive014/run index 2fc948b6d6..f48eb78c45 100755 --- a/tests/idris2/interactive014/run +++ b/tests/idris2/interactive014/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner case.idr < input +. ../../testutils.sh +idris2 case.idr < input diff --git a/tests/idris2/interactive015/run b/tests/idris2/interactive015/run index a5983e604a..ec050c77b1 100755 --- a/tests/idris2/interactive015/run +++ b/tests/idris2/interactive015/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner IEdit.idr < input +. ../../testutils.sh +idris2 IEdit.idr < input diff --git a/tests/idris2/interactive016/run b/tests/idris2/interactive016/run index 76f198ec86..bd22d7ef88 100755 --- a/tests/idris2/interactive016/run +++ b/tests/idris2/interactive016/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-banner --no-color --console-width 0 Cont.idr < input +. ../../testutils.sh +idris2 Cont.idr < input diff --git a/tests/idris2/interactive017/run b/tests/idris2/interactive017/run index 139e5957af..ddc5e342c0 100755 --- a/tests/idris2/interactive017/run +++ b/tests/idris2/interactive017/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner RLE.idr < input +. ../../testutils.sh +idris2 RLE.idr < input diff --git a/tests/idris2/interactive018/run b/tests/idris2/interactive018/run index 935671bb20..7ae00127b5 100755 --- a/tests/idris2/interactive018/run +++ b/tests/idris2/interactive018/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-banner --no-color --console-width 0 PlusPrf.idr < input +. ../../testutils.sh +idris2 PlusPrf.idr < input diff --git a/tests/idris2/interactive019/run b/tests/idris2/interactive019/run index cb046c64f1..2a211674dd 100755 --- a/tests/idris2/interactive019/run +++ b/tests/idris2/interactive019/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-banner --no-color --console-width 0 TypeSearch.idr < input +. ../../testutils.sh +idris2 TypeSearch.idr < input diff --git a/tests/idris2/interactive020/run b/tests/idris2/interactive020/run index 6f696478d0..a86589613d 100755 --- a/tests/idris2/interactive020/run +++ b/tests/idris2/interactive020/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-banner --no-color --console-width 0 Issue835.idr < input +. ../../testutils.sh +idris2 Issue835.idr < input diff --git a/tests/idris2/interactive021/run b/tests/idris2/interactive021/run index ae9de69d95..86c2a5d98b 100644 --- a/tests/idris2/interactive021/run +++ b/tests/idris2/interactive021/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner TypeAtDoNotation.idr < input +. ../../testutils.sh +idris2 TypeAtDoNotation.idr < input diff --git a/tests/idris2/interactive022/run b/tests/idris2/interactive022/run index afb9a56f7d..d928f12b7e 100644 --- a/tests/idris2/interactive022/run +++ b/tests/idris2/interactive022/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner TypeAtBangSyntax.idr < input +. ../../testutils.sh +idris2 TypeAtBangSyntax.idr < input diff --git a/tests/idris2/interactive023/run b/tests/idris2/interactive023/run index 06875e3ac1..ea627ff5fd 100644 --- a/tests/idris2/interactive023/run +++ b/tests/idris2/interactive023/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner TypeAtLambda.idr < input +. ../../testutils.sh +idris2 TypeAtLambda.idr < input diff --git a/tests/idris2/interactive024/run b/tests/idris2/interactive024/run index a372b32070..171f2a86e0 100644 --- a/tests/idris2/interactive024/run +++ b/tests/idris2/interactive024/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner TypeAtAsPatterns.idr < input +. ../../testutils.sh +idris2 TypeAtAsPatterns.idr < input diff --git a/tests/idris2/interactive025/run b/tests/idris2/interactive025/run index 7d28d868a8..c0bae1656e 100644 --- a/tests/idris2/interactive025/run +++ b/tests/idris2/interactive025/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner TypeAtInterfaces.idr < input +. ../../testutils.sh +idris2 TypeAtInterfaces.idr < input diff --git a/tests/idris2/interactive026/run b/tests/idris2/interactive026/run index 431c96b157..2f390b9fa7 100644 --- a/tests/idris2/interactive026/run +++ b/tests/idris2/interactive026/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner TypeAtRecords.idr < input +. ../../testutils.sh +idris2 TypeAtRecords.idr < input diff --git a/tests/idris2/interactive027/run b/tests/idris2/interactive027/run index fb9fcda04e..0d48d91c3d 100644 --- a/tests/idris2/interactive027/run +++ b/tests/idris2/interactive027/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner TypeAtLocalVars.idr < input +. ../../testutils.sh +idris2 TypeAtLocalVars.idr < input diff --git a/tests/idris2/interactive028/run b/tests/idris2/interactive028/run index 47008ffee7..fecfb4b107 100644 --- a/tests/idris2/interactive028/run +++ b/tests/idris2/interactive028/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner < input +. ../../testutils.sh +idris2 < input diff --git a/tests/idris2/interactive029/run b/tests/idris2/interactive029/run index 6d0eaccaef..a56b732f5f 100644 --- a/tests/idris2/interactive029/run +++ b/tests/idris2/interactive029/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner Issue834.idr < input +. ../../testutils.sh +idris2 Issue834.idr < input diff --git a/tests/idris2/interactive030/run b/tests/idris2/interactive030/run index 47008ffee7..fecfb4b107 100644 --- a/tests/idris2/interactive030/run +++ b/tests/idris2/interactive030/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner < input +. ../../testutils.sh +idris2 < input diff --git a/tests/idris2/interactive031/run b/tests/idris2/interactive031/run index b55b11be97..9242884368 100644 --- a/tests/idris2/interactive031/run +++ b/tests/idris2/interactive031/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner Signatures.idr < input +. ../../testutils.sh +idris2 Signatures.idr < input diff --git a/tests/idris2/interactive032/run b/tests/idris2/interactive032/run index 03f1054b6d..5a5252bddb 100755 --- a/tests/idris2/interactive032/run +++ b/tests/idris2/interactive032/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner Uninh.idr < input +. ../../testutils.sh +idris2 Uninh.idr < input diff --git a/tests/idris2/interactive033/run b/tests/idris2/interactive033/run index 656899a894..1dfbad70a1 100755 --- a/tests/idris2/interactive033/run +++ b/tests/idris2/interactive033/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner UninhIndent.idr < input +. ../../testutils.sh +idris2 UninhIndent.idr < input diff --git a/tests/idris2/interactive034/run b/tests/idris2/interactive034/run index f3aca2d4fa..7758f3e19e 100755 --- a/tests/idris2/interactive034/run +++ b/tests/idris2/interactive034/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner timeout.idr < input +. ../../testutils.sh +idris2 timeout.idr < input diff --git a/tests/idris2/interactive035/run b/tests/idris2/interactive035/run index 0f03f0c10f..654715325d 100755 --- a/tests/idris2/interactive035/run +++ b/tests/idris2/interactive035/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner unify.idr < input +. ../../testutils.sh +idris2 unify.idr < input diff --git a/tests/idris2/interactive036/run b/tests/idris2/interactive036/run index a267b92e6d..29c5f319c3 100755 --- a/tests/idris2/interactive036/run +++ b/tests/idris2/interactive036/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner casefn.idr < input +. ../../testutils.sh +idris2 casefn.idr < input diff --git a/tests/idris2/interactive037/run b/tests/idris2/interactive037/run index d82682cfbd..e33955b1a0 100755 --- a/tests/idris2/interactive037/run +++ b/tests/idris2/interactive037/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --no-banner Holes.idr < input +idris2 Holes.idr < input diff --git a/tests/idris2/interactive038/run b/tests/idris2/interactive038/run index a5983e604a..ec050c77b1 100755 --- a/tests/idris2/interactive038/run +++ b/tests/idris2/interactive038/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner IEdit.idr < input +. ../../testutils.sh +idris2 IEdit.idr < input diff --git a/tests/idris2/interactive039/run b/tests/idris2/interactive039/run index b74c809852..f0660f1109 100755 --- a/tests/idris2/interactive039/run +++ b/tests/idris2/interactive039/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --no-banner CS_Syntax.idr < input +idris2 CS_Syntax.idr < input diff --git a/tests/idris2/interactive040/run b/tests/idris2/interactive040/run index 30854d92ee..fecfb4b107 100755 --- a/tests/idris2/interactive040/run +++ b/tests/idris2/interactive040/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --no-banner < input +idris2 < input diff --git a/tests/idris2/interactive041/run b/tests/idris2/interactive041/run index 758f1ca8f0..db6d86a927 100755 --- a/tests/idris2/interactive041/run +++ b/tests/idris2/interactive041/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --no-banner Issue1741.idr < input +idris2 Issue1741.idr < input diff --git a/tests/idris2/interactive042/run b/tests/idris2/interactive042/run index fa6424e488..323393e8f4 100755 --- a/tests/idris2/interactive042/run +++ b/tests/idris2/interactive042/run @@ -1,6 +1,6 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --no-banner Issue35.idr < input -$1 --no-color --console-width 0 --no-banner --check Issue35-2.idr || true -$1 --no-color --console-width 0 --no-banner --show-machine-names --check Issue35-2.idr || true -$1 --no-color --console-width 0 --no-banner --show-namespaces --show-machine-names --check Issue35-2.idr || true \ No newline at end of file +idris2 Issue35.idr < input +check Issue35-2.idr || true +check --show-machine-names Issue35-2.idr || true +check --show-namespaces --show-machine-names Issue35-2.idr || true diff --git a/tests/idris2/interactive043/run b/tests/idris2/interactive043/run index 218413ab21..d2631795ff 100755 --- a/tests/idris2/interactive043/run +++ b/tests/idris2/interactive043/run @@ -1,10 +1,10 @@ -rm -rf build +. ../../testutils.sh rm -f ImplicitSplitsGen.idr cp ImplicitSplits.idr ImplicitSplitsGen.idr -$1 --no-color --console-width 0 --no-banner ImplicitSplitsGen.idr < input +idris2 ImplicitSplitsGen.idr < input # now that we've inline edited it, make sure it still type checks: -$1 --no-color --console-width 0 --no-banner --check ImplicitSplitsGen.idr +check ImplicitSplitsGen.idr rm ImplicitSplitsGen.idr diff --git a/tests/idris2/interactive044/run b/tests/idris2/interactive044/run index 9a0ae21f9a..71d5a34bd4 100755 --- a/tests/idris2/interactive044/run +++ b/tests/idris2/interactive044/run @@ -1,9 +1,9 @@ -rm -rf build +. ../../testutils.sh rm -f SplitShadowGen.idr cp SplitShadow.idr SplitShadowGen.idr sleep 1 -$1 --no-color --console-width 0 --no-banner SplitShadowGen.idr < input +idris2 SplitShadowGen.idr < input rm SplitShadowGen.idr diff --git a/tests/idris2/interactive045/run b/tests/idris2/interactive045/run index a5c8c12688..9041703d8a 100755 --- a/tests/idris2/interactive045/run +++ b/tests/idris2/interactive045/run @@ -1,9 +1,9 @@ -rm -rf build +. ../../testutils.sh rm -f Issue1742Gen.idr cp Issue1742.idr Issue1742Gen.idr sleep 1 -$1 --no-color --console-width 0 --no-banner Issue1742Gen.idr < input +idris2 Issue1742Gen.idr < input rm Issue1742Gen.idr diff --git a/tests/idris2/interactive046/run b/tests/idris2/interactive046/run index d90fe64772..dd6bd50c7f 100755 --- a/tests/idris2/interactive046/run +++ b/tests/idris2/interactive046/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner --check Issue2712.idr +. ../../testutils.sh +check Issue2712.idr diff --git a/tests/idris2/interface001/run b/tests/idris2/interface001/run index a806a6dde9..cca7bbd7b7 100755 --- a/tests/idris2/interface001/run +++ b/tests/idris2/interface001/run @@ -1,5 +1,4 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner --no-prelude IFace.idr < input -$1 --no-color --console-width 0 --no-banner --no-prelude IFace1.idr < input1 +. ../../testutils.sh +idris2 --no-prelude IFace.idr < input +idris2 --no-prelude IFace1.idr < input1 diff --git a/tests/idris2/interface002/run b/tests/idris2/interface002/run index 5563286fab..19265959fa 100755 --- a/tests/idris2/interface002/run +++ b/tests/idris2/interface002/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner --no-prelude Functor.idr < input +. ../../testutils.sh +idris2 --no-prelude Functor.idr < input diff --git a/tests/idris2/interface003/run b/tests/idris2/interface003/run index 618c69a3ff..070cba614b 100755 --- a/tests/idris2/interface003/run +++ b/tests/idris2/interface003/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner Do.idr < input +. ../../testutils.sh +idris2 Do.idr < input diff --git a/tests/idris2/interface004/run b/tests/idris2/interface004/run index 618c69a3ff..070cba614b 100755 --- a/tests/idris2/interface004/run +++ b/tests/idris2/interface004/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner Do.idr < input +. ../../testutils.sh +idris2 Do.idr < input diff --git a/tests/idris2/interface005/run b/tests/idris2/interface005/run index 409c08cdc0..5b4317229e 100755 --- a/tests/idris2/interface005/run +++ b/tests/idris2/interface005/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --check Deps.idr +. ../../testutils.sh +check Deps.idr diff --git a/tests/idris2/interface006/run b/tests/idris2/interface006/run index b58b8f8798..fe3c8cd87c 100755 --- a/tests/idris2/interface006/run +++ b/tests/idris2/interface006/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --check Bimonad.idr +. ../../testutils.sh +check Bimonad.idr diff --git a/tests/idris2/interface007/run b/tests/idris2/interface007/run index a0e29bcd4c..24086a6c87 100755 --- a/tests/idris2/interface007/run +++ b/tests/idris2/interface007/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --debug-elab-check --check A.idr +. ../../testutils.sh +check --debug-elab-check A.idr diff --git a/tests/idris2/interface008/run b/tests/idris2/interface008/run index 409c08cdc0..5b4317229e 100755 --- a/tests/idris2/interface008/run +++ b/tests/idris2/interface008/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --check Deps.idr +. ../../testutils.sh +check Deps.idr diff --git a/tests/idris2/interface009/run b/tests/idris2/interface009/run index 4109236655..1118c65bcc 100755 --- a/tests/idris2/interface009/run +++ b/tests/idris2/interface009/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner Odd.idr < input +. ../../testutils.sh +idris2 Odd.idr < input diff --git a/tests/idris2/interface010/run b/tests/idris2/interface010/run index fbe1c502eb..c7c671d744 100755 --- a/tests/idris2/interface010/run +++ b/tests/idris2/interface010/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 Dep.idr --check +. ../../testutils.sh +check Dep.idr diff --git a/tests/idris2/interface011/run b/tests/idris2/interface011/run index bc8bdeb63f..e4c61ef462 100755 --- a/tests/idris2/interface011/run +++ b/tests/idris2/interface011/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 FuncImpl.idr --check +. ../../testutils.sh +check FuncImpl.idr diff --git a/tests/idris2/interface012/run b/tests/idris2/interface012/run index 5a1e4f1ff8..82ae339542 100755 --- a/tests/idris2/interface012/run +++ b/tests/idris2/interface012/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 Defmeth.idr --check +. ../../testutils.sh +check Defmeth.idr diff --git a/tests/idris2/interface013/run b/tests/idris2/interface013/run index db4956b195..9bb60668eb 100755 --- a/tests/idris2/interface013/run +++ b/tests/idris2/interface013/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 TypeInt.idr --check +. ../../testutils.sh +check TypeInt.idr diff --git a/tests/idris2/interface014/run b/tests/idris2/interface014/run index 801e16c9e6..d12a274b1c 100755 --- a/tests/idris2/interface014/run +++ b/tests/idris2/interface014/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 DepInt.idr --check +. ../../testutils.sh +check DepInt.idr diff --git a/tests/idris2/interface015/run b/tests/idris2/interface015/run index e1dd75df0f..d1a25333b9 100755 --- a/tests/idris2/interface015/run +++ b/tests/idris2/interface015/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 gnu.idr --check +. ../../testutils.sh +check gnu.idr diff --git a/tests/idris2/interface016/run b/tests/idris2/interface016/run index 3e8ceebd64..25ba0b8b70 100755 --- a/tests/idris2/interface016/run +++ b/tests/idris2/interface016/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 TwoNum.idr --check +. ../../testutils.sh +check TwoNum.idr diff --git a/tests/idris2/interface017/run b/tests/idris2/interface017/run index 4ab7e153f9..fdc050c0bf 100755 --- a/tests/idris2/interface017/run +++ b/tests/idris2/interface017/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 Tricho.idr --check +. ../../testutils.sh +check Tricho.idr diff --git a/tests/idris2/interface018/run b/tests/idris2/interface018/run index ff9a43b9a7..f29f8f6faf 100755 --- a/tests/idris2/interface018/run +++ b/tests/idris2/interface018/run @@ -1,6 +1,5 @@ -rm -rf build - -$1 --no-color --no-banner --console-width 0 Sized.idr < input -$1 --no-color --console-width 0 Sized2.idr --check -$1 --no-color --console-width 0 Sized3.idr --check +. ../../testutils.sh +idris2 Sized.idr < input +check Sized2.idr +check Sized3.idr diff --git a/tests/idris2/interface019/run b/tests/idris2/interface019/run index ef3ceaa016..6d8e1e5605 100755 --- a/tests/idris2/interface019/run +++ b/tests/idris2/interface019/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 LocalHints.idr --check +. ../../testutils.sh +check LocalHints.idr diff --git a/tests/idris2/interface020/run b/tests/idris2/interface020/run index a808fb39b9..b6cf755f88 100755 --- a/tests/idris2/interface020/run +++ b/tests/idris2/interface020/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 LocalInterface.idr --check +. ../../testutils.sh +check LocalInterface.idr diff --git a/tests/idris2/interface021/run b/tests/idris2/interface021/run index daecd89cd4..a0611e6d32 100755 --- a/tests/idris2/interface021/run +++ b/tests/idris2/interface021/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 LocalHint.idr --check +. ../../testutils.sh +check LocalHint.idr diff --git a/tests/idris2/interface022/run b/tests/idris2/interface022/run index 72292dfe35..273dd0c211 100755 --- a/tests/idris2/interface022/run +++ b/tests/idris2/interface022/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --no-banner --console-width 0 DotMethod.idr < input +. ../../testutils.sh +idris2 DotMethod.idr < input diff --git a/tests/idris2/interface023/run b/tests/idris2/interface023/run index 459af48614..b8b0938790 100644 --- a/tests/idris2/interface023/run +++ b/tests/idris2/interface023/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-banner --no-color --console-width 0 AppComp.idr < input +. ../../testutils.sh +idris2 AppComp.idr < input diff --git a/tests/idris2/interface024/run b/tests/idris2/interface024/run index e4d3a1bd62..4db6d1aee3 100644 --- a/tests/idris2/interface024/run +++ b/tests/idris2/interface024/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-banner --no-color --check EH.idr -p contrib +. ../../testutils.sh +check EH.idr -p contrib diff --git a/tests/idris2/interface025/run b/tests/idris2/interface025/run index 77aa67d358..4493d93c10 100644 --- a/tests/idris2/interface025/run +++ b/tests/idris2/interface025/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-banner --no-color --console-width 0 AutoSearchHide2.idr < input +. ../../testutils.sh +idris2 AutoSearchHide2.idr < input diff --git a/tests/idris2/interface026/run b/tests/idris2/interface026/run index 2569bc5bfb..c3d09e3f6c 100644 --- a/tests/idris2/interface026/run +++ b/tests/idris2/interface026/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-banner --no-color --console-width 0 UninhabitedRec.idr < input +. ../../testutils.sh +idris2 UninhabitedRec.idr < input diff --git a/tests/idris2/interface027/run b/tests/idris2/interface027/run index 07f8afa558..adb525f312 100644 --- a/tests/idris2/interface027/run +++ b/tests/idris2/interface027/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-banner --no-color --console-width 0 params.idr < input +. ../../testutils.sh +idris2 params.idr < input diff --git a/tests/idris2/interface028/run b/tests/idris2/interface028/run index 92c558321c..cec8cd1b4b 100644 --- a/tests/idris2/interface028/run +++ b/tests/idris2/interface028/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-banner --no-color --console-width 0 --check InterfaceArgs.idr +check InterfaceArgs.idr diff --git a/tests/idris2/interface029/run b/tests/idris2/interface029/run index 9c7acf2f70..9c15937d55 100644 --- a/tests/idris2/interface029/run +++ b/tests/idris2/interface029/run @@ -1,5 +1,5 @@ -rm -rf build +. ../../testutils.sh -$1 --no-banner --no-color --console-width 0 -x main ForwardImpl1.idr -$1 --no-banner --no-color --console-width 0 --check ForwardImpl2.idr -$1 --no-banner --no-color --console-width 0 --check ForwardImpl3.idr +run ForwardImpl1.idr +check ForwardImpl2.idr +check ForwardImpl3.idr diff --git a/tests/idris2/interpolation001/run b/tests/idris2/interpolation001/run index 4e02e7eb75..bf0b168bef 100755 --- a/tests/idris2/interpolation001/run +++ b/tests/idris2/interpolation001/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner --check --alt-error-count 10 IfThenElse.idr +. ../../testutils.sh +check --alt-error-count 10 IfThenElse.idr diff --git a/tests/idris2/interpolation002/run b/tests/idris2/interpolation002/run index 30c1a0ac17..fa944260bc 100755 --- a/tests/idris2/interpolation002/run +++ b/tests/idris2/interpolation002/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner --exec test StringLiteral.idr +. ../../testutils.sh +idris2 --exec test StringLiteral.idr diff --git a/tests/idris2/interpolation003/run b/tests/idris2/interpolation003/run index d2c00b30c9..441b64ba30 100755 --- a/tests/idris2/interpolation003/run +++ b/tests/idris2/interpolation003/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner Test.idr < input +. ../../testutils.sh +idris2 Test.idr < input diff --git a/tests/idris2/interpolation004/run b/tests/idris2/interpolation004/run index 6e8254e391..b66578cdae 100755 --- a/tests/idris2/interpolation004/run +++ b/tests/idris2/interpolation004/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --no-banner --check StringLiteral.idr \ No newline at end of file +check StringLiteral.idr diff --git a/tests/idris2/interpreter001/run b/tests/idris2/interpreter001/run index d54b883a79..fecfb4b107 100755 --- a/tests/idris2/interpreter001/run +++ b/tests/idris2/interpreter001/run @@ -1 +1,3 @@ -$1 --no-color --console-width 0 --no-banner < input +. ../../testutils.sh + +idris2 < input diff --git a/tests/idris2/interpreter002/run b/tests/idris2/interpreter002/run index f34822def4..72182e4d4e 100755 --- a/tests/idris2/interpreter002/run +++ b/tests/idris2/interpreter002/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner --no-prelude < input +. ../../testutils.sh +idris2 --no-prelude < input diff --git a/tests/idris2/interpreter003/run b/tests/idris2/interpreter003/run index d54b883a79..fecfb4b107 100755 --- a/tests/idris2/interpreter003/run +++ b/tests/idris2/interpreter003/run @@ -1 +1,3 @@ -$1 --no-color --console-width 0 --no-banner < input +. ../../testutils.sh + +idris2 < input diff --git a/tests/idris2/interpreter004/run b/tests/idris2/interpreter004/run index d54b883a79..fecfb4b107 100755 --- a/tests/idris2/interpreter004/run +++ b/tests/idris2/interpreter004/run @@ -1 +1,3 @@ -$1 --no-color --console-width 0 --no-banner < input +. ../../testutils.sh + +idris2 < input diff --git a/tests/idris2/interpreter005/run b/tests/idris2/interpreter005/run index 80ed52e1c8..efb1e49621 100755 --- a/tests/idris2/interpreter005/run +++ b/tests/idris2/interpreter005/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-banner --no-color --console-width 0 Issue37.lidr < input +. ../../testutils.sh +idris2 Issue37.lidr < input diff --git a/tests/idris2/interpreter006/run b/tests/idris2/interpreter006/run index d54b883a79..fecfb4b107 100755 --- a/tests/idris2/interpreter006/run +++ b/tests/idris2/interpreter006/run @@ -1 +1,3 @@ -$1 --no-color --console-width 0 --no-banner < input +. ../../testutils.sh + +idris2 < input diff --git a/tests/idris2/interpreter007/run b/tests/idris2/interpreter007/run index d54b883a79..fecfb4b107 100755 --- a/tests/idris2/interpreter007/run +++ b/tests/idris2/interpreter007/run @@ -1 +1,3 @@ -$1 --no-color --console-width 0 --no-banner < input +. ../../testutils.sh + +idris2 < input diff --git a/tests/idris2/interpreter008/run b/tests/idris2/interpreter008/run index a10f5150ad..d2d8c793c9 100755 --- a/tests/idris2/interpreter008/run +++ b/tests/idris2/interpreter008/run @@ -1 +1,3 @@ -$1 --no-color --console-width 0 --no-banner Issue2041.idr < input +. ../../testutils.sh + +idris2 Issue2041.idr < input diff --git a/tests/idris2/lazy001/run b/tests/idris2/lazy001/run index cd63b62cb8..781b1d3526 100755 --- a/tests/idris2/lazy001/run +++ b/tests/idris2/lazy001/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner --no-prelude Lazy.idr < input +. ../../testutils.sh +idris2 --no-prelude Lazy.idr < input diff --git a/tests/idris2/lazy002/run b/tests/idris2/lazy002/run index 7e8f8f32d0..96bad64a01 100755 --- a/tests/idris2/lazy002/run +++ b/tests/idris2/lazy002/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner -p contrib LazyFoldlM.idr < input +. ../../testutils.sh +idris2 -p contrib LazyFoldlM.idr < input diff --git a/tests/idris2/lazy003/run b/tests/idris2/lazy003/run index 464bb5bab5..c33f96d18b 100644 --- a/tests/idris2/lazy003/run +++ b/tests/idris2/lazy003/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner DelayLam.idr < input +. ../../testutils.sh +idris2 DelayLam.idr < input diff --git a/tests/idris2/linear001/run b/tests/idris2/linear001/run index 27dea045ce..d2e5d9d81f 100755 --- a/tests/idris2/linear001/run +++ b/tests/idris2/linear001/run @@ -1,4 +1,3 @@ -rm -rf build - -echo ':q' | $1 --no-color --console-width 0 --no-banner --no-prelude Door.idr +. ../../testutils.sh +echo ':q' | idris2 --no-prelude Door.idr diff --git a/tests/idris2/linear002/run b/tests/idris2/linear002/run index 7af7dc6996..2708c0d900 100755 --- a/tests/idris2/linear002/run +++ b/tests/idris2/linear002/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner --no-prelude Door.idr < input +. ../../testutils.sh +idris2 --no-prelude Door.idr < input diff --git a/tests/idris2/linear003/run b/tests/idris2/linear003/run index 4ce1973ac4..f67775b529 100755 --- a/tests/idris2/linear003/run +++ b/tests/idris2/linear003/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner --no-prelude Linear.idr < input +. ../../testutils.sh +idris2 --no-prelude Linear.idr < input diff --git a/tests/idris2/linear004/run b/tests/idris2/linear004/run index 95b260a9d4..4f4a5a77bc 100755 --- a/tests/idris2/linear004/run +++ b/tests/idris2/linear004/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner --no-prelude Erase.idr < input +. ../../testutils.sh +idris2 --no-prelude Erase.idr < input diff --git a/tests/idris2/linear005/run b/tests/idris2/linear005/run index 770797118d..e6c63e9986 100755 --- a/tests/idris2/linear005/run +++ b/tests/idris2/linear005/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner Door.idr < input +. ../../testutils.sh +idris2 Door.idr < input diff --git a/tests/idris2/linear006/run b/tests/idris2/linear006/run index 526c8f15a1..29a7dcd48e 100755 --- a/tests/idris2/linear006/run +++ b/tests/idris2/linear006/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner ZFun.idr < input +. ../../testutils.sh +idris2 ZFun.idr < input diff --git a/tests/idris2/linear007/run b/tests/idris2/linear007/run index 411d28f244..a74a408b45 100755 --- a/tests/idris2/linear007/run +++ b/tests/idris2/linear007/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --check LCase.idr +. ../../testutils.sh +check LCase.idr diff --git a/tests/idris2/linear008/run b/tests/idris2/linear008/run index 928911b9e4..8f9c6d68cf 100644 --- a/tests/idris2/linear008/run +++ b/tests/idris2/linear008/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --check Door.idr +. ../../testutils.sh +check Door.idr diff --git a/tests/idris2/linear009/run b/tests/idris2/linear009/run index f961f713ac..68e408f2ff 100644 --- a/tests/idris2/linear009/run +++ b/tests/idris2/linear009/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner qtt.idr < input +. ../../testutils.sh +idris2 qtt.idr < input diff --git a/tests/idris2/linear010/run b/tests/idris2/linear010/run index 7fe618cb20..683ee35b90 100644 --- a/tests/idris2/linear010/run +++ b/tests/idris2/linear010/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --check Door.idr -p linear +check Door.idr -p linear diff --git a/tests/idris2/linear011/run b/tests/idris2/linear011/run index e12a2a7c51..c2f4ab208a 100644 --- a/tests/idris2/linear011/run +++ b/tests/idris2/linear011/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --no-banner Network.idr -p linear -p network < input +idris2 Network.idr -p linear -p network < input diff --git a/tests/idris2/linear012/run b/tests/idris2/linear012/run index 5cb88637c0..a4b65dd45f 100644 --- a/tests/idris2/linear012/run +++ b/tests/idris2/linear012/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner linholes.idr < input +. ../../testutils.sh +idris2 linholes.idr < input diff --git a/tests/idris2/linear013/run b/tests/idris2/linear013/run index 49f1857c2d..b5082d31ac 100644 --- a/tests/idris2/linear013/run +++ b/tests/idris2/linear013/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner Issue758.idr < input +. ../../testutils.sh +idris2 Issue758.idr < input diff --git a/tests/idris2/linear014/run b/tests/idris2/linear014/run index f4f3329fb4..18b94f9894 100644 --- a/tests/idris2/linear014/run +++ b/tests/idris2/linear014/run @@ -1,3 +1,3 @@ -$1 --no-color --console-width 0 --no-banner --check Issue55.idr +. ../../testutils.sh -rm -r build +check Issue55.idr diff --git a/tests/idris2/linear015/run b/tests/idris2/linear015/run index f2df502737..b27c98a40c 100644 --- a/tests/idris2/linear015/run +++ b/tests/idris2/linear015/run @@ -1,3 +1,3 @@ -$1 --no-color --console-width 0 --no-banner --check Issue1861.idr +. ../../testutils.sh -rm -r build \ No newline at end of file +check Issue1861.idr diff --git a/tests/idris2/linear016/run b/tests/idris2/linear016/run index 9096b82189..2bf190b2d0 100644 --- a/tests/idris2/linear016/run +++ b/tests/idris2/linear016/run @@ -1,4 +1,4 @@ -$1 --no-color --console-width 0 --no-banner --check Issue2895.idr -$1 --no-color --console-width 0 --no-banner --check Issue2895_2.idr +. ../../testutils.sh -rm -r build +check Issue2895.idr +check Issue2895_2.idr diff --git a/tests/idris2/linear017/run b/tests/idris2/linear017/run index 4b4d889dd2..e5c3405335 100644 --- a/tests/idris2/linear017/run +++ b/tests/idris2/linear017/run @@ -1,3 +1,3 @@ -$1 --no-color --console-width 0 --no-banner --check As.idr +. ../../testutils.sh -rm -r build +check As.idr diff --git a/tests/idris2/literate001/run b/tests/idris2/literate001/run index 1157580b9a..b2447bcbe5 100755 --- a/tests/idris2/literate001/run +++ b/tests/idris2/literate001/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner IEdit.lidr < input +. ../../testutils.sh +idris2 IEdit.lidr < input diff --git a/tests/idris2/literate002/run b/tests/idris2/literate002/run index f4553c00a3..8da12871cc 100755 --- a/tests/idris2/literate002/run +++ b/tests/idris2/literate002/run @@ -1,5 +1,4 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner IEdit.lidr < input -$1 --no-color --console-width 0 --no-banner IEdit2.lidr < input2 +. ../../testutils.sh +idris2 IEdit.lidr < input +idris2 IEdit2.lidr < input2 diff --git a/tests/idris2/literate003/run b/tests/idris2/literate003/run index 1157580b9a..b2447bcbe5 100755 --- a/tests/idris2/literate003/run +++ b/tests/idris2/literate003/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner IEdit.lidr < input +. ../../testutils.sh +idris2 IEdit.lidr < input diff --git a/tests/idris2/literate004/run b/tests/idris2/literate004/run index 1157580b9a..b2447bcbe5 100755 --- a/tests/idris2/literate004/run +++ b/tests/idris2/literate004/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner IEdit.lidr < input +. ../../testutils.sh +idris2 IEdit.lidr < input diff --git a/tests/idris2/literate005/run b/tests/idris2/literate005/run index 1157580b9a..b2447bcbe5 100755 --- a/tests/idris2/literate005/run +++ b/tests/idris2/literate005/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner IEdit.lidr < input +. ../../testutils.sh +idris2 IEdit.lidr < input diff --git a/tests/idris2/literate006/run b/tests/idris2/literate006/run index 753346bbf8..25014df805 100755 --- a/tests/idris2/literate006/run +++ b/tests/idris2/literate006/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner Door.lidr < input +. ../../testutils.sh +idris2 Door.lidr < input diff --git a/tests/idris2/literate007/run b/tests/idris2/literate007/run index 242ed5b5c8..728d112866 100755 --- a/tests/idris2/literate007/run +++ b/tests/idris2/literate007/run @@ -1,5 +1,4 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner IEdit.lidr < input -$1 --no-color --console-width 0 --no-banner IEditOrg.org < input2 +. ../../testutils.sh +idris2 IEdit.lidr < input +idris2 IEditOrg.org < input2 diff --git a/tests/idris2/literate008/run b/tests/idris2/literate008/run index 1157580b9a..b2447bcbe5 100755 --- a/tests/idris2/literate008/run +++ b/tests/idris2/literate008/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner IEdit.lidr < input +. ../../testutils.sh +idris2 IEdit.lidr < input diff --git a/tests/idris2/literate009/run b/tests/idris2/literate009/run index 30d3987b2f..988f2bef06 100755 --- a/tests/idris2/literate009/run +++ b/tests/idris2/literate009/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner WithLift.lidr < input +. ../../testutils.sh +idris2 WithLift.lidr < input diff --git a/tests/idris2/literate010/run b/tests/idris2/literate010/run index 2eb4783fa3..ef39e14ed1 100755 --- a/tests/idris2/literate010/run +++ b/tests/idris2/literate010/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --check --no-banner MyFirstIdrisProgram.org +. ../../testutils.sh +check MyFirstIdrisProgram.org diff --git a/tests/idris2/literate011/run b/tests/idris2/literate011/run index ec6152229e..4c62db0e31 100755 --- a/tests/idris2/literate011/run +++ b/tests/idris2/literate011/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner IEdit.md < input +. ../../testutils.sh +idris2 IEdit.md < input diff --git a/tests/idris2/literate012/run b/tests/idris2/literate012/run index a5104b95b9..5a24d321f9 100755 --- a/tests/idris2/literate012/run +++ b/tests/idris2/literate012/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner IEdit.org < input +. ../../testutils.sh +idris2 IEdit.org < input diff --git a/tests/idris2/literate013/run b/tests/idris2/literate013/run index 829bc5c3f8..7106e0f5d7 100755 --- a/tests/idris2/literate013/run +++ b/tests/idris2/literate013/run @@ -1,4 +1,4 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --no-banner --check Lit.lidr -$1 --no-color --console-width 0 --no-banner --check LitTeX.tex +check Lit.lidr +check LitTeX.tex diff --git a/tests/idris2/literate014/run b/tests/idris2/literate014/run index 347151fded..685423ab3f 100755 --- a/tests/idris2/literate014/run +++ b/tests/idris2/literate014/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner with.lidr < input +. ../../testutils.sh +idris2 with.lidr < input diff --git a/tests/idris2/literate015/run b/tests/idris2/literate015/run index 542885797e..55a453de6d 100755 --- a/tests/idris2/literate015/run +++ b/tests/idris2/literate015/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner case.lidr < input +. ../../testutils.sh +idris2 case.lidr < input diff --git a/tests/idris2/literate016/run b/tests/idris2/literate016/run index a5104b95b9..5a24d321f9 100755 --- a/tests/idris2/literate016/run +++ b/tests/idris2/literate016/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner IEdit.org < input +. ../../testutils.sh +idris2 IEdit.org < input diff --git a/tests/idris2/literate017/run b/tests/idris2/literate017/run index 9679c639e2..55708b1e07 100755 --- a/tests/idris2/literate017/run +++ b/tests/idris2/literate017/run @@ -1,4 +1,6 @@ -rm -rf build project-output.ipkg +. ../../testutils.sh -$1 --init project-output.ipkg < input -diff project-output.ipkg project-expected.ipkg \ No newline at end of file +rm project-output.ipkg + +idris2 --init project-output.ipkg < input +diff project-output.ipkg project-expected.ipkg diff --git a/tests/idris2/namespace001/run b/tests/idris2/namespace001/run index 3c80c945cd..e6d0e3e284 100644 --- a/tests/idris2/namespace001/run +++ b/tests/idris2/namespace001/run @@ -1,5 +1,4 @@ -rm -rf build - -echo ':q' | $1 --no-color --console-width 0 --no-banner Dup.idr -echo ':t Test' | $1 --no-color --console-width 0 --no-banner Scope.idr +. ../../testutils.sh +echo ':q' | idris2 Dup.idr +echo ':t Test' | idris2 Scope.idr diff --git a/tests/idris2/namespace002/run b/tests/idris2/namespace002/run index 3de7ec7937..fdffd3d555 100644 --- a/tests/idris2/namespace002/run +++ b/tests/idris2/namespace002/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --check Issue1313.idr || true +check Issue1313.idr || true diff --git a/tests/idris2/namespace003/run b/tests/idris2/namespace003/run index b1d4093bef..248924d6b2 100644 --- a/tests/idris2/namespace003/run +++ b/tests/idris2/namespace003/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --check Test.idr || true +check Test.idr || true diff --git a/tests/idris2/namespace004/run b/tests/idris2/namespace004/run index 1b3ee1c87b..3e447e08b5 100644 --- a/tests/idris2/namespace004/run +++ b/tests/idris2/namespace004/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --check Export.idr \ No newline at end of file +check Export.idr diff --git a/tests/idris2/namespace005/run b/tests/idris2/namespace005/run index 390bbccc40..5f4086519a 100755 --- a/tests/idris2/namespace005/run +++ b/tests/idris2/namespace005/run @@ -1,20 +1,18 @@ -rm -rf ./build +. ../../testutils.sh +check Main0.idr +check Main1.idr +check MainFail.idr +check Main3.idr -$1 --no-color --console-width 0 --check Main0.idr -$1 --no-color --console-width 0 --check Main1.idr -$1 --no-color --console-width 0 --check MainFail.idr -$1 --no-color --console-width 0 --check Main3.idr +run Main0.idr +run Main1.idr -$1 --no-color --console-width 0 --exec main Main0.idr -$1 --no-color --console-width 0 --exec main Main1.idr +check MainPre0.idr +check MainPre1.idr -$1 --no-color --console-width 0 --check MainPre0.idr -$1 --no-color --console-width 0 --check MainPre1.idr - -$1 --no-color --console-width 0 --exec main MainPre0.idr -$1 --no-color --console-width 0 --exec main MainPre1.idr - -$1 --no-color --console-width 0 --check MainConflict.idr -$1 --no-color --console-width 0 --exec exec MainConflict.idr +run MainPre0.idr +run MainPre1.idr +check MainConflict.idr +idris2 --exec exec MainConflict.idr diff --git a/tests/idris2/params001/run b/tests/idris2/params001/run index 522748355f..15a9a0a8db 100755 --- a/tests/idris2/params001/run +++ b/tests/idris2/params001/run @@ -1,5 +1,4 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner --check param.idr -$1 --no-color --console-width 0 --no-banner --check parambad.idr +. ../../testutils.sh +check param.idr +check parambad.idr diff --git a/tests/idris2/params002/run b/tests/idris2/params002/run index b042792654..1984e14548 100755 --- a/tests/idris2/params002/run +++ b/tests/idris2/params002/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner ParamsPrint.idr < input +. ../../testutils.sh +idris2 ParamsPrint.idr < input diff --git a/tests/idris2/params003/run b/tests/idris2/params003/run index bf83222dfb..8d17d8e6ce 100755 --- a/tests/idris2/params003/run +++ b/tests/idris2/params003/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner casesplit.idr < input +. ../../testutils.sh +idris2 casesplit.idr < input diff --git a/tests/idris2/perf001/run b/tests/idris2/perf001/run index e328ebd8dc..f9651336dd 100755 --- a/tests/idris2/perf001/run +++ b/tests/idris2/perf001/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --check Big.idr +. ../../testutils.sh +check Big.idr diff --git a/tests/idris2/perf002/run b/tests/idris2/perf002/run index e328ebd8dc..f9651336dd 100755 --- a/tests/idris2/perf002/run +++ b/tests/idris2/perf002/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --check Big.idr +. ../../testutils.sh +check Big.idr diff --git a/tests/idris2/perf003/run b/tests/idris2/perf003/run index 4990086762..2398c31034 100755 --- a/tests/idris2/perf003/run +++ b/tests/idris2/perf003/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --check Auto.idr +. ../../testutils.sh +check Auto.idr diff --git a/tests/idris2/perf004/run b/tests/idris2/perf004/run index 71fd0c56f2..2af4edcc4e 100755 --- a/tests/idris2/perf004/run +++ b/tests/idris2/perf004/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --check bigdpair.idr +. ../../testutils.sh +check bigdpair.idr diff --git a/tests/idris2/perf005/run b/tests/idris2/perf005/run index 67bfd9e533..1bc7122df1 100755 --- a/tests/idris2/perf005/run +++ b/tests/idris2/perf005/run @@ -1,10 +1,9 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --check Lambda.idr +check Lambda.idr -$1 --no-color --console-width 0 --check Bad1.idr +check Bad1.idr -$1 --no-color --console-width 0 --check Bad2.idr - -$1 --no-color --console-width 0 --check Bad3.idr +check Bad2.idr +check Bad3.idr diff --git a/tests/idris2/perf007/run b/tests/idris2/perf007/run index 63c6f890d0..78ac337031 100644 --- a/tests/idris2/perf007/run +++ b/tests/idris2/perf007/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --check Slooow.idr +. ../../testutils.sh +check Slooow.idr diff --git a/tests/idris2/perf008/run b/tests/idris2/perf008/run index 7a0729f981..3f89ed77ae 100644 --- a/tests/idris2/perf008/run +++ b/tests/idris2/perf008/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 FinPerf.idr --exec main +. ../../testutils.sh +run FinPerf.idr diff --git a/tests/idris2/perf009/run b/tests/idris2/perf009/run index 605ef5cffe..4a6e58ce7f 100644 --- a/tests/idris2/perf009/run +++ b/tests/idris2/perf009/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --check C.idr +check C.idr diff --git a/tests/idris2/perf010/run b/tests/idris2/perf010/run index 0235fd3a0d..8d7dc963a0 100644 --- a/tests/idris2/perf010/run +++ b/tests/idris2/perf010/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --check Printf.idr +check Printf.idr diff --git a/tests/idris2/perf011/run b/tests/idris2/perf011/run index 605ef5cffe..4a6e58ce7f 100644 --- a/tests/idris2/perf011/run +++ b/tests/idris2/perf011/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --check C.idr +check C.idr diff --git a/tests/idris2/perf012/run b/tests/idris2/perf012/run index 522a6cd1a3..6fb7ad593f 100644 --- a/tests/idris2/perf012/run +++ b/tests/idris2/perf012/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 Main.idr -o main +idris2 Main.idr -o main diff --git a/tests/idris2/perf2202/run b/tests/idris2/perf2202/run index 31891ade86..f0e81801b6 100755 --- a/tests/idris2/perf2202/run +++ b/tests/idris2/perf2202/run @@ -1,4 +1,4 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --check many10000.idr -$1 --no-color --console-width 0 --check Church.idr \ No newline at end of file +check many10000.idr +check Church.idr diff --git a/tests/idris2/perror001/run b/tests/idris2/perror001/run index 3412449191..9af9ec0f8e 100755 --- a/tests/idris2/perror001/run +++ b/tests/idris2/perror001/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --check PError.idr +. ../../testutils.sh +check PError.idr diff --git a/tests/idris2/perror002/run b/tests/idris2/perror002/run index 3412449191..9af9ec0f8e 100755 --- a/tests/idris2/perror002/run +++ b/tests/idris2/perror002/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --check PError.idr +. ../../testutils.sh +check PError.idr diff --git a/tests/idris2/perror003/run b/tests/idris2/perror003/run index 500fb90527..222a9b1b14 100755 --- a/tests/idris2/perror003/run +++ b/tests/idris2/perror003/run @@ -1,5 +1,4 @@ -rm -rf build - -$1 --no-color --console-width 0 --check PError.idr -$1 --no-color --console-width 0 --check PError2.idr +. ../../testutils.sh +check PError.idr +check PError2.idr diff --git a/tests/idris2/perror004/run b/tests/idris2/perror004/run index 3412449191..9af9ec0f8e 100755 --- a/tests/idris2/perror004/run +++ b/tests/idris2/perror004/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --check PError.idr +. ../../testutils.sh +check PError.idr diff --git a/tests/idris2/perror005/run b/tests/idris2/perror005/run index 3412449191..9af9ec0f8e 100755 --- a/tests/idris2/perror005/run +++ b/tests/idris2/perror005/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --check PError.idr +. ../../testutils.sh +check PError.idr diff --git a/tests/idris2/perror006/run b/tests/idris2/perror006/run index 3412449191..9af9ec0f8e 100755 --- a/tests/idris2/perror006/run +++ b/tests/idris2/perror006/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --check PError.idr +. ../../testutils.sh +check PError.idr diff --git a/tests/idris2/perror007/run b/tests/idris2/perror007/run index 951c2ccd1a..9b41cb8191 100755 --- a/tests/idris2/perror007/run +++ b/tests/idris2/perror007/run @@ -1,15 +1,14 @@ -rm -rf build - -$1 --no-color --console-width 0 --check StrError1.idr || true -$1 --no-color --console-width 0 --check StrError2.idr || true -$1 --no-color --console-width 0 --check StrError3.idr || true -$1 --no-color --console-width 0 --check StrError4.idr || true -$1 --no-color --console-width 0 --check StrError5.idr || true -$1 --no-color --console-width 0 --check StrError6.idr || true -$1 --no-color --console-width 0 --check StrError7.idr || true -$1 --no-color --console-width 0 --check StrError8.idr || true -$1 --no-color --console-width 0 --check StrError9.idr || true -$1 --no-color --console-width 0 --check StrError10.idr || true -$1 --no-color --console-width 0 --check StrError11.idr || true -$1 --no-color --console-width 0 --check StrError12.idr || true +. ../../testutils.sh +check StrError1.idr || true +check StrError2.idr || true +check StrError3.idr || true +check StrError4.idr || true +check StrError5.idr || true +check StrError6.idr || true +check StrError7.idr || true +check StrError8.idr || true +check StrError9.idr || true +check StrError10.idr || true +check StrError11.idr || true +check StrError12.idr || true diff --git a/tests/idris2/perror008/run b/tests/idris2/perror008/run index 9464c8f203..86453dab6f 100755 --- a/tests/idris2/perror008/run +++ b/tests/idris2/perror008/run @@ -1,12 +1,11 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --check Issue710a.idr || true -$1 --no-color --console-width 0 --check Issue710b.idr || true -$1 --no-color --console-width 0 --check Issue710c.idr || true -$1 --no-color --console-width 0 --check Issue710d.idr || true -$1 --no-color --console-width 0 --check Issue710e.idr || true -$1 --no-color --console-width 0 --check Issue710f.idr || true - -$1 --no-color --console-width 0 --check Issue1224a.idr || true -$1 --no-color --console-width 0 --check Issue1224b.idr || true +check Issue710a.idr || true +check Issue710b.idr || true +check Issue710c.idr || true +check Issue710d.idr || true +check Issue710e.idr || true +check Issue710f.idr || true +check Issue1224a.idr || true +check Issue1224b.idr || true diff --git a/tests/idris2/perror009/run b/tests/idris2/perror009/run index 045b3b74f6..99b7a73375 100644 --- a/tests/idris2/perror009/run +++ b/tests/idris2/perror009/run @@ -1,3 +1,3 @@ -$1 --no-color --console-width 0 --no-banner -Werror --check Error1.idr +. ../../testutils.sh -rm -rf build/ +check -Werror Error1.idr diff --git a/tests/idris2/perror010/run b/tests/idris2/perror010/run index 6002160ea7..a8b988361a 100755 --- a/tests/idris2/perror010/run +++ b/tests/idris2/perror010/run @@ -1,7 +1,7 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --check NamedReturn1.idr || true -$1 --no-color --console-width 0 --check NamedReturn2.idr || true -$1 --no-color --console-width 0 --check NamedReturn3.idr || true -$1 --no-color --console-width 0 --check NamedReturn4.idr || true -$1 --no-color --console-width 0 --no-banner TrailingLam.idr < input +check NamedReturn1.idr || true +check NamedReturn2.idr || true +check NamedReturn3.idr || true +check NamedReturn4.idr || true +idris2 TrailingLam.idr < input diff --git a/tests/idris2/perror011/run b/tests/idris2/perror011/run index de244844b6..7d65c2dd7d 100755 --- a/tests/idris2/perror011/run +++ b/tests/idris2/perror011/run @@ -1,6 +1,6 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --check Issue1345.idr || true -$1 --no-color --console-width 0 --check Issue1496-1.idr || true -$1 --no-color --console-width 0 --check Issue1496-2.idr || true -$1 --build foo.ipkg \ No newline at end of file +check Issue1345.idr || true +check Issue1496-1.idr || true +check Issue1496-2.idr || true +idris2 --build foo.ipkg diff --git a/tests/idris2/perror012/run b/tests/idris2/perror012/run index 149945fcca..3ff2996284 100755 --- a/tests/idris2/perror012/run +++ b/tests/idris2/perror012/run @@ -1,4 +1,4 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --check LamParseError.idr -$1 --no-color --console-width 0 --check CaseParseError.idr +check LamParseError.idr +check CaseParseError.idr diff --git a/tests/idris2/perror013/run b/tests/idris2/perror013/run index 23367a69a9..a0f3d3abdc 100755 --- a/tests/idris2/perror013/run +++ b/tests/idris2/perror013/run @@ -1,6 +1,6 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --check EmptyFailing.idr || true -$1 --no-color --console-width 0 --check EmptyMutual.idr || true -$1 --no-color --console-width 0 --check EmptyUsing.idr || true -$1 --no-color --console-width 0 --check EmptyParameters.idr || true +check EmptyFailing.idr || true +check EmptyMutual.idr || true +check EmptyUsing.idr || true +check EmptyParameters.idr || true diff --git a/tests/idris2/perror014/run b/tests/idris2/perror014/run index e4e73aca04..cb78507392 100755 --- a/tests/idris2/perror014/run +++ b/tests/idris2/perror014/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --check ParseList.idr || true +check ParseList.idr || true diff --git a/tests/idris2/perror015/run b/tests/idris2/perror015/run index 41f17f08e9..9588fd16a7 100755 --- a/tests/idris2/perror015/run +++ b/tests/idris2/perror015/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --check ParseWith.idr || true +check ParseWith.idr || true diff --git a/tests/idris2/perror016/run b/tests/idris2/perror016/run index 49a5c257b4..c35719fae6 100755 --- a/tests/idris2/perror016/run +++ b/tests/idris2/perror016/run @@ -1,6 +1,6 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --check ParseIf.idr || true -$1 --no-color --console-width 0 --check ParseIf2.idr || true -$1 --no-color --console-width 0 --check ParseIf3.idr || true -$1 --no-color --console-width 0 --check ParseIf4.idr || true +check ParseIf.idr || true +check ParseIf2.idr || true +check ParseIf3.idr || true +check ParseIf4.idr || true diff --git a/tests/idris2/perror017/run b/tests/idris2/perror017/run index 0c6ddc979e..03b6a7cb94 100755 --- a/tests/idris2/perror017/run +++ b/tests/idris2/perror017/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --check ParseImpl.idr || true +check ParseImpl.idr || true diff --git a/tests/idris2/perror018/run b/tests/idris2/perror018/run index 6f5dd6ab3b..42e0ccfebb 100755 --- a/tests/idris2/perror018/run +++ b/tests/idris2/perror018/run @@ -1,5 +1,5 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --check ParseRecord.idr || true -$1 --no-color --console-width 0 --check ParseRecord2.idr || true -$1 --no-color --console-width 0 --check ParseRecord3.idr || true +check ParseRecord.idr || true +check ParseRecord2.idr || true +check ParseRecord3.idr || true diff --git a/tests/idris2/perror019/run b/tests/idris2/perror019/run index 7f11eb841d..e91b494540 100644 --- a/tests/idris2/perror019/run +++ b/tests/idris2/perror019/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-banner --no-color --console-width 0 --check ImplError.idr +check ImplError.idr diff --git a/tests/idris2/perror020/run b/tests/idris2/perror020/run index ae477a872f..0847511ad1 100644 --- a/tests/idris2/perror020/run +++ b/tests/idris2/perror020/run @@ -1,4 +1,4 @@ -rm -rf build +. ../../testutils.sh -$1 --no-banner --no-color --console-width 0 --check Issue2769.idr -$1 --no-banner --no-color --console-width 0 --check Issue2769b.idr +check Issue2769.idr +check Issue2769b.idr diff --git a/tests/idris2/perror021/run b/tests/idris2/perror021/run index 05dc66508e..096de90fd3 100644 --- a/tests/idris2/perror021/run +++ b/tests/idris2/perror021/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-banner --no-color --console-width 0 --check Implicit.idr +. ../../testutils.sh +check Implicit.idr diff --git a/tests/idris2/perror022/run b/tests/idris2/perror022/run index 366aef3225..2e9711032e 100644 --- a/tests/idris2/perror022/run +++ b/tests/idris2/perror022/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-banner --no-color --console-width 0 --check Indent.idr +. ../../testutils.sh +check Indent.idr diff --git a/tests/idris2/perror023/run b/tests/idris2/perror023/run index 35820caef3..a5fb5a99c7 100755 --- a/tests/idris2/perror023/run +++ b/tests/idris2/perror023/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --check ParseError.idr +check ParseError.idr diff --git a/tests/idris2/perror024/run b/tests/idris2/perror024/run index 35820caef3..a5fb5a99c7 100755 --- a/tests/idris2/perror024/run +++ b/tests/idris2/perror024/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --check ParseError.idr +check ParseError.idr diff --git a/tests/idris2/perror025/run b/tests/idris2/perror025/run index 0a14ca1f39..6c3df8ceed 100755 --- a/tests/idris2/perror025/run +++ b/tests/idris2/perror025/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --check DataWhere.idr +check DataWhere.idr diff --git a/tests/idris2/perror026/run b/tests/idris2/perror026/run index c30f5e98d1..ea25311502 100755 --- a/tests/idris2/perror026/run +++ b/tests/idris2/perror026/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --check Micro.idr +check Micro.idr diff --git a/tests/idris2/perror027/run b/tests/idris2/perror027/run index a8e90ec01e..422d2c27c5 100755 --- a/tests/idris2/perror027/run +++ b/tests/idris2/perror027/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --check Outdent.idr +check Outdent.idr diff --git a/tests/idris2/perror028/run b/tests/idris2/perror028/run index 692ea278f5..a9617307b4 100644 --- a/tests/idris2/perror028/run +++ b/tests/idris2/perror028/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --check LetInDo.idr \ No newline at end of file +check LetInDo.idr diff --git a/tests/idris2/perror029/run b/tests/idris2/perror029/run index de74a0feea..36ad497cca 100644 --- a/tests/idris2/perror029/run +++ b/tests/idris2/perror029/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --check DelayParse.idr +check DelayParse.idr diff --git a/tests/idris2/pkg001/run b/tests/idris2/pkg001/run index fa9604f3f3..3d99b32b7c 100755 --- a/tests/idris2/pkg001/run +++ b/tests/idris2/pkg001/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --build dummy.ipkg +. ../../testutils.sh +idris2 --build dummy.ipkg diff --git a/tests/idris2/pkg002/run b/tests/idris2/pkg002/run index 74a38aba0b..3cb8823af2 100755 --- a/tests/idris2/pkg002/run +++ b/tests/idris2/pkg002/run @@ -1,6 +1,5 @@ -rm -rf build +. ../../testutils.sh cd src/Top -$1 --find-ipkg --check Dummy.idr +check --find-ipkg Dummy.idr cd ../.. - diff --git a/tests/idris2/pkg003/run b/tests/idris2/pkg003/run index 7222684746..61526aceb7 100755 --- a/tests/idris2/pkg003/run +++ b/tests/idris2/pkg003/run @@ -1,9 +1,11 @@ -$1 --build testpkg.ipkg +. ../../testutils.sh + +idris2 --build testpkg.ipkg rm -rf build/ -$1 --build testpkg.ipkg --quiet -$1 --build testpkg.ipkg --verbose -$1 --build testpkg.ipkg --codegen gambit -$1 --build testpkg.ipkg --ide-mode -$1 --build malformed-package-name -$1 --build non-existent-package.ipkg +idris2 --build testpkg.ipkg --quiet +idris2 --build testpkg.ipkg --verbose +idris2 --build testpkg.ipkg --codegen gambit +idris2 --build testpkg.ipkg --ide-mode +idris2 --build malformed-package-name +idris2 --build non-existent-package.ipkg rm -rf build/ diff --git a/tests/idris2/pkg004/run b/tests/idris2/pkg004/run index f27e74a2d4..b32a1a5bc6 100755 --- a/tests/idris2/pkg004/run +++ b/tests/idris2/pkg004/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --repl dummy.ipkg < input +. ../../testutils.sh +idris2 --repl dummy.ipkg < input diff --git a/tests/idris2/pkg005/run b/tests/idris2/pkg005/run index 0b6a055a91..1e1df3c0b9 100755 --- a/tests/idris2/pkg005/run +++ b/tests/idris2/pkg005/run @@ -1,6 +1,5 @@ -rm -rf build - -$1 --clean foo.ipkg -$1 --repl foo.ipkg < input -$1 --repl foo.ipkg < input +. ../../testutils.sh +idris2 --clean foo.ipkg +idris2 --repl foo.ipkg < input +idris2 --repl foo.ipkg < input diff --git a/tests/idris2/pkg006/run b/tests/idris2/pkg006/run index 414137e6d2..3c62534d24 100755 --- a/tests/idris2/pkg006/run +++ b/tests/idris2/pkg006/run @@ -1,8 +1,7 @@ -rm -rf build - -$1 --no-color --build test1.ipkg -$1 --no-color --build test2.ipkg -$1 --no-color --build test3.ipkg -$1 --no-color --build test4.ipkg -$1 --no-color --build test5.ipkg +. ../../testutils.sh +idris2 --build test1.ipkg +idris2 --build test2.ipkg +idris2 --build test3.ipkg +idris2 --build test4.ipkg +idris2 --build test5.ipkg diff --git a/tests/idris2/pkg007/run b/tests/idris2/pkg007/run index 40c33a9da3..834d0f0ca7 100755 --- a/tests/idris2/pkg007/run +++ b/tests/idris2/pkg007/run @@ -1,9 +1,11 @@ -$1 --init < input -$1 --build cool.ipkg +. ../../testutils.sh + +idris2 --init < input +idris2 --build cool.ipkg rm -Rf build/ -$1 --init cool2.ipkg < input2 -$1 --build cool2.ipkg +idris2 --init cool2.ipkg < input2 +idris2 --build cool2.ipkg rm -Rf build/ rm -f cool.ipkg cool2.ipkg diff --git a/tests/idris2/pkg008/run b/tests/idris2/pkg008/run index bce7d6afd9..652be410e4 100755 --- a/tests/idris2/pkg008/run +++ b/tests/idris2/pkg008/run @@ -1,2 +1,3 @@ -$1 --build foo.ipkg --build bar.ipkg -rm -rf build/ +. ../../testutils.sh + +idris2 --build foo.ipkg --build bar.ipkg diff --git a/tests/idris2/pkg009/run b/tests/idris2/pkg009/run index 50ddb019d7..83fb782eeb 100644 --- a/tests/idris2/pkg009/run +++ b/tests/idris2/pkg009/run @@ -1,2 +1,4 @@ -$1 --build testpkg/testpkg.ipkg +. ../../testutils.sh + +idris2 --build testpkg/testpkg.ipkg rm -rf testpkg/build/ diff --git a/tests/idris2/pkg010/run b/tests/idris2/pkg010/run index ebaa477314..32df9792a3 100755 --- a/tests/idris2/pkg010/run +++ b/tests/idris2/pkg010/run @@ -1,3 +1,5 @@ +. ../../testutils.sh + if [ "$OS" = "windows" ]; then MY_PWD="$(cygpath -m "$(pwd)")\\\\" else @@ -11,7 +13,7 @@ export IDRIS2_PREFIX=${MY_PWD}currently/nonexistent/dir/ export IDRIS2_INC_CGS= -$1 --install ./testpkg.ipkg | sed -r "s/.([0-9]){10}//g" +idris2 --install ./testpkg.ipkg | sed -r "s/.([0-9]){10}//g" # ../ is there for some extra safety for using rm -rf rm -rf ../pkg010/build ../pkg010/currently diff --git a/tests/idris2/pkg011/run b/tests/idris2/pkg011/run index 320087b564..78172fc980 100755 --- a/tests/idris2/pkg011/run +++ b/tests/idris2/pkg011/run @@ -1,23 +1,25 @@ +. ../../testutils.sh + for t in dot dot-slash dot-slash-dot; do cd $t echo $t - $1 --build testpkg.ipkg + idris2 --build testpkg.ipkg rm -rf build/ - $1 --check --find-ipkg Main.idr + check --find-ipkg Main.idr rm -rf build/ cd .. done echo dot-slash-name-slash cd dot-slash-name-slash -$1 --build testpkg.ipkg +idris2 --build testpkg.ipkg rm -rf build/ -$1 --check --find-ipkg src/Main.idr +check --find-ipkg src/Main.idr rm -rf build/ cd .. echo sibling/pkg cd sibling/pkg -$1 --build testpkg.ipkg +idris2 --build testpkg.ipkg rm -rf build/ cd ../.. diff --git a/tests/idris2/pkg012/run b/tests/idris2/pkg012/run index 362349ad75..f5469d88ab 100755 --- a/tests/idris2/pkg012/run +++ b/tests/idris2/pkg012/run @@ -1,3 +1,3 @@ -rm -rf ./build/ +. ../../testutils.sh -$1 --build testpkg.ipkg +idris2 --build testpkg.ipkg diff --git a/tests/idris2/pkg013/run b/tests/idris2/pkg013/run index e2ec795e77..84b6a3e573 100644 --- a/tests/idris2/pkg013/run +++ b/tests/idris2/pkg013/run @@ -1,3 +1,5 @@ +. ../../testutils.sh + if [ "$OS" = "windows" ]; then MY_PWD="$(cygpath -m "$(pwd)")\\\\" else @@ -6,4 +8,4 @@ fi MY_PWD="${MY_PWD}" ./gen_expected.sh -$1 --build test.ipkg --log package.depends:10 +idris2 --build test.ipkg --log package.depends:10 diff --git a/tests/idris2/pkg014/run b/tests/idris2/pkg014/run index e2ec795e77..84b6a3e573 100644 --- a/tests/idris2/pkg014/run +++ b/tests/idris2/pkg014/run @@ -1,3 +1,5 @@ +. ../../testutils.sh + if [ "$OS" = "windows" ]; then MY_PWD="$(cygpath -m "$(pwd)")\\\\" else @@ -6,4 +8,4 @@ fi MY_PWD="${MY_PWD}" ./gen_expected.sh -$1 --build test.ipkg --log package.depends:10 +idris2 --build test.ipkg --log package.depends:10 diff --git a/tests/idris2/pkg015/run b/tests/idris2/pkg015/run index 6dc46021d7..a14384b5c0 100644 --- a/tests/idris2/pkg015/run +++ b/tests/idris2/pkg015/run @@ -1 +1,3 @@ -$1 --build test.ipkg --log package.depends:10 +. ../../../testutils.sh + +idris2 --build test.ipkg --log package.depends:10 diff --git a/tests/idris2/pkg016/run b/tests/idris2/pkg016/run index 7d5399b2c2..7b26a4531a 100644 --- a/tests/idris2/pkg016/run +++ b/tests/idris2/pkg016/run @@ -1,11 +1,12 @@ -$1 --install foo.ipkg > /dev/null -$1 --install bar.ipkg > /dev/null -$1 --install baz.ipkg > /dev/null -$1 --build test.ipkg +. ../../../testutils.sh + +idris2 --install foo.ipkg > /dev/null +idris2 --install bar.ipkg > /dev/null +idris2 --install baz.ipkg > /dev/null +idris2 --build test.ipkg build/exec/test rm -r "${IDRIS2_PREFIX}"/idris2-*/foo-0 rm -r "${IDRIS2_PREFIX}"/idris2-*/bar-0 rm -r "${IDRIS2_PREFIX}"/idris2-*/baz-0 -rm -r build diff --git a/tests/idris2/pkg017/run b/tests/idris2/pkg017/run index b5fbf48c0e..3cbf19eb7e 100644 --- a/tests/idris2/pkg017/run +++ b/tests/idris2/pkg017/run @@ -1,3 +1,5 @@ +. ../../testutils.sh + if [ "$OS" = "windows" ]; then MY_PWD="$(cygpath -m "$(pwd)")\\\\" else @@ -8,11 +10,11 @@ MY_PWD="${MY_PWD}" ./gen_expected.sh mkdir prefix -IDRIS2_PREFIX="${MY_PWD}/prefix" $1 --install-with-src a1/a1.ipkg > /dev/null -IDRIS2_PREFIX="${MY_PWD}/prefix" $1 --install-with-src a2/a2.ipkg > /dev/null +IDRIS2_PREFIX="${MY_PWD}/prefix" idris2 --install-with-src a1/a1.ipkg > /dev/null +IDRIS2_PREFIX="${MY_PWD}/prefix" idris2 --install-with-src a2/a2.ipkg > /dev/null -IDRIS2_PREFIX="${MY_PWD}/prefix" $1 --no-prelude --no-color --console-width 0 --ide-mode < input1 | grep -v ":highlight-source" | cut -c 7- -IDRIS2_PREFIX="${MY_PWD}/prefix" $1 --no-prelude --no-color --console-width 0 --ide-mode < input2 | grep -v ":highlight-source" | cut -c 7- +IDRIS2_PREFIX="${MY_PWD}/prefix" idris2 --no-prelude --ide-mode < input1 | grep -v ":highlight-source" | cut -c 7- +IDRIS2_PREFIX="${MY_PWD}/prefix" idris2 --no-prelude --ide-mode < input2 | grep -v ":highlight-source" | cut -c 7- rm -r a1/build a2/build b1/build b2/build rm -rf prefix diff --git a/tests/idris2/positivity001/run b/tests/idris2/positivity001/run index 37a2509d4f..96ab89ccf9 100644 --- a/tests/idris2/positivity001/run +++ b/tests/idris2/positivity001/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-banner --no-color --console-width 0 Issue660.idr --check +. ../../testutils.sh +check Issue660.idr diff --git a/tests/idris2/positivity002/run b/tests/idris2/positivity002/run index 37a2509d4f..96ab89ccf9 100644 --- a/tests/idris2/positivity002/run +++ b/tests/idris2/positivity002/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-banner --no-color --console-width 0 Issue660.idr --check +. ../../testutils.sh +check Issue660.idr diff --git a/tests/idris2/positivity003/run b/tests/idris2/positivity003/run index 37a2509d4f..96ab89ccf9 100644 --- a/tests/idris2/positivity003/run +++ b/tests/idris2/positivity003/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-banner --no-color --console-width 0 Issue660.idr --check +. ../../testutils.sh +check Issue660.idr diff --git a/tests/idris2/positivity004/run b/tests/idris2/positivity004/run index 1db4c278f6..274b479254 100644 --- a/tests/idris2/positivity004/run +++ b/tests/idris2/positivity004/run @@ -1,5 +1,5 @@ -rm -rf build +. ../../testutils.sh -$1 --no-banner --no-color --console-width 0 Issue1771-1.idr --check || true -$1 --no-banner --no-color --console-width 0 Issue1771-2.idr --check || true -$1 --no-banner --no-color --console-width 0 Issue1771-3.idr --check || true \ No newline at end of file +check Issue1771-1.idr || true +check Issue1771-2.idr || true +check Issue1771-3.idr || true diff --git a/tests/idris2/pretty001/run b/tests/idris2/pretty001/run index 674b63cfb8..7b69792ef2 100644 --- a/tests/idris2/pretty001/run +++ b/tests/idris2/pretty001/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --no-banner Issue1328A.idr < input +idris2 Issue1328A.idr < input diff --git a/tests/idris2/pretty002/run b/tests/idris2/pretty002/run index 30854d92ee..fecfb4b107 100644 --- a/tests/idris2/pretty002/run +++ b/tests/idris2/pretty002/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --no-banner < input +idris2 < input diff --git a/tests/idris2/primloop/run b/tests/idris2/primloop/run index c71cd4dbc1..33107d5859 100644 --- a/tests/idris2/primloop/run +++ b/tests/idris2/primloop/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --no-banner PrimLoop.idr < input +idris2 PrimLoop.idr < input diff --git a/tests/idris2/quantifiers001/run b/tests/idris2/quantifiers001/run index 384583246b..5ea4e26bd1 100755 --- a/tests/idris2/quantifiers001/run +++ b/tests/idris2/quantifiers001/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --no-banner TestQuantifiers.idr < input +idris2 TestQuantifiers.idr < input diff --git a/tests/idris2/real001/run b/tests/idris2/real001/run index 7108a56ac3..dbedfe8f05 100644 --- a/tests/idris2/real001/run +++ b/tests/idris2/real001/run @@ -1,5 +1,4 @@ -rm -rf build - -$1 --no-color --console-width 0 --check TestProto.idr -#$1 --no-color --console-width 0 --check MakeChans.idr +. ../../testutils.sh +check TestProto.idr +#check MakeChans.idr diff --git a/tests/idris2/real002/run b/tests/idris2/real002/run index 447a958c91..3bb3bdd4b3 100644 --- a/tests/idris2/real002/run +++ b/tests/idris2/real002/run @@ -1,5 +1,4 @@ -rm -rf build - -$1 --no-color --console-width 0 --check Store.idr -$1 --no-color --console-width 0 --check StoreL.idr +. ../../testutils.sh +check Store.idr +check StoreL.idr diff --git a/tests/idris2/record001/run b/tests/idris2/record001/run index 77e8ba283a..62e148929e 100755 --- a/tests/idris2/record001/run +++ b/tests/idris2/record001/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner Record.idr < input +. ../../testutils.sh +idris2 Record.idr < input diff --git a/tests/idris2/record002/run b/tests/idris2/record002/run index 77e8ba283a..62e148929e 100755 --- a/tests/idris2/record002/run +++ b/tests/idris2/record002/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner Record.idr < input +. ../../testutils.sh +idris2 Record.idr < input diff --git a/tests/idris2/record003/run b/tests/idris2/record003/run index d5f64b07d5..80e434e591 100755 --- a/tests/idris2/record003/run +++ b/tests/idris2/record003/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 -c --no-color --console-width 0 --no-banner Record.idr +. ../../testutils.sh +idris2 -c Record.idr diff --git a/tests/idris2/record004/run b/tests/idris2/record004/run index 6002294133..37beba4789 100755 --- a/tests/idris2/record004/run +++ b/tests/idris2/record004/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner Main.idr < input +. ../../testutils.sh +idris2 Main.idr < input diff --git a/tests/idris2/record005/run b/tests/idris2/record005/run index 339b3ef5a2..da8d57ecaf 100755 --- a/tests/idris2/record005/run +++ b/tests/idris2/record005/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner Fld.idr < input +. ../../testutils.sh +idris2 Fld.idr < input diff --git a/tests/idris2/record006/run b/tests/idris2/record006/run index 339b3ef5a2..da8d57ecaf 100755 --- a/tests/idris2/record006/run +++ b/tests/idris2/record006/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner Fld.idr < input +. ../../testutils.sh +idris2 Fld.idr < input diff --git a/tests/idris2/record007/run b/tests/idris2/record007/run index 52df0913a1..7b1104dc1d 100755 --- a/tests/idris2/record007/run +++ b/tests/idris2/record007/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner Bond.idr < input +. ../../testutils.sh +idris2 Bond.idr < input diff --git a/tests/idris2/record008/run b/tests/idris2/record008/run index dc4e51ad4b..64619990c7 100755 --- a/tests/idris2/record008/run +++ b/tests/idris2/record008/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner Postfix.idr < input +. ../../testutils.sh +idris2 Postfix.idr < input diff --git a/tests/idris2/record009/run b/tests/idris2/record009/run index 2dec80a337..23ab35bdff 100755 --- a/tests/idris2/record009/run +++ b/tests/idris2/record009/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner record.idr < input +. ../../testutils.sh +idris2 record.idr < input diff --git a/tests/idris2/record010/run b/tests/idris2/record010/run index 79cd4ba44e..df93277329 100755 --- a/tests/idris2/record010/run +++ b/tests/idris2/record010/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner --check record.idr +. ../../testutils.sh +check record.idr diff --git a/tests/idris2/record011/run b/tests/idris2/record011/run index aa88885547..19ee100b0b 100755 --- a/tests/idris2/record011/run +++ b/tests/idris2/record011/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --no-banner --check Issue2095.idr +check Issue2095.idr diff --git a/tests/idris2/record012/run b/tests/idris2/record012/run index a3cc956ce9..c0ea24df05 100755 --- a/tests/idris2/record012/run +++ b/tests/idris2/record012/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --no-banner --check Issue2065.idr +check Issue2065.idr diff --git a/tests/idris2/record013/run b/tests/idris2/record013/run index b958b384c3..4574c57288 100755 --- a/tests/idris2/record013/run +++ b/tests/idris2/record013/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --no-banner Issue1945.idr < input +idris2 Issue1945.idr < input diff --git a/tests/idris2/record014/run b/tests/idris2/record014/run index 717182e6e6..a5c8d5d776 100755 --- a/tests/idris2/record014/run +++ b/tests/idris2/record014/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --no-banner --check Issue1404.idr +check Issue1404.idr diff --git a/tests/idris2/record015/run b/tests/idris2/record015/run index b14678126f..de9479f270 100755 --- a/tests/idris2/record015/run +++ b/tests/idris2/record015/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --no-banner --exec main Issue2176.idr +run Issue2176.idr diff --git a/tests/idris2/record016/run b/tests/idris2/record016/run index 5d4614a7c9..e4a5065fc5 100755 --- a/tests/idris2/record016/run +++ b/tests/idris2/record016/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --no-banner HoleRecord.idr < input +idris2 HoleRecord.idr < input diff --git a/tests/idris2/record017/run b/tests/idris2/record017/run index 796e1c37c2..db5de30ae5 100755 --- a/tests/idris2/record017/run +++ b/tests/idris2/record017/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --no-banner -c RecordOptions.idr +idris2 -c RecordOptions.idr diff --git a/tests/idris2/record018/run b/tests/idris2/record018/run index 85017a2fd0..ce2877d08b 100755 --- a/tests/idris2/record018/run +++ b/tests/idris2/record018/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 mut.idr --check +. ../../testutils.sh +check mut.idr diff --git a/tests/idris2/record019/run b/tests/idris2/record019/run index 0de098e19e..20aaadf3c9 100755 --- a/tests/idris2/record019/run +++ b/tests/idris2/record019/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --no-banner -c BindParams.idr +idris2 -c BindParams.idr diff --git a/tests/idris2/reflection001/run b/tests/idris2/reflection001/run index eecdb30db5..59012e6d4f 100755 --- a/tests/idris2/reflection001/run +++ b/tests/idris2/reflection001/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner quote.idr < input +. ../../testutils.sh +idris2 quote.idr < input diff --git a/tests/idris2/reflection002/run b/tests/idris2/reflection002/run index ce864d2d0e..8b5520427c 100755 --- a/tests/idris2/reflection002/run +++ b/tests/idris2/reflection002/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner power.idr < input +. ../../testutils.sh +idris2 power.idr < input diff --git a/tests/idris2/reflection003/run b/tests/idris2/reflection003/run index 4317b72258..5c4359d7a3 100755 --- a/tests/idris2/reflection003/run +++ b/tests/idris2/reflection003/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 refprims.idr --check +. ../../testutils.sh +check refprims.idr diff --git a/tests/idris2/reflection004/run b/tests/idris2/reflection004/run index 0429c5cc30..e86f7b9b4f 100755 --- a/tests/idris2/reflection004/run +++ b/tests/idris2/reflection004/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner refdecl.idr < input +. ../../testutils.sh +idris2 refdecl.idr < input diff --git a/tests/idris2/reflection005/run b/tests/idris2/reflection005/run index 0429c5cc30..e86f7b9b4f 100755 --- a/tests/idris2/reflection005/run +++ b/tests/idris2/reflection005/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner refdecl.idr < input +. ../../testutils.sh +idris2 refdecl.idr < input diff --git a/tests/idris2/reflection006/run b/tests/idris2/reflection006/run index 535b44e96c..f70c2e6a82 100755 --- a/tests/idris2/reflection006/run +++ b/tests/idris2/reflection006/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --check refleq.idr +. ../../testutils.sh +check refleq.idr diff --git a/tests/idris2/reflection007/run b/tests/idris2/reflection007/run index 1dbad30530..be06b99c47 100755 --- a/tests/idris2/reflection007/run +++ b/tests/idris2/reflection007/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner NatExpr.idr < input +. ../../testutils.sh +idris2 NatExpr.idr < input diff --git a/tests/idris2/reflection008/run b/tests/idris2/reflection008/run index 43847fecc3..43638a3c41 100755 --- a/tests/idris2/reflection008/run +++ b/tests/idris2/reflection008/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner Interp.idr < input +. ../../testutils.sh +idris2 Interp.idr < input diff --git a/tests/idris2/reflection009/run b/tests/idris2/reflection009/run index 6fe7bf1092..c938b03d82 100755 --- a/tests/idris2/reflection009/run +++ b/tests/idris2/reflection009/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --check perf.idr +. ../../testutils.sh +check perf.idr diff --git a/tests/idris2/reflection010/run b/tests/idris2/reflection010/run index e02d592c8f..bc5391405b 100755 --- a/tests/idris2/reflection010/run +++ b/tests/idris2/reflection010/run @@ -1,3 +1,3 @@ -$1 --no-color --console-width 0 --no-banner --check Name.idr | sed -E "s/\.[0-9]+:[0-9]+/\.N:N/" +. ../../testutils.sh -rm -rf build +check Name.idr | sed -E "s/\.[0-9]+:[0-9]+/\.N:N/" diff --git a/tests/idris2/reflection011/run b/tests/idris2/reflection011/run index 5632d4d2e1..e7143a45d0 100755 --- a/tests/idris2/reflection011/run +++ b/tests/idris2/reflection011/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner tryref.idr < input +. ../../testutils.sh +idris2 tryref.idr < input diff --git a/tests/idris2/reflection012/run b/tests/idris2/reflection012/run index 34bd99454c..28c36f5f1d 100755 --- a/tests/idris2/reflection012/run +++ b/tests/idris2/reflection012/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner nameinfo.idr < input +. ../../testutils.sh +idris2 nameinfo.idr < input diff --git a/tests/idris2/reflection013/run b/tests/idris2/reflection013/run index 05d35edeaf..7a48fcc609 100755 --- a/tests/idris2/reflection013/run +++ b/tests/idris2/reflection013/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner --check WithUnambig.idr +. ../../testutils.sh +check WithUnambig.idr diff --git a/tests/idris2/reflection014/run b/tests/idris2/reflection014/run index 0429c5cc30..e86f7b9b4f 100755 --- a/tests/idris2/reflection014/run +++ b/tests/idris2/reflection014/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner refdecl.idr < input +. ../../testutils.sh +idris2 refdecl.idr < input diff --git a/tests/idris2/reflection015/run b/tests/idris2/reflection015/run index 4c330832aa..2e3dd06660 100755 --- a/tests/idris2/reflection015/run +++ b/tests/idris2/reflection015/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner --check MacroRetFunc.idr +. ../../testutils.sh +check MacroRetFunc.idr diff --git a/tests/idris2/reflection016/run b/tests/idris2/reflection016/run index a5acd1fce1..c7981f9d56 100755 --- a/tests/idris2/reflection016/run +++ b/tests/idris2/reflection016/run @@ -1,5 +1,4 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner Eta.idr < input -$1 --no-color --console-width 0 --no-banner --check BindElabScBug.idr +. ../../testutils.sh +idris2 Eta.idr < input +check BindElabScBug.idr diff --git a/tests/idris2/reflection017/run b/tests/idris2/reflection017/run index 42f9e04a6c..d2458e5ed4 100755 --- a/tests/idris2/reflection017/run +++ b/tests/idris2/reflection017/run @@ -1,4 +1,4 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --no-banner --check CanElabType.idr -$1 --no-color --console-width 0 --no-banner --check StillCantEscape.idr +check CanElabType.idr +check StillCantEscape.idr diff --git a/tests/idris2/reflection018/run b/tests/idris2/reflection018/run index a6195f4a79..4abf199262 100755 --- a/tests/idris2/reflection018/run +++ b/tests/idris2/reflection018/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --no-banner --check AtTypeLevel.idr +check AtTypeLevel.idr diff --git a/tests/idris2/reflection019/run b/tests/idris2/reflection019/run index a1e4637cd8..c779442ba2 100755 --- a/tests/idris2/reflection019/run +++ b/tests/idris2/reflection019/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --no-banner ElabScriptWarning.idr < input +idris2 ElabScriptWarning.idr < input diff --git a/tests/idris2/reflection020/run b/tests/idris2/reflection020/run index 9306d1e8d7..c82eda21a5 100644 --- a/tests/idris2/reflection020/run +++ b/tests/idris2/reflection020/run @@ -1,6 +1,6 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --no-banner FromTTImp.idr << EOF +idris2 FromTTImp.idr << EOF :printdef natExprMacroTest :printdef natExprPrecedenceTest :printdef natExprScriptTest @@ -8,14 +8,14 @@ $1 --no-color --console-width 0 --no-banner FromTTImp.idr << EOF :q EOF -$1 --no-color --console-width 0 --no-banner FromName.idr << EOF +idris2 FromName.idr << EOF :printdef myNameMacroTest :printdef myNameScriptTest myNameFunctionTest :q EOF -$1 --no-color --console-width 0 --no-banner FromDecls.idr << EOF +idris2 FromDecls.idr << EOF :printdef natDeclsMacroTest :printdef natDeclsScriptTest natDeclsFunctionTest diff --git a/tests/idris2/reflection021/run b/tests/idris2/reflection021/run index ce510e716e..4ca6d240e0 100755 --- a/tests/idris2/reflection021/run +++ b/tests/idris2/reflection021/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --no-banner --check QuoteSearch.idr +check QuoteSearch.idr diff --git a/tests/idris2/reg001/run b/tests/idris2/reg001/run index 0b9d5f35ea..e927e12c0b 100755 --- a/tests/idris2/reg001/run +++ b/tests/idris2/reg001/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --check D.idr +. ../../testutils.sh +check D.idr diff --git a/tests/idris2/reg002/run b/tests/idris2/reg002/run index c0d757d48f..c2084f482a 100755 --- a/tests/idris2/reg002/run +++ b/tests/idris2/reg002/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --check linm.idr +. ../../testutils.sh +check linm.idr diff --git a/tests/idris2/reg003/run b/tests/idris2/reg003/run index d82682cfbd..e33955b1a0 100755 --- a/tests/idris2/reg003/run +++ b/tests/idris2/reg003/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --no-banner Holes.idr < input +idris2 Holes.idr < input diff --git a/tests/idris2/reg004/run b/tests/idris2/reg004/run index 1d2365fb4f..10632b2def 100755 --- a/tests/idris2/reg004/run +++ b/tests/idris2/reg004/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --check ambig.idr +. ../../testutils.sh +check ambig.idr diff --git a/tests/idris2/reg005/run b/tests/idris2/reg005/run index 2c5473229d..6d1af86408 100755 --- a/tests/idris2/reg005/run +++ b/tests/idris2/reg005/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --check iftype.idr +. ../../testutils.sh +check iftype.idr diff --git a/tests/idris2/reg006/run b/tests/idris2/reg006/run index 8c77da2d31..05d093b40b 100755 --- a/tests/idris2/reg006/run +++ b/tests/idris2/reg006/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --check Cmd.idr +. ../../testutils.sh +check Cmd.idr diff --git a/tests/idris2/reg007/run b/tests/idris2/reg007/run index c3f47a417c..64c80a5b21 100755 --- a/tests/idris2/reg007/run +++ b/tests/idris2/reg007/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --check Main.idr +. ../../testutils.sh +check Main.idr diff --git a/tests/idris2/reg008/run b/tests/idris2/reg008/run index fcb4499324..7457509d1f 100755 --- a/tests/idris2/reg008/run +++ b/tests/idris2/reg008/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 Vending.idr --no-banner --debug-elab-check < input +. ../../testutils.sh +idris2 Vending.idr --debug-elab-check < input diff --git a/tests/idris2/reg009/run b/tests/idris2/reg009/run index 2f746b23a7..190bd11b91 100755 --- a/tests/idris2/reg009/run +++ b/tests/idris2/reg009/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 Case.idr --check +. ../../testutils.sh +check Case.idr diff --git a/tests/idris2/reg010/run b/tests/idris2/reg010/run index f1bacdfa5e..33f2d1ada4 100755 --- a/tests/idris2/reg010/run +++ b/tests/idris2/reg010/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 Recordname.idr --check +. ../../testutils.sh +check Recordname.idr diff --git a/tests/idris2/reg011/run b/tests/idris2/reg011/run index 85017a2fd0..ce2877d08b 100755 --- a/tests/idris2/reg011/run +++ b/tests/idris2/reg011/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 mut.idr --check +. ../../testutils.sh +check mut.idr diff --git a/tests/idris2/reg012/run b/tests/idris2/reg012/run index dd7af0aee6..8b984ed2bc 100755 --- a/tests/idris2/reg012/run +++ b/tests/idris2/reg012/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 Foo.idr --check +. ../../testutils.sh +check Foo.idr diff --git a/tests/idris2/reg013/run b/tests/idris2/reg013/run index 0687bb1ec9..8b47b02e13 100755 --- a/tests/idris2/reg013/run +++ b/tests/idris2/reg013/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 UnboundImplicits.idr --check +. ../../testutils.sh +check UnboundImplicits.idr diff --git a/tests/idris2/reg014/run b/tests/idris2/reg014/run index 95660f1414..be0ede11d9 100755 --- a/tests/idris2/reg014/run +++ b/tests/idris2/reg014/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 casecase.idr --check --debug-elab-check +. ../../testutils.sh +check casecase.idr --debug-elab-check diff --git a/tests/idris2/reg015/run b/tests/idris2/reg015/run index c03ffb24c5..581897bbad 100755 --- a/tests/idris2/reg015/run +++ b/tests/idris2/reg015/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 anyfail.idr --check +. ../../testutils.sh +check anyfail.idr diff --git a/tests/idris2/reg016/run b/tests/idris2/reg016/run index 24bd5e2402..29afa74be8 100755 --- a/tests/idris2/reg016/run +++ b/tests/idris2/reg016/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 Using.idr --check +. ../../testutils.sh +check Using.idr diff --git a/tests/idris2/reg017/run b/tests/idris2/reg017/run index 096ac67bed..5465fa2dfc 100755 --- a/tests/idris2/reg017/run +++ b/tests/idris2/reg017/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 lammult.idr --check +. ../../testutils.sh +check lammult.idr diff --git a/tests/idris2/reg018/run b/tests/idris2/reg018/run index 3ff6b91ee5..508b8282e7 100755 --- a/tests/idris2/reg018/run +++ b/tests/idris2/reg018/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 cycle.idr --check +. ../../testutils.sh +check cycle.idr diff --git a/tests/idris2/reg019/run b/tests/idris2/reg019/run index f4c962db62..546e07fcd2 100755 --- a/tests/idris2/reg019/run +++ b/tests/idris2/reg019/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 lazybug.idr --check +. ../../testutils.sh +check lazybug.idr diff --git a/tests/idris2/reg020/run b/tests/idris2/reg020/run index 2488030b0c..9de98b0472 100755 --- a/tests/idris2/reg020/run +++ b/tests/idris2/reg020/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 matchlits.idr --no-banner < input +. ../../testutils.sh +idris2 matchlits.idr < input diff --git a/tests/idris2/reg021/run b/tests/idris2/reg021/run index 2204818c02..6f7f117224 100755 --- a/tests/idris2/reg021/run +++ b/tests/idris2/reg021/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --check case.idr +. ../../testutils.sh +check case.idr diff --git a/tests/idris2/reg022/run b/tests/idris2/reg022/run index 2204818c02..6f7f117224 100755 --- a/tests/idris2/reg022/run +++ b/tests/idris2/reg022/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --check case.idr +. ../../testutils.sh +check case.idr diff --git a/tests/idris2/reg023/run b/tests/idris2/reg023/run index 5a44b08701..3f140f93b2 100755 --- a/tests/idris2/reg023/run +++ b/tests/idris2/reg023/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --check boom.idr +. ../../testutils.sh +check boom.idr diff --git a/tests/idris2/reg024/run b/tests/idris2/reg024/run index 456d6ed0ea..17f6211da4 100755 --- a/tests/idris2/reg024/run +++ b/tests/idris2/reg024/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner split.idr < input +. ../../testutils.sh +idris2 split.idr < input diff --git a/tests/idris2/reg025/run b/tests/idris2/reg025/run index 718347d056..157940aa4b 100755 --- a/tests/idris2/reg025/run +++ b/tests/idris2/reg025/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner lift.idr < input +. ../../testutils.sh +idris2 lift.idr < input diff --git a/tests/idris2/reg026/run b/tests/idris2/reg026/run index c42b04ff7b..f5d5dd5d96 100755 --- a/tests/idris2/reg026/run +++ b/tests/idris2/reg026/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --check Meh.idr +. ../../testutils.sh +check Meh.idr diff --git a/tests/idris2/reg027/run b/tests/idris2/reg027/run index ac597bc184..4d6e1ff77b 100755 --- a/tests/idris2/reg027/run +++ b/tests/idris2/reg027/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --check pwhere.idr +. ../../testutils.sh +check pwhere.idr diff --git a/tests/idris2/reg028/run b/tests/idris2/reg028/run index 586e9799fe..64b10e3f8e 100755 --- a/tests/idris2/reg028/run +++ b/tests/idris2/reg028/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --check Test.idr +. ../../testutils.sh +check Test.idr diff --git a/tests/idris2/reg029/run b/tests/idris2/reg029/run index dccce02910..fcbc83b59b 100755 --- a/tests/idris2/reg029/run +++ b/tests/idris2/reg029/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --check lqueue.idr +. ../../testutils.sh +check lqueue.idr diff --git a/tests/idris2/reg030/run b/tests/idris2/reg030/run index 7131440592..e8267b3ab2 100755 --- a/tests/idris2/reg030/run +++ b/tests/idris2/reg030/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --check A.idr +. ../../testutils.sh +check A.idr diff --git a/tests/idris2/reg031/run b/tests/idris2/reg031/run index 990276edb2..a08b0bef9c 100755 --- a/tests/idris2/reg031/run +++ b/tests/idris2/reg031/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --check dpair.idr +. ../../testutils.sh +check dpair.idr diff --git a/tests/idris2/reg032/run b/tests/idris2/reg032/run index 68d2f832b7..72cff3a29c 100755 --- a/tests/idris2/reg032/run +++ b/tests/idris2/reg032/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --check recupdate.idr +. ../../testutils.sh +check recupdate.idr diff --git a/tests/idris2/reg033/run b/tests/idris2/reg033/run index c886d939ff..d0f513a107 100755 --- a/tests/idris2/reg033/run +++ b/tests/idris2/reg033/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --check test.idr +. ../../testutils.sh +check test.idr diff --git a/tests/idris2/reg034/run b/tests/idris2/reg034/run index 091b7eb34f..1092d1eb15 100755 --- a/tests/idris2/reg034/run +++ b/tests/idris2/reg034/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --check void.idr +. ../../testutils.sh +check void.idr diff --git a/tests/idris2/reg035/run b/tests/idris2/reg035/run index 3021eaf6ec..096de90fd3 100755 --- a/tests/idris2/reg035/run +++ b/tests/idris2/reg035/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 Implicit.idr --check +. ../../testutils.sh +check Implicit.idr diff --git a/tests/idris2/reg036/run b/tests/idris2/reg036/run index 8bf84b6b0a..64b10e3f8e 100755 --- a/tests/idris2/reg036/run +++ b/tests/idris2/reg036/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 Test.idr --check +. ../../testutils.sh +check Test.idr diff --git a/tests/idris2/reg037/run b/tests/idris2/reg037/run index 8bf84b6b0a..64b10e3f8e 100755 --- a/tests/idris2/reg037/run +++ b/tests/idris2/reg037/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 Test.idr --check +. ../../testutils.sh +check Test.idr diff --git a/tests/idris2/reg038/run b/tests/idris2/reg038/run index a8d8e647f0..50feabb225 100755 --- a/tests/idris2/reg038/run +++ b/tests/idris2/reg038/run @@ -1,5 +1,4 @@ -rm -rf build - -$1 --no-color --console-width 0 Test1.idr --check -$1 --no-color --console-width 0 Test2.idr --check +. ../../testutils.sh +check Test1.idr +check Test2.idr diff --git a/tests/idris2/reg039/run b/tests/idris2/reg039/run index c34c5ac5b1..52897db3e3 100755 --- a/tests/idris2/reg039/run +++ b/tests/idris2/reg039/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --check dupdup.idr +. ../../testutils.sh +check dupdup.idr diff --git a/tests/idris2/reg040/run b/tests/idris2/reg040/run index c897e5dce1..4c3a475e32 100755 --- a/tests/idris2/reg040/run +++ b/tests/idris2/reg040/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --check CoverBug.idr +. ../../testutils.sh +check CoverBug.idr diff --git a/tests/idris2/reg041/run b/tests/idris2/reg041/run index 33a85b5e2b..00fb744d9e 100755 --- a/tests/idris2/reg041/run +++ b/tests/idris2/reg041/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --check tuple.idr +. ../../testutils.sh +check tuple.idr diff --git a/tests/idris2/reg042/run b/tests/idris2/reg042/run index f031736885..2f7509220c 100644 --- a/tests/idris2/reg042/run +++ b/tests/idris2/reg042/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner NatOpts.idr < input +. ../../testutils.sh +idris2 NatOpts.idr < input diff --git a/tests/idris2/reg043/run b/tests/idris2/reg043/run index 48f91c0572..740fd28d27 100755 --- a/tests/idris2/reg043/run +++ b/tests/idris2/reg043/run @@ -1,7 +1,7 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --check NotFake.idr +check NotFake.idr echo "TT2This Is A Fake TTC" > build/ttc/$(ls build/ttc)/Fake.ttc -$1 --no-color --console-width 0 --check TestFake.idr +check TestFake.idr diff --git a/tests/idris2/reg044/run b/tests/idris2/reg044/run index d4edee6e1f..d459abd54a 100755 --- a/tests/idris2/reg044/run +++ b/tests/idris2/reg044/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --check Methods.idr +. ../../testutils.sh +check Methods.idr diff --git a/tests/idris2/reg045/run b/tests/idris2/reg045/run index b5842e31d7..8426db0e77 100755 --- a/tests/idris2/reg045/run +++ b/tests/idris2/reg045/run @@ -1,3 +1,3 @@ -$1 --no-color --console-width 0 --check withparams.idr +. ../../testutils.sh -rm -rf build +check withparams.idr diff --git a/tests/idris2/reg046/run b/tests/idris2/reg046/run index 51d5193d20..1b63d28c3d 100755 --- a/tests/idris2/reg046/run +++ b/tests/idris2/reg046/run @@ -1,3 +1,3 @@ -$1 --no-color --console-width 0 --no-banner Postpone.idr < input +. ../../testutils.sh -rm -rf build +idris2 Postpone.idr < input diff --git a/tests/idris2/reg047/run b/tests/idris2/reg047/run index b163270623..d1c2a87cbb 100755 --- a/tests/idris2/reg047/run +++ b/tests/idris2/reg047/run @@ -1,2 +1,3 @@ -$1 --no-banner --no-color --console-width 0 QualifiedDoBang.idr < input -rm -rf build +. ../../testutils.sh + +idris2 QualifiedDoBang.idr < input diff --git a/tests/idris2/reg048/run b/tests/idris2/reg048/run index 099e1ea35f..d5473efbc4 100755 --- a/tests/idris2/reg048/run +++ b/tests/idris2/reg048/run @@ -1,2 +1,2 @@ -rm -rf build -$1 --no-banner --no-color --console-width 0 inferror.idr --check -p contrib +. ../../testutils.sh +check inferror.idr -p contrib diff --git a/tests/idris2/reg049/run b/tests/idris2/reg049/run index d10f3201c1..e5dc521b16 100755 --- a/tests/idris2/reg049/run +++ b/tests/idris2/reg049/run @@ -1,2 +1,2 @@ -rm -rf build -$1 --no-banner --no-color --console-width 0 lettype.idr --check -p contrib +. ../../testutils.sh +check lettype.idr -p contrib diff --git a/tests/idris2/reg050/run b/tests/idris2/reg050/run index 52395de146..66707ed52f 100755 --- a/tests/idris2/reg050/run +++ b/tests/idris2/reg050/run @@ -1,2 +1,2 @@ -rm -rf build -$1 --no-banner --no-color --console-width 0 loopy.idr --check -p contrib +. ../../testutils.sh +check loopy.idr -p contrib diff --git a/tests/idris2/reg051/run b/tests/idris2/reg051/run index b44dda73ae..2b332393bf 100755 --- a/tests/idris2/reg051/run +++ b/tests/idris2/reg051/run @@ -1,2 +1,2 @@ -rm -rf build -$1 --no-banner --no-color --console-width 0 BigFins.idr --check +. ../../testutils.sh +check BigFins.idr diff --git a/tests/idris2/reg052/run b/tests/idris2/reg052/run index 6279478fe9..bdd7a10f4e 100755 --- a/tests/idris2/reg052/run +++ b/tests/idris2/reg052/run @@ -1,2 +1,2 @@ -rm -rf build -$1 --no-banner --no-color --console-width 0 DPairQuote.idr --check +. ../../testutils.sh +check DPairQuote.idr diff --git a/tests/idris2/repl001/run b/tests/idris2/repl001/run index d54b883a79..fecfb4b107 100755 --- a/tests/idris2/repl001/run +++ b/tests/idris2/repl001/run @@ -1 +1,3 @@ -$1 --no-color --console-width 0 --no-banner < input +. ../../testutils.sh + +idris2 < input diff --git a/tests/idris2/repl002/run b/tests/idris2/repl002/run index d54b883a79..fecfb4b107 100755 --- a/tests/idris2/repl002/run +++ b/tests/idris2/repl002/run @@ -1 +1,3 @@ -$1 --no-color --console-width 0 --no-banner < input +. ../../testutils.sh + +idris2 < input diff --git a/tests/idris2/repl003/run b/tests/idris2/repl003/run index d54b883a79..fecfb4b107 100755 --- a/tests/idris2/repl003/run +++ b/tests/idris2/repl003/run @@ -1 +1,3 @@ -$1 --no-color --console-width 0 --no-banner < input +. ../../testutils.sh + +idris2 < input diff --git a/tests/idris2/repl004/run b/tests/idris2/repl004/run index d54b883a79..fecfb4b107 100755 --- a/tests/idris2/repl004/run +++ b/tests/idris2/repl004/run @@ -1 +1,3 @@ -$1 --no-color --console-width 0 --no-banner < input +. ../../testutils.sh + +idris2 < input diff --git a/tests/idris2/repl005/run b/tests/idris2/repl005/run index d54b883a79..fecfb4b107 100755 --- a/tests/idris2/repl005/run +++ b/tests/idris2/repl005/run @@ -1 +1,3 @@ -$1 --no-color --console-width 0 --no-banner < input +. ../../testutils.sh + +idris2 < input diff --git a/tests/idris2/repl006/run b/tests/idris2/repl006/run index d54b883a79..fecfb4b107 100755 --- a/tests/idris2/repl006/run +++ b/tests/idris2/repl006/run @@ -1 +1,3 @@ -$1 --no-color --console-width 0 --no-banner < input +. ../../testutils.sh + +idris2 < input diff --git a/tests/idris2/rewrite001/run b/tests/idris2/rewrite001/run index 40284ee225..48e0430d41 100755 --- a/tests/idris2/rewrite001/run +++ b/tests/idris2/rewrite001/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --no-banner --check Issue2573.idr +check Issue2573.idr diff --git a/tests/idris2/schemeeval001/run b/tests/idris2/schemeeval001/run index 47008ffee7..fecfb4b107 100755 --- a/tests/idris2/schemeeval001/run +++ b/tests/idris2/schemeeval001/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner < input +. ../../testutils.sh +idris2 < input diff --git a/tests/idris2/schemeeval002/run b/tests/idris2/schemeeval002/run index 47008ffee7..fecfb4b107 100755 --- a/tests/idris2/schemeeval002/run +++ b/tests/idris2/schemeeval002/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner < input +. ../../testutils.sh +idris2 < input diff --git a/tests/idris2/schemeeval003/run b/tests/idris2/schemeeval003/run index 47008ffee7..fecfb4b107 100755 --- a/tests/idris2/schemeeval003/run +++ b/tests/idris2/schemeeval003/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner < input +. ../../testutils.sh +idris2 < input diff --git a/tests/idris2/schemeeval004/run b/tests/idris2/schemeeval004/run index 24353950f2..6d6d987b11 100755 --- a/tests/idris2/schemeeval004/run +++ b/tests/idris2/schemeeval004/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner list.idr < input +. ../../testutils.sh +idris2 list.idr < input diff --git a/tests/idris2/schemeeval005/run b/tests/idris2/schemeeval005/run index 171913dd5b..a854f30fd6 100755 --- a/tests/idris2/schemeeval005/run +++ b/tests/idris2/schemeeval005/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner Printf.idr < input +. ../../testutils.sh +idris2 Printf.idr < input diff --git a/tests/idris2/schemeeval006/run b/tests/idris2/schemeeval006/run index 216c8e05c7..fecfb4b107 100644 --- a/tests/idris2/schemeeval006/run +++ b/tests/idris2/schemeeval006/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-banner --no-color --console-width 0 < input +idris2 < input diff --git a/tests/idris2/spec001/run b/tests/idris2/spec001/run index 74c2cf2c50..7306f953da 100644 --- a/tests/idris2/spec001/run +++ b/tests/idris2/spec001/run @@ -1,8 +1,8 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --no-banner --log specialise:5 -c Mult3.idr -$1 --no-color --console-width 0 --no-banner --cg node -o mult3.js -c Mult3.idr +idris2 --log specialise:5 -c Mult3.idr +idris2 --cg node -o mult3.js -c Mult3.idr awk -v RS= '/\/\* Main.main/' build/exec/mult3.js -$1 --no-color --console-width 0 --no-banner --log specialise:5 -c Desc.idr -$1 --no-color --console-width 0 --no-banner --log specialise:5 -c Desc2.idr -$1 --no-color --console-width 0 --no-banner --log specialise:5 -c Identity.idr \ No newline at end of file +idris2 --log specialise:5 -c Desc.idr +idris2 --log specialise:5 -c Desc2.idr +idris2 --log specialise:5 -c Identity.idr diff --git a/tests/idris2/termination001/run b/tests/idris2/termination001/run index b7086c3a56..a035eb51c1 100755 --- a/tests/idris2/termination001/run +++ b/tests/idris2/termination001/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --no-banner -c AgdaIssue6059.idr +idris2 -c AgdaIssue6059.idr diff --git a/tests/idris2/total001/run b/tests/idris2/total001/run index 4f97a2d2d7..c1204f255d 100755 --- a/tests/idris2/total001/run +++ b/tests/idris2/total001/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner Total.idr < input +. ../../testutils.sh +idris2 Total.idr < input diff --git a/tests/idris2/total002/run b/tests/idris2/total002/run index 4f97a2d2d7..c1204f255d 100755 --- a/tests/idris2/total002/run +++ b/tests/idris2/total002/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner Total.idr < input +. ../../testutils.sh +idris2 Total.idr < input diff --git a/tests/idris2/total003/run b/tests/idris2/total003/run index 4f97a2d2d7..c1204f255d 100755 --- a/tests/idris2/total003/run +++ b/tests/idris2/total003/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner Total.idr < input +. ../../testutils.sh +idris2 Total.idr < input diff --git a/tests/idris2/total004/run b/tests/idris2/total004/run index 4f97a2d2d7..c1204f255d 100755 --- a/tests/idris2/total004/run +++ b/tests/idris2/total004/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner Total.idr < input +. ../../testutils.sh +idris2 Total.idr < input diff --git a/tests/idris2/total005/run b/tests/idris2/total005/run index 4f97a2d2d7..c1204f255d 100755 --- a/tests/idris2/total005/run +++ b/tests/idris2/total005/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner Total.idr < input +. ../../testutils.sh +idris2 Total.idr < input diff --git a/tests/idris2/total006/run b/tests/idris2/total006/run index 4f97a2d2d7..c1204f255d 100755 --- a/tests/idris2/total006/run +++ b/tests/idris2/total006/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner Total.idr < input +. ../../testutils.sh +idris2 Total.idr < input diff --git a/tests/idris2/total007/run b/tests/idris2/total007/run index d1891c919e..5c09e30a92 100755 --- a/tests/idris2/total007/run +++ b/tests/idris2/total007/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner partial.idr --check +. ../../testutils.sh +check partial.idr diff --git a/tests/idris2/total008/run b/tests/idris2/total008/run index d1891c919e..5c09e30a92 100755 --- a/tests/idris2/total008/run +++ b/tests/idris2/total008/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner partial.idr --check +. ../../testutils.sh +check partial.idr diff --git a/tests/idris2/total009/run b/tests/idris2/total009/run index 65d1dce1ea..e33f3104e3 100755 --- a/tests/idris2/total009/run +++ b/tests/idris2/total009/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 tree.idr --check +. ../../testutils.sh +check tree.idr diff --git a/tests/idris2/total010/run b/tests/idris2/total010/run index a31050712e..e21261e515 100755 --- a/tests/idris2/total010/run +++ b/tests/idris2/total010/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 PartialWith.idr --check +. ../../testutils.sh +check PartialWith.idr diff --git a/tests/idris2/total011/run b/tests/idris2/total011/run index be1ee5673e..fa22291fc8 100755 --- a/tests/idris2/total011/run +++ b/tests/idris2/total011/run @@ -1,7 +1,7 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 Issue1782.idr --check -$1 --no-color --console-width 0 Issue1460.idr --check -$1 --no-color --console-width 0 Issue1859.idr --check -$1 --no-color --console-width 0 Issue1859-2.idr --check -$1 --no-color --console-width 0 Issue1828.idr --check +check Issue1782.idr +check Issue1460.idr +check Issue1859.idr +check Issue1859-2.idr +check Issue1828.idr diff --git a/tests/idris2/total012/run b/tests/idris2/total012/run index 9ce204c224..aa40ac525d 100755 --- a/tests/idris2/total012/run +++ b/tests/idris2/total012/run @@ -1,12 +1,12 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 Issue1828.idr --check -$1 --no-color --console-width 0 Issue1828.idr --check --total --log totality.requirement:20 | sed -r "s/.([0-9]){10}//g" +check Issue1828.idr +check Issue1828.idr --total --log totality.requirement:20 | sed -r "s/.([0-9]){10}//g" -$1 --no-color --console-width 0 Issue1828.idr --check -$1 --no-color --console-width 0 Issue1828.idr --check --log totality.requirement:20 | sed -r "s/.([0-9]){10}//g" +check Issue1828.idr +check Issue1828.idr --log totality.requirement:20 | sed -r "s/.([0-9]){10}//g" -$1 --no-color --console-width 0 TotallyTotal.idr --check -$1 --no-color --console-width 0 TotallyTotal.idr --check --log totality.requirement:20 | sed -r "s/.([0-9]){10}//g" -$1 --no-color --console-width 0 TotallyTotal.idr --check --total --log totality.requirement:20 | sed -r "s/.([0-9]){10}//g" \ No newline at end of file +check TotallyTotal.idr +check TotallyTotal.idr --log totality.requirement:20 | sed -r "s/.([0-9]){10}//g" +check TotallyTotal.idr --total --log totality.requirement:20 | sed -r "s/.([0-9]){10}//g" diff --git a/tests/idris2/total013/run b/tests/idris2/total013/run index 717182e6e6..a5c8d5d776 100755 --- a/tests/idris2/total013/run +++ b/tests/idris2/total013/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --no-banner --check Issue1404.idr +check Issue1404.idr diff --git a/tests/idris2/total014/run b/tests/idris2/total014/run index 645c8f240c..f2e9b12faa 100755 --- a/tests/idris2/total014/run +++ b/tests/idris2/total014/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --no-banner --check FunCompTC.idr +check FunCompTC.idr diff --git a/tests/idris2/total015/run b/tests/idris2/total015/run index a8c62c7e59..c2aaa1a2b4 100755 --- a/tests/idris2/total015/run +++ b/tests/idris2/total015/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --no-banner --check CoveringData.idr +check CoveringData.idr diff --git a/tests/idris2/total016/run b/tests/idris2/total016/run index c9d0605dcd..f7a0e2fe14 100755 --- a/tests/idris2/total016/run +++ b/tests/idris2/total016/run @@ -1,4 +1,4 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --no-banner --check AssertPositivity.idr -$1 --no-color --console-width 0 --no-banner --check LazyPositivityCheck.idr +check AssertPositivity.idr +check LazyPositivityCheck.idr diff --git a/tests/idris2/total017/run b/tests/idris2/total017/run index f7ba081d77..ffc5516b50 100755 --- a/tests/idris2/total017/run +++ b/tests/idris2/total017/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --no-banner --check Paper.idr +check Paper.idr diff --git a/tests/idris2/total018/run b/tests/idris2/total018/run index 672e4172ac..320b54769f 100755 --- a/tests/idris2/total018/run +++ b/tests/idris2/total018/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --no-banner --check Issue2448.idr +check Issue2448.idr diff --git a/tests/idris2/total019/run b/tests/idris2/total019/run index bfc998660c..afe2e9700b 100755 --- a/tests/idris2/total019/run +++ b/tests/idris2/total019/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --no-banner --check Check.idr +check Check.idr diff --git a/tests/idris2/total020/run b/tests/idris2/total020/run index bfc998660c..afe2e9700b 100755 --- a/tests/idris2/total020/run +++ b/tests/idris2/total020/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --no-banner --check Check.idr +check Check.idr diff --git a/tests/idris2/unification001/run b/tests/idris2/unification001/run index 23ed7dd08f..0921a91ff0 100755 --- a/tests/idris2/unification001/run +++ b/tests/idris2/unification001/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --no-banner --check Issue647.idr +check Issue647.idr diff --git a/tests/idris2/warning001/run b/tests/idris2/warning001/run index 22324f1d37..2c0748ecfd 100755 --- a/tests/idris2/warning001/run +++ b/tests/idris2/warning001/run @@ -1,8 +1,7 @@ -rm -rf build - -$1 --no-color --console-width 0 --check Issue539.idr -$1 --no-color --console-width 0 --check Issue621.idr -$1 --no-color --console-width 0 --check Issue1401.idr -$1 --no-color --console-width 0 --check PR1407.idr -$1 --no-color --console-width 0 --check Holes.idr +. ../../testutils.sh +check Issue539.idr +check Issue621.idr +check Issue1401.idr +check PR1407.idr +check Holes.idr diff --git a/tests/idris2/warning002/run b/tests/idris2/warning002/run index c3e3b212de..7542784c76 100755 --- a/tests/idris2/warning002/run +++ b/tests/idris2/warning002/run @@ -1,8 +1,7 @@ -rm -rf ./build +. ../../testutils.sh -$1 --no-color --console-width 0 --build deprecated.ipkg -$1 --no-color --no-banner --console-width 0 Foo.idr < /dev/null +idris2 --cg node -o node_args TestArgs.idr > /dev/null node ./build/exec/node_args a b node ./build/exec/node_args c - diff --git a/tests/node/bitops/run b/tests/node/bitops/run index 5d135d8ec8..d5069e73e0 100644 --- a/tests/node/bitops/run +++ b/tests/node/bitops/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --cg node --no-banner --no-color --console-width 0 BitOps.idr < input +. ../../testutils.sh +idris2 --cg node BitOps.idr < input diff --git a/tests/node/casts/run b/tests/node/casts/run index c795c59c76..0f3c79c675 100644 --- a/tests/node/casts/run +++ b/tests/node/casts/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --cg node --no-banner --no-color --console-width 0 Casts.idr < input +. ../../testutils.sh +idris2 --cg node Casts.idr < input diff --git a/tests/node/doubles/run b/tests/node/doubles/run index 39d34b93f1..32bc5b3598 100644 --- a/tests/node/doubles/run +++ b/tests/node/doubles/run @@ -1,5 +1,4 @@ -rm -rf build +. ../../testutils.sh -$1 --no-banner --no-color --console-width 0 --cg node -o node_doubles.js Doubles.idr > /dev/null +idris2 --cg node -o node_doubles.js Doubles.idr > /dev/null node build/exec/node_doubles.js - diff --git a/tests/node/fastConcat/run b/tests/node/fastConcat/run index 127caa5994..de7425502b 100644 --- a/tests/node/fastConcat/run +++ b/tests/node/fastConcat/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --cg node --no-banner --no-color --console-width 0 FastConcat.idr < input +. ../../testutils.sh +idris2 --cg node FastConcat.idr < input diff --git a/tests/node/fix1839/run b/tests/node/fix1839/run index e30015854b..60f3f7f7d7 100644 --- a/tests/node/fix1839/run +++ b/tests/node/fix1839/run @@ -1,4 +1,4 @@ -rm -rf build +. ../../testutils.sh -$1 --cg node --no-banner --no-color --console-width 0 OS.idr < input -$1 --cg javascript --no-banner --no-color --console-width 0 -o os.js OS.idr +idris2 --cg node OS.idr < input +idris2 --cg javascript -o os.js OS.idr diff --git a/tests/node/idiom001/run b/tests/node/idiom001/run index cb2af60efd..9e47f7b816 100644 --- a/tests/node/idiom001/run +++ b/tests/node/idiom001/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --cg node --no-banner Main.idr < input +. ../../testutils.sh +idris2 --cg node Main.idr < input diff --git a/tests/node/integer_array/run b/tests/node/integer_array/run index 482634d2b9..a67f8abb85 100755 --- a/tests/node/integer_array/run +++ b/tests/node/integer_array/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --cg node --no-banner array.idr < input +idris2 --cg node array.idr < input diff --git a/tests/node/integers/run b/tests/node/integers/run index 80c2f5566d..27effc59f8 100644 --- a/tests/node/integers/run +++ b/tests/node/integers/run @@ -1,5 +1,4 @@ -rm -rf build +. ../../testutils.sh -$1 --no-banner --no-color --console-width 0 --cg node -o node_integers.js TestIntegers.idr > /dev/null +idris2 --cg node -o node_integers.js TestIntegers.idr > /dev/null node build/exec/node_integers.js - diff --git a/tests/node/memo/run b/tests/node/memo/run index 32630c1850..32e8075391 100644 --- a/tests/node/memo/run +++ b/tests/node/memo/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --cg node --no-banner --no-color --console-width 0 Memo.idr < input +. ../../testutils.sh +idris2 --cg node Memo.idr < input diff --git a/tests/node/newints/run b/tests/node/newints/run index d62fdda2a6..16900feb15 100644 --- a/tests/node/newints/run +++ b/tests/node/newints/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --cg node --no-banner --no-color --console-width 0 IntOps.idr < input +. ../../testutils.sh +idris2 --cg node IntOps.idr < input diff --git a/tests/node/node001/run b/tests/node/node001/run index 1a43c17472..a044ce7c09 100755 --- a/tests/node/node001/run +++ b/tests/node/node001/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --cg node --no-banner Total.idr < input +. ../../testutils.sh +idris2 --cg node Total.idr < input diff --git a/tests/node/node002/run b/tests/node/node002/run index 596d8ce0b6..7d83820f8d 100755 --- a/tests/node/node002/run +++ b/tests/node/node002/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --cg node --no-banner Pythag.idr < input +idris2 --cg node Pythag.idr < input diff --git a/tests/node/node003/run b/tests/node/node003/run index 1ee4e49371..20afe1fe11 100755 --- a/tests/node/node003/run +++ b/tests/node/node003/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --cg node --no-banner IORef.idr < input +idris2 --cg node IORef.idr < input diff --git a/tests/node/node004/run b/tests/node/node004/run index b327678a28..4731309fd7 100755 --- a/tests/node/node004/run +++ b/tests/node/node004/run @@ -1,2 +1,5 @@ -$1 --no-color --console-width 0 --cg node --no-banner Buffer.idr < input -rm -rf build test.buf +. ../../testutils.sh + +rm test.buf + +idris2 --cg node Buffer.idr < input diff --git a/tests/node/node005/run b/tests/node/node005/run index a20d525f9c..ec47e16de0 100755 --- a/tests/node/node005/run +++ b/tests/node/node005/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --cg node --no-banner Filter.idr < input +. ../../testutils.sh +idris2 --cg node Filter.idr < input diff --git a/tests/node/node006/run b/tests/node/node006/run index 679e4677a5..53f8c44f3c 100755 --- a/tests/node/node006/run +++ b/tests/node/node006/run @@ -1,5 +1,4 @@ -rm -rf build - -$1 --no-color --console-width 0 --cg node --no-banner TypeCase.idr < input -$1 --no-color --console-width 0 --cg node --no-banner TypeCase2.idr --check +. ../../testutils.sh +idris2 --cg node TypeCase.idr < input +check --cg node TypeCase2.idr diff --git a/tests/node/node007/run b/tests/node/node007/run index 2f13221fff..7d554a8d79 100755 --- a/tests/node/node007/run +++ b/tests/node/node007/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --cg node --no-banner TypeCase.idr < input +. ../../testutils.sh +idris2 --cg node TypeCase.idr < input diff --git a/tests/node/node008/run b/tests/node/node008/run index cfd2932ed2..d7524b7005 100755 --- a/tests/node/node008/run +++ b/tests/node/node008/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --cg node --no-banner Nat.idr < input +. ../../testutils.sh +idris2 --cg node Nat.idr < input diff --git a/tests/node/node009/run b/tests/node/node009/run index 6fc0d7289b..e153ebec0e 100755 --- a/tests/node/node009/run +++ b/tests/node/node009/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --cg node --no-banner uni.idr < input +. ../../testutils.sh +idris2 --cg node uni.idr < input diff --git a/tests/node/node011/run b/tests/node/node011/run index 83676264a9..fe2b80cdd7 100644 --- a/tests/node/node011/run +++ b/tests/node/node011/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --cg node --no-banner bangs.idr < input +. ../../testutils.sh +idris2 --cg node bangs.idr < input diff --git a/tests/node/node012/run b/tests/node/node012/run index 482634d2b9..a67f8abb85 100755 --- a/tests/node/node012/run +++ b/tests/node/node012/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --cg node --no-banner array.idr < input +idris2 --cg node array.idr < input diff --git a/tests/node/node014/run b/tests/node/node014/run index 513db346aa..cbe7920e1c 100755 --- a/tests/node/node014/run +++ b/tests/node/node014/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --cg node --no-banner -p network Echo.idr < input +. ../../testutils.sh +idris2 --cg node -p network Echo.idr < input diff --git a/tests/node/node015/run b/tests/node/node015/run index e533c19be2..cf34c951c8 100755 --- a/tests/node/node015/run +++ b/tests/node/node015/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --cg node --no-banner Numbers.idr < input +. ../../testutils.sh +idris2 --cg node Numbers.idr < input diff --git a/tests/node/node017/run b/tests/node/node017/run index 742e3497fc..05c005b8ae 100755 --- a/tests/node/node017/run +++ b/tests/node/node017/run @@ -1,4 +1,7 @@ +. ../../testutils.sh + +rm -rf testdir + ./gen_expected.sh -$1 --no-color --console-width 0 --cg node --no-banner dir.idr < input +idris2 --cg node dir.idr < input cat testdir/test.txt -rm -rf build testdir diff --git a/tests/node/node018/run b/tests/node/node018/run index d0b5f2eeb7..f5cc69a816 100755 --- a/tests/node/node018/run +++ b/tests/node/node018/run @@ -1,3 +1,5 @@ -$1 --no-color --console-width 0 --cg node --no-banner File.idr < input +. ../../testutils.sh -rm -rf build testout.txt +rm testout.txt + +idris2 --cg node File.idr < input diff --git a/tests/node/node019/run b/tests/node/node019/run index 3d76c06493..22b79e49c8 100755 --- a/tests/node/node019/run +++ b/tests/node/node019/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --cg node --no-banner partial.idr < input +. ../../testutils.sh +idris2 --cg node partial.idr < input diff --git a/tests/node/node020/run b/tests/node/node020/run index 81cdebd5bd..809dc5fc59 100644 --- a/tests/node/node020/run +++ b/tests/node/node020/run @@ -1,4 +1,3 @@ -rm -rf build - -POPEN_CMD="$1 --version" $1 --no-color --console-width 0 --cg node --no-banner Popen.idr < input +. ../../testutils.sh +POPEN_CMD="$idris2 --version" idris2 --cg node Popen.idr < input diff --git a/tests/node/node021/run b/tests/node/node021/run index d9ad200187..5dd72645c5 100644 --- a/tests/node/node021/run +++ b/tests/node/node021/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --cg node --no-banner Bits.idr < input +. ../../testutils.sh +idris2 --cg node Bits.idr < input diff --git a/tests/node/node022/run b/tests/node/node022/run index 2732a3e760..624f06658f 100644 --- a/tests/node/node022/run +++ b/tests/node/node022/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --cg node --no-banner --no-color --console-width 0 BitCasts.idr < input +. ../../testutils.sh +idris2 --cg node BitCasts.idr < input diff --git a/tests/node/node023/run b/tests/node/node023/run index c795c59c76..0f3c79c675 100644 --- a/tests/node/node023/run +++ b/tests/node/node023/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --cg node --no-banner --no-color --console-width 0 Casts.idr < input +. ../../testutils.sh +idris2 --cg node Casts.idr < input diff --git a/tests/node/node024/run b/tests/node/node024/run index 5d135d8ec8..d5069e73e0 100644 --- a/tests/node/node024/run +++ b/tests/node/node024/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --cg node --no-banner --no-color --console-width 0 BitOps.idr < input +. ../../testutils.sh +idris2 --cg node BitOps.idr < input diff --git a/tests/node/node025/run b/tests/node/node025/run index 593a952eb5..0fb0d30f9f 100644 --- a/tests/node/node025/run +++ b/tests/node/node025/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --cg node --no-banner --no-color --console-width 0 Fix1037.idr < input +. ../../testutils.sh +idris2 --cg node Fix1037.idr < input diff --git a/tests/node/node026/run b/tests/node/node026/run index a6f98621c1..bf7de0183e 100644 --- a/tests/node/node026/run +++ b/tests/node/node026/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 -p contrib --cg node --no-banner --no-color --console-width 0 Fix1795.idr < input +. ../../testutils.sh +idris2 -p contrib --cg node Fix1795.idr < input diff --git a/tests/node/node027/run b/tests/node/node027/run index fc557d93b0..57df102cf8 100644 --- a/tests/node/node027/run +++ b/tests/node/node027/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --cg node --no-banner --no-color --console-width 0 PID.idr < input +idris2 --cg node PID.idr < input diff --git a/tests/node/nomangle001/run b/tests/node/nomangle001/run index 1177a59c04..8f3d635efd 100644 --- a/tests/node/nomangle001/run +++ b/tests/node/nomangle001/run @@ -1,7 +1,7 @@ -rm -rf build +. ../../testutils.sh -$1 nomangle.idr --no-color --no-banner --console-width 0 -o test1 --cg node -$1 nomangle.idr --no-color --no-banner --console-width 0 -o test2 --cg node --directive minimal +idris2 nomangle.idr -o test1 --cg node +idris2 nomangle.idr -o test2 --cg node --directive minimal cat build/exec/test1 | grep -e foo -e baz -e another_name cat build/exec/test2 | grep -e foo -e baz -e another_name diff --git a/tests/node/nomangle002/run b/tests/node/nomangle002/run index 2bf8b50e54..9d8a5c156b 100644 --- a/tests/node/nomangle002/run +++ b/tests/node/nomangle002/run @@ -1,4 +1,4 @@ -rm -rf build +. ../../testutils.sh -$1 nomangle1.idr --no-color --no-banner --console-width 0 -o test --cg node -$1 nomangle2.idr --no-color --no-banner --console-width 0 -o test --cg node +idris2 nomangle1.idr -o test --cg node +idris2 nomangle2.idr -o test --cg node diff --git a/tests/node/perf001/run b/tests/node/perf001/run index 6b4ef75318..09892f812e 100644 --- a/tests/node/perf001/run +++ b/tests/node/perf001/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --cg node --no-banner --no-color --console-width 0 Span.idr < input +. ../../testutils.sh +idris2 --cg node Span.idr < input diff --git a/tests/node/reg001/run b/tests/node/reg001/run index f6e106f166..dd8069709b 100755 --- a/tests/node/reg001/run +++ b/tests/node/reg001/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --cg node numbers.idr -x main +. ../../testutils.sh +run --cg node numbers.idr diff --git a/tests/node/reg002/run b/tests/node/reg002/run index 707ec9a7de..58498217ac 100755 --- a/tests/node/reg002/run +++ b/tests/node/reg002/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --cg node Issue1843.idr -x main +. ../../testutils.sh +run --cg node Issue1843.idr diff --git a/tests/node/stringcast/run b/tests/node/stringcast/run index 11dd056fc6..d6166759d3 100644 --- a/tests/node/stringcast/run +++ b/tests/node/stringcast/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --cg node --no-banner --no-color --console-width 0 StringCast.idr < input +. ../../testutils.sh +idris2 --cg node StringCast.idr < input diff --git a/tests/node/syntax001/run b/tests/node/syntax001/run index df458eac0c..9449b9c950 100644 --- a/tests/node/syntax001/run +++ b/tests/node/syntax001/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --cg node caseBlock.idr -x main +. ../../testutils.sh +run --cg node caseBlock.idr diff --git a/tests/node/tailrec001/run b/tests/node/tailrec001/run index 3e63f60f06..2a0746dde4 100755 --- a/tests/node/tailrec001/run +++ b/tests/node/tailrec001/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --cg node tailrec.idr -x main +. ../../testutils.sh +run --cg node tailrec.idr diff --git a/tests/node/tailrec002/run b/tests/node/tailrec002/run index 86c2062ba3..92679e9c1a 100755 --- a/tests/node/tailrec002/run +++ b/tests/node/tailrec002/run @@ -1,4 +1,4 @@ -rm -rf build -$1 --no-color --console-width 0 --no-banner --cg node EvenOdd.idr -o app.js -$1 --no-color --console-width 0 --no-banner Main.idr -x main +. ../../testutils.sh +idris2 --cg node EvenOdd.idr -o app.js +run Main.idr diff --git a/tests/node/tailrec_libs/run b/tests/node/tailrec_libs/run index 30a8d78d71..2a0746dde4 100644 --- a/tests/node/tailrec_libs/run +++ b/tests/node/tailrec_libs/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --cg node --no-banner --no-color --console-width 0 tailrec.idr -x main +run --cg node tailrec.idr diff --git a/tests/prelude/bind001/run b/tests/prelude/bind001/run index 0df264d40a..bcaf1b9b4f 100755 --- a/tests/prelude/bind001/run +++ b/tests/prelude/bind001/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --no-banner bind.idr < input +idris2 bind.idr < input diff --git a/tests/prelude/char001/run b/tests/prelude/char001/run index dc11cf71f6..0678b8e5ea 100755 --- a/tests/prelude/char001/run +++ b/tests/prelude/char001/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --check chars.idr +check chars.idr diff --git a/tests/prelude/double001/run b/tests/prelude/double001/run index c5fd56e3c6..34962defeb 100755 --- a/tests/prelude/double001/run +++ b/tests/prelude/double001/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --no-banner --no-color --console-width 0 Types.idr < input +idris2 Types.idr < input diff --git a/tests/prelude/nat001/run b/tests/prelude/nat001/run index 6c631040cb..3038799a9e 100755 --- a/tests/prelude/nat001/run +++ b/tests/prelude/nat001/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../testutils.sh -$1 --check nats.idr +check nats.idr diff --git a/tests/prelude/reg001/run b/tests/prelude/reg001/run index f286371315..5f8605d9bc 100755 --- a/tests/prelude/reg001/run +++ b/tests/prelude/reg001/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --check fixity.idr +. ../../testutils.sh +check fixity.idr diff --git a/tests/racket/barrier001/run b/tests/racket/barrier001/run index 93b2c4cac8..108c440fda 100644 --- a/tests/racket/barrier001/run +++ b/tests/racket/barrier001/run @@ -1 +1,3 @@ -$1 --no-banner --no-color --console-width 0 --cg chez Main.idr --exec main +. ../../testutils.sh + +run --cg chez Main.idr diff --git a/tests/racket/conditions001/run b/tests/racket/conditions001/run index 84c3a35328..d0aff72ed6 100644 --- a/tests/racket/conditions001/run +++ b/tests/racket/conditions001/run @@ -1 +1,3 @@ -$1 --no-banner --no-color --console-width 0 --cg racket Main.idr --exec main +. ../../testutils.sh + +run --cg racket Main.idr diff --git a/tests/racket/conditions002/run b/tests/racket/conditions002/run index 84c3a35328..d0aff72ed6 100644 --- a/tests/racket/conditions002/run +++ b/tests/racket/conditions002/run @@ -1 +1,3 @@ -$1 --no-banner --no-color --console-width 0 --cg racket Main.idr --exec main +. ../../testutils.sh + +run --cg racket Main.idr diff --git a/tests/racket/conditions003/run b/tests/racket/conditions003/run index 84c3a35328..d0aff72ed6 100644 --- a/tests/racket/conditions003/run +++ b/tests/racket/conditions003/run @@ -1 +1,3 @@ -$1 --no-banner --no-color --console-width 0 --cg racket Main.idr --exec main +. ../../testutils.sh + +run --cg racket Main.idr diff --git a/tests/racket/conditions004/run b/tests/racket/conditions004/run index 84c3a35328..d0aff72ed6 100644 --- a/tests/racket/conditions004/run +++ b/tests/racket/conditions004/run @@ -1 +1,3 @@ -$1 --no-banner --no-color --console-width 0 --cg racket Main.idr --exec main +. ../../testutils.sh + +run --cg racket Main.idr diff --git a/tests/racket/conditions005/run b/tests/racket/conditions005/run index 84c3a35328..d0aff72ed6 100644 --- a/tests/racket/conditions005/run +++ b/tests/racket/conditions005/run @@ -1 +1,3 @@ -$1 --no-banner --no-color --console-width 0 --cg racket Main.idr --exec main +. ../../testutils.sh + +run --cg racket Main.idr diff --git a/tests/racket/conditions006/run b/tests/racket/conditions006/run index 84c3a35328..d0aff72ed6 100644 --- a/tests/racket/conditions006/run +++ b/tests/racket/conditions006/run @@ -1 +1,3 @@ -$1 --no-banner --no-color --console-width 0 --cg racket Main.idr --exec main +. ../../testutils.sh + +run --cg racket Main.idr diff --git a/tests/racket/conditions007/run b/tests/racket/conditions007/run index 84c3a35328..d0aff72ed6 100644 --- a/tests/racket/conditions007/run +++ b/tests/racket/conditions007/run @@ -1 +1,3 @@ -$1 --no-banner --no-color --console-width 0 --cg racket Main.idr --exec main +. ../../testutils.sh + +run --cg racket Main.idr diff --git a/tests/racket/ffi001/run b/tests/racket/ffi001/run index fc4c5e3b66..825a99daca 100644 --- a/tests/racket/ffi001/run +++ b/tests/racket/ffi001/run @@ -1 +1,3 @@ -$1 --no-banner --no-color --console-width 0 --cg racket RacketLib.idr --exec main +. ../../testutils.sh + +run --cg racket RacketLib.idr diff --git a/tests/racket/forkjoin001/run b/tests/racket/forkjoin001/run index 84c3a35328..d0aff72ed6 100644 --- a/tests/racket/forkjoin001/run +++ b/tests/racket/forkjoin001/run @@ -1 +1,3 @@ -$1 --no-banner --no-color --console-width 0 --cg racket Main.idr --exec main +. ../../testutils.sh + +run --cg racket Main.idr diff --git a/tests/racket/futures001/run b/tests/racket/futures001/run index deec40623a..9f97ae23c9 100644 --- a/tests/racket/futures001/run +++ b/tests/racket/futures001/run @@ -1,4 +1,4 @@ -$1 --no-banner --no-color --console-width 0 --cg racket Futures.idr -p contrib --exec constant -$1 --no-banner --no-color --console-width 0 --cg racket Futures.idr -p contrib --exec map +. ../../testutils.sh -rm -r build +idris2 --cg racket Futures.idr -p contrib --exec constant +idris2 --cg racket Futures.idr -p contrib --exec map diff --git a/tests/racket/mutex001/run b/tests/racket/mutex001/run index 84c3a35328..d0aff72ed6 100644 --- a/tests/racket/mutex001/run +++ b/tests/racket/mutex001/run @@ -1 +1,3 @@ -$1 --no-banner --no-color --console-width 0 --cg racket Main.idr --exec main +. ../../testutils.sh + +run --cg racket Main.idr diff --git a/tests/racket/mutex002/run b/tests/racket/mutex002/run index 84c3a35328..d0aff72ed6 100644 --- a/tests/racket/mutex002/run +++ b/tests/racket/mutex002/run @@ -1 +1,3 @@ -$1 --no-banner --no-color --console-width 0 --cg racket Main.idr --exec main +. ../../testutils.sh + +run --cg racket Main.idr diff --git a/tests/racket/mutex003/run b/tests/racket/mutex003/run index 84c3a35328..d0aff72ed6 100644 --- a/tests/racket/mutex003/run +++ b/tests/racket/mutex003/run @@ -1 +1,3 @@ -$1 --no-banner --no-color --console-width 0 --cg racket Main.idr --exec main +. ../../testutils.sh + +run --cg racket Main.idr diff --git a/tests/racket/mutex004/run b/tests/racket/mutex004/run index 84c3a35328..d0aff72ed6 100644 --- a/tests/racket/mutex004/run +++ b/tests/racket/mutex004/run @@ -1 +1,3 @@ -$1 --no-banner --no-color --console-width 0 --cg racket Main.idr --exec main +. ../../testutils.sh + +run --cg racket Main.idr diff --git a/tests/racket/mutex005/run b/tests/racket/mutex005/run index 84c3a35328..d0aff72ed6 100644 --- a/tests/racket/mutex005/run +++ b/tests/racket/mutex005/run @@ -1 +1,3 @@ -$1 --no-banner --no-color --console-width 0 --cg racket Main.idr --exec main +. ../../testutils.sh + +run --cg racket Main.idr diff --git a/tests/racket/semaphores001/run b/tests/racket/semaphores001/run index 84c3a35328..d0aff72ed6 100644 --- a/tests/racket/semaphores001/run +++ b/tests/racket/semaphores001/run @@ -1 +1,3 @@ -$1 --no-banner --no-color --console-width 0 --cg racket Main.idr --exec main +. ../../testutils.sh + +run --cg racket Main.idr diff --git a/tests/racket/semaphores002/run b/tests/racket/semaphores002/run index 84c3a35328..d0aff72ed6 100644 --- a/tests/racket/semaphores002/run +++ b/tests/racket/semaphores002/run @@ -1 +1,3 @@ -$1 --no-banner --no-color --console-width 0 --cg racket Main.idr --exec main +. ../../testutils.sh + +run --cg racket Main.idr diff --git a/tests/refc/args/run b/tests/refc/args/run index df2974a9da..1915f72984 100644 --- a/tests/refc/args/run +++ b/tests/refc/args/run @@ -1,6 +1,5 @@ -rm -rf build +. ../../testutils.sh -$1 --no-banner --no-color --console-width 0 --cg refc -o refc_args TestArgs.idr > /dev/null +idris2 --cg refc -o refc_args TestArgs.idr > /dev/null ./build/exec/refc_args a b ./build/exec/refc_args c - diff --git a/tests/refc/buffer/run b/tests/refc/buffer/run index df1958e15e..20e4046dfe 100644 --- a/tests/refc/buffer/run +++ b/tests/refc/buffer/run @@ -1,6 +1,6 @@ -rm -rf build +. ../../testutils.sh -$1 --no-banner --no-color --console-width 0 --cg refc -o refc_buffer TestBuffer.idr > /dev/null +idris2 --cg refc -o refc_buffer TestBuffer.idr > /dev/null ./build/exec/refc_buffer base64 -i testWrite.buf diff --git a/tests/refc/clock/run b/tests/refc/clock/run index 91a64df313..06a0c52017 100644 --- a/tests/refc/clock/run +++ b/tests/refc/clock/run @@ -1,5 +1,4 @@ -rm -rf build +. ../../testutils.sh -$1 --no-banner --no-color --console-width 0 --cg refc -o refc_clock TestClock.idr > /dev/null +idris2 --cg refc -o refc_clock TestClock.idr > /dev/null ./build/exec/refc_clock - diff --git a/tests/refc/doubles/run b/tests/refc/doubles/run index 637e2c6296..764fdac189 100644 --- a/tests/refc/doubles/run +++ b/tests/refc/doubles/run @@ -1,5 +1,4 @@ -rm -rf build +. ../../testutils.sh -$1 --no-banner --no-color --console-width 0 --cg refc -o refc_doubles TestDoubles.idr > /dev/null +idris2 --cg refc -o refc_doubles TestDoubles.idr > /dev/null ./build/exec/refc_doubles - diff --git a/tests/refc/garbageCollect/run b/tests/refc/garbageCollect/run index e4c5dbb27f..8f2475a9ee 100644 --- a/tests/refc/garbageCollect/run +++ b/tests/refc/garbageCollect/run @@ -1,5 +1,5 @@ -rm -rf build +. ../../testutils.sh -$1 --no-banner --no-color --console-width 0 --cg refc -o testGC TestGarbageCollect.idr > /dev/null +idris2 --cg refc -o testGC TestGarbageCollect.idr > /dev/null ./build/exec/testGC diff --git a/tests/refc/integers/run b/tests/refc/integers/run index 8ad0dbb089..dd49b87c50 100644 --- a/tests/refc/integers/run +++ b/tests/refc/integers/run @@ -1,5 +1,4 @@ -rm -rf build +. ../../testutils.sh -$1 --no-banner --no-color --console-width 0 --cg refc -o refc_integers TestIntegers.idr > /dev/null +idris2 --cg refc -o refc_integers TestIntegers.idr > /dev/null ./build/exec/refc_integers - diff --git a/tests/refc/issue1778/run b/tests/refc/issue1778/run index abc3abd236..2e449ef715 100644 --- a/tests/refc/issue1778/run +++ b/tests/refc/issue1778/run @@ -1,5 +1,4 @@ -rm -rf build +. ../../testutils.sh -$1 --no-banner --no-color --console-width 0 --cg refc -o reverse Reverse.idr > /dev/null +idris2 --cg refc -o reverse Reverse.idr > /dev/null ./build/exec/reverse - diff --git a/tests/refc/issue2424/run b/tests/refc/issue2424/run index 01928ff14b..1defe762de 100644 --- a/tests/refc/issue2424/run +++ b/tests/refc/issue2424/run @@ -1,5 +1,5 @@ -rm -rf build +. ../../testutils.sh -$1 --no-banner --no-color --console-width 0 --cg refc -o test ControlAppMonadTest.idr > /dev/null +idris2 --cg refc -o test ControlAppMonadTest.idr > /dev/null ./build/exec/test diff --git a/tests/refc/issue2452/run b/tests/refc/issue2452/run index 0734ea225c..1a9e3d222b 100644 --- a/tests/refc/issue2452/run +++ b/tests/refc/issue2452/run @@ -1,4 +1,4 @@ -rm -rf build +. ../../testutils.sh -$1 --no-banner --no-color --console-width 0 --cg refc -o bits_case BitsCase.idr > /dev/null +idris2 --cg refc -o bits_case BitsCase.idr > /dev/null ./build/exec/bits_case diff --git a/tests/refc/refc001/run b/tests/refc/refc001/run index ae7b8087bd..8e6f2201bd 100644 --- a/tests/refc/refc001/run +++ b/tests/refc/refc001/run @@ -1,5 +1,4 @@ -rm -rf build +. ../../testutils.sh -$1 --no-banner --no-color --console-width 0 --cg refc -o refc001 Tail.idr > /dev/null +idris2 --cg refc -o refc001 Tail.idr > /dev/null ./build/exec/refc001 - diff --git a/tests/refc/refc002/run b/tests/refc/refc002/run index e3099b7c10..29eea06878 100644 --- a/tests/refc/refc002/run +++ b/tests/refc/refc002/run @@ -1,5 +1,4 @@ -rm -rf build +. ../../testutils.sh -$1 --no-banner --no-color --console-width 0 --cg refc -o refc002 RecordProjection.idr > /dev/null +idris2 --cg refc -o refc002 RecordProjection.idr > /dev/null ./build/exec/refc002 - diff --git a/tests/refc/refc003/run b/tests/refc/refc003/run index 569771fd4d..b9e893bf7b 100644 --- a/tests/refc/refc003/run +++ b/tests/refc/refc003/run @@ -1,4 +1,4 @@ -rm -rf build +. ../../testutils.sh -$1 --no-banner --no-color --console-width 0 --cg refc -o refc003 Issue1191.idr > /dev/null +idris2 --cg refc -o refc003 Issue1191.idr > /dev/null ./build/exec/refc003 diff --git a/tests/refc/reg001/run b/tests/refc/reg001/run index 910978f2fc..36982f2559 100755 --- a/tests/refc/reg001/run +++ b/tests/refc/reg001/run @@ -1,5 +1,5 @@ -rm -rf build +. ../../testutils.sh -$1 --no-color --console-width 0 --cg refc Issue1843.idr -o test +idris2 --cg refc Issue1843.idr -o test ./build/exec/test diff --git a/tests/refc/strings/run b/tests/refc/strings/run index 73ae9cf5ca..332dc1f35e 100644 --- a/tests/refc/strings/run +++ b/tests/refc/strings/run @@ -1,5 +1,4 @@ -rm -rf build +. ../../testutils.sh -$1 --no-banner --no-color --console-width 0 --cg refc -p contrib -o refc_strings TestStrings.idr > /dev/null +idris2 --cg refc -p contrib -o refc_strings TestStrings.idr > /dev/null ./build/exec/refc_strings - diff --git a/tests/templates/simple-test/notes.md b/tests/templates/simple-test/notes.md index c656f8fac8..22a1341ab9 100644 --- a/tests/templates/simple-test/notes.md +++ b/tests/templates/simple-test/notes.md @@ -6,9 +6,9 @@ template `with-ipkg`. ### Mandatory steps * Create a new subdirectory for the tests. Try to adhere to the naming scheme of existing tests. * Update the `tests/Main.idr`, adding the new subdirectory to the lists of tests. -* In the `run` script, the compiler location is bound to `$1` and all calls must have the `--no-color --console-width 0 ---no-banner` flags to avoid system dependent differences. Also keep the `rm -rf build` line as keeping build files can -lead unexpected differences in the output, due to compiler logs and not errors. +* The `testutils.sh` script defines the `idris2` name binding, and some testing utility functions. +* Automate deletion of any temporary files created by your test. This prevents inconsistencies between runs. The `build` + directory is deleted automatically by `testutils.sh`. * If the tests open a REPL session, remember to put a quit command, `:q`in the input file. * The files named `run`, `expected`, `output` are reserved by the test runner, do not overwrite them by running the test. diff --git a/tests/templates/simple-test/run b/tests/templates/simple-test/run index d2c00b30c9..441b64ba30 100755 --- a/tests/templates/simple-test/run +++ b/tests/templates/simple-test/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner Test.idr < input +. ../../testutils.sh +idris2 Test.idr < input diff --git a/tests/templates/ttimp/notes.md b/tests/templates/ttimp/notes.md index 25f4b8b7f7..049558fe91 100644 --- a/tests/templates/ttimp/notes.md +++ b/tests/templates/ttimp/notes.md @@ -6,8 +6,9 @@ Template for tests on the `TTImp` representation. * Create a new subdirectory for the tests. Try to adhere to the naming scheme of existing tests. For `TTImp` files the convention is to use the `.yaff` extension. * Update the `tests/Main.idr`, adding the new subdirectory to the lists of tests. -* In the `run` script, Each invocation to the `TTImp` REPL must be prefixed by the `--yaffle` flag. Also keep the `rm -rf build` line as keeping build files can -lead unexpected differences in the output, due to compiler logs and not errors. +* The `testutils.sh` script defines the `idris2` name binding, and some testing utility functions. +* Automate deletion of any temporary files created by your test. This prevents inconsistencies between runs. The `build` + directory is deleted automatically by `testutils.sh`. * If the tests open a REPL session, remember to put a quit command, `:q`in the input file. * The files named `run`, `expected`, `output` are reserved by the test runner, do not overwrite them by running the test. diff --git a/tests/templates/ttimp/run b/tests/templates/ttimp/run index d7e98d2836..8165669e3e 100755 --- a/tests/templates/ttimp/run +++ b/tests/templates/ttimp/run @@ -1,4 +1,4 @@ -rm -rf build +. ../../testutils.sh -$1 --yaffle Interp.yaff < input -$1 --yaffle build/ttc/$(ls build/ttc/)/Interp.ttc < input +idris2 --yaffle Interp.yaff < input +idris2 --yaffle build/ttc/$(ls build/ttc/)/Interp.ttc < input diff --git a/tests/templates/with-ipkg/notes.md b/tests/templates/with-ipkg/notes.md index b15926fc4a..329c036fe2 100644 --- a/tests/templates/with-ipkg/notes.md +++ b/tests/templates/with-ipkg/notes.md @@ -5,10 +5,9 @@ Template for tests that requires a package file. ### Mandatory steps * Create a new subdirectory for the tests. Try to adhere to the naming scheme of existing tests. * Update the `tests/Main.idr`, adding the new subdirectory to the lists of tests. -* In the `ipkg` always keep the `--no-color --console-width 0` flags in the `opts` field to avoid system dependent -differences. -* In the `run` script, keep the `rm -rf build` line as keeping build files can -lead unexpected differences in the output, due to compiler logs and not errors. +* The `testutils.sh` script defines the `idris2` name binding, and some testing utility functions. +* Automate deletion of any temporary files created by your test. This prevents inconsistencies between runs. The `build` + directory is deleted automatically by `testutils.sh`. * If the tests open a REPL session, remember to put a quit command, `:q`in the input file. * The files named `run`, `expected`, `output` are reserved by the test runner, do not overwrite them by running the test. diff --git a/tests/templates/with-ipkg/run b/tests/templates/with-ipkg/run index fa9604f3f3..3d99b32b7c 100755 --- a/tests/templates/with-ipkg/run +++ b/tests/templates/with-ipkg/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --build dummy.ipkg +. ../../testutils.sh +idris2 --build dummy.ipkg diff --git a/tests/testutils.sh b/tests/testutils.sh new file mode 100755 index 0000000000..24951eb143 --- /dev/null +++ b/tests/testutils.sh @@ -0,0 +1,23 @@ +#!/bin/sh + +# This script is intended to be sourced from test scripts. +# It provides a number of test utilities. +# Usage: . ../../testutils.sh + +idris2="$1" + +# Delete build files between runs to prevent unexpected differences. +# As this is at the top-level, this is run when this script is imported. +rm -rf build + +idris2() { + $idris2 --no-banner --console-width 0 --no-color "$@" +} + +check() { + idris2 --check "$@" +} + +run() { + idris2 --exec main "$@" +} diff --git a/tests/ttimp/basic001/run b/tests/ttimp/basic001/run index ce479ad0d3..f561b3c7ed 100755 --- a/tests/ttimp/basic001/run +++ b/tests/ttimp/basic001/run @@ -1,4 +1,4 @@ -rm -rf build +. ../../testutils.sh -$1 --yaffle Interp.yaff < input -$1 --yaffle build/ttc/*/Interp.ttc < input +idris2 --yaffle Interp.yaff < input +idris2 --yaffle build/ttc/*/Interp.ttc < input diff --git a/tests/ttimp/basic002/run b/tests/ttimp/basic002/run index 7404d9d938..89ec13e862 100755 --- a/tests/ttimp/basic002/run +++ b/tests/ttimp/basic002/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --yaffle Adder.yaff < input +. ../../testutils.sh +idris2 --yaffle Adder.yaff < input diff --git a/tests/ttimp/basic003/run b/tests/ttimp/basic003/run index 390f217736..14ee5abe1d 100755 --- a/tests/ttimp/basic003/run +++ b/tests/ttimp/basic003/run @@ -1,4 +1,4 @@ -rm -rf build +. ../../testutils.sh -$1 --yaffle Hole.yaff < input -$1 --yaffle build/ttc/*/Hole.ttc < input +idris2 --yaffle Hole.yaff < input +idris2 --yaffle build/ttc/*/Hole.ttc < input diff --git a/tests/ttimp/basic004/run b/tests/ttimp/basic004/run index 5d67d83244..83bc075b87 100755 --- a/tests/ttimp/basic004/run +++ b/tests/ttimp/basic004/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --yaffle AsPat.yaff < input +. ../../testutils.sh +idris2 --yaffle AsPat.yaff < input diff --git a/tests/ttimp/basic005/run b/tests/ttimp/basic005/run index 7d0918388b..02bdc51622 100755 --- a/tests/ttimp/basic005/run +++ b/tests/ttimp/basic005/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --yaffle Ambig.yaff < input +. ../../testutils.sh +idris2 --yaffle Ambig.yaff < input diff --git a/tests/ttimp/basic006/run b/tests/ttimp/basic006/run index 7d0918388b..02bdc51622 100755 --- a/tests/ttimp/basic006/run +++ b/tests/ttimp/basic006/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --yaffle Ambig.yaff < input +. ../../testutils.sh +idris2 --yaffle Ambig.yaff < input diff --git a/tests/ttimp/coverage002/run b/tests/ttimp/coverage002/run index 4f68a0b59d..014c3ba642 100755 --- a/tests/ttimp/coverage002/run +++ b/tests/ttimp/coverage002/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --yaffle Vect.yaff < input +. ../../testutils.sh +idris2 --yaffle Vect.yaff < input diff --git a/tests/ttimp/dot001/run b/tests/ttimp/dot001/run index 75611ab8f4..d4375a85dd 100755 --- a/tests/ttimp/dot001/run +++ b/tests/ttimp/dot001/run @@ -1,7 +1,6 @@ -rm -rf build - -echo ':q' | $1 --yaffle Dot.yaff -echo ':q' | $1 --yaffle Dot2.yaff -echo ':q' | $1 --yaffle Dot3.yaff -echo ':q' | $1 --yaffle Dot4.yaff +. ../../testutils.sh +echo ':q' | idris2 --yaffle Dot.yaff +echo ':q' | idris2 --yaffle Dot2.yaff +echo ':q' | idris2 --yaffle Dot3.yaff +echo ':q' | idris2 --yaffle Dot4.yaff diff --git a/tests/ttimp/eta001/run b/tests/ttimp/eta001/run index 995d4b651b..206622709e 100755 --- a/tests/ttimp/eta001/run +++ b/tests/ttimp/eta001/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --yaffle Eta.yaff < input +. ../../testutils.sh +idris2 --yaffle Eta.yaff < input diff --git a/tests/ttimp/lazy001/run b/tests/ttimp/lazy001/run index bd5cc7c9fb..22ce0e9f0a 100755 --- a/tests/ttimp/lazy001/run +++ b/tests/ttimp/lazy001/run @@ -1,5 +1,5 @@ -rm -rf build +. ../../testutils.sh -$1 --yaffle Lazy.yaff < input -$1 --yaffle LazyInf.yaff < input -$1 --yaffle build/ttc/*/LazyInf.ttc < input +idris2 --yaffle Lazy.yaff < input +idris2 --yaffle LazyInf.yaff < input +idris2 --yaffle build/ttc/*/LazyInf.ttc < input diff --git a/tests/ttimp/nest001/run b/tests/ttimp/nest001/run index 0ce0d07fc6..89e797d17d 100755 --- a/tests/ttimp/nest001/run +++ b/tests/ttimp/nest001/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --yaffle Let.yaff < input +. ../../testutils.sh +idris2 --yaffle Let.yaff < input diff --git a/tests/ttimp/nest002/run b/tests/ttimp/nest002/run index 41582a5a71..0f315c5543 100755 --- a/tests/ttimp/nest002/run +++ b/tests/ttimp/nest002/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --yaffle Case.yaff < input +. ../../testutils.sh +idris2 --yaffle Case.yaff < input diff --git a/tests/ttimp/perf001/run b/tests/ttimp/perf001/run index 3af6832228..6f2d29be1d 100755 --- a/tests/ttimp/perf001/run +++ b/tests/ttimp/perf001/run @@ -1,4 +1,3 @@ -rm -rf build - -echo ':q' | $1 --yaffle bigsuc.yaff +. ../../testutils.sh +echo ':q' | idris2 --yaffle bigsuc.yaff diff --git a/tests/ttimp/perf002/run b/tests/ttimp/perf002/run index 971cf00849..09580f77eb 100755 --- a/tests/ttimp/perf002/run +++ b/tests/ttimp/perf002/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --yaffle BigVect.yaff < input +. ../../testutils.sh +idris2 --yaffle BigVect.yaff < input diff --git a/tests/ttimp/perf003/run b/tests/ttimp/perf003/run index ce9825d231..e636ea3c57 100755 --- a/tests/ttimp/perf003/run +++ b/tests/ttimp/perf003/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --yaffle Id.yaff < input +. ../../testutils.sh +idris2 --yaffle Id.yaff < input diff --git a/tests/ttimp/qtt001/run b/tests/ttimp/qtt001/run index f520ad29ce..5c432b18c1 100755 --- a/tests/ttimp/qtt001/run +++ b/tests/ttimp/qtt001/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --yaffle QTT.yaff < input +. ../../testutils.sh +idris2 --yaffle QTT.yaff < input diff --git a/tests/ttimp/qtt003/run b/tests/ttimp/qtt003/run index 7bcf505677..b7c4e5f490 100755 --- a/tests/ttimp/qtt003/run +++ b/tests/ttimp/qtt003/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --yaffle QTTEq.yaff < input +. ../../testutils.sh +idris2 --yaffle QTTEq.yaff < input diff --git a/tests/ttimp/record001/run b/tests/ttimp/record001/run index afcde9aeb6..20a55e42c1 100755 --- a/tests/ttimp/record001/run +++ b/tests/ttimp/record001/run @@ -1,4 +1,4 @@ -rm -rf build +. ../../testutils.sh -$1 --yaffle Record.yaff < input -$1 --yaffle build/ttc/*/Record.ttc < input +idris2 --yaffle Record.yaff < input +idris2 --yaffle build/ttc/*/Record.ttc < input diff --git a/tests/ttimp/record002/run b/tests/ttimp/record002/run index 03483aa5ab..59170145c2 100755 --- a/tests/ttimp/record002/run +++ b/tests/ttimp/record002/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --yaffle Record.yaff < input +. ../../testutils.sh +idris2 --yaffle Record.yaff < input diff --git a/tests/ttimp/record003/run b/tests/ttimp/record003/run index afcde9aeb6..20a55e42c1 100755 --- a/tests/ttimp/record003/run +++ b/tests/ttimp/record003/run @@ -1,4 +1,4 @@ -rm -rf build +. ../../testutils.sh -$1 --yaffle Record.yaff < input -$1 --yaffle build/ttc/*/Record.ttc < input +idris2 --yaffle Record.yaff < input +idris2 --yaffle build/ttc/*/Record.ttc < input diff --git a/tests/ttimp/record004/run b/tests/ttimp/record004/run index afcde9aeb6..20a55e42c1 100644 --- a/tests/ttimp/record004/run +++ b/tests/ttimp/record004/run @@ -1,4 +1,4 @@ -rm -rf build +. ../../testutils.sh -$1 --yaffle Record.yaff < input -$1 --yaffle build/ttc/*/Record.ttc < input +idris2 --yaffle Record.yaff < input +idris2 --yaffle build/ttc/*/Record.ttc < input diff --git a/tests/ttimp/total001/run b/tests/ttimp/total001/run index 4f68a0b59d..014c3ba642 100755 --- a/tests/ttimp/total001/run +++ b/tests/ttimp/total001/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --yaffle Vect.yaff < input +. ../../testutils.sh +idris2 --yaffle Vect.yaff < input diff --git a/tests/ttimp/total002/run b/tests/ttimp/total002/run index dfa7a43e01..36c913715d 100755 --- a/tests/ttimp/total002/run +++ b/tests/ttimp/total002/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --yaffle Total.yaff < input +. ../../testutils.sh +idris2 --yaffle Total.yaff < input diff --git a/tests/ttimp/total003/run b/tests/ttimp/total003/run index ba3176f3c1..02d8a38704 100755 --- a/tests/ttimp/total003/run +++ b/tests/ttimp/total003/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --yaffle Bad.yaff < input +. ../../testutils.sh +idris2 --yaffle Bad.yaff < input diff --git a/tests/typedd-book/chapter01/run b/tests/typedd-book/chapter01/run index 0a5fc2f800..2827bb9c3a 100755 --- a/tests/typedd-book/chapter01/run +++ b/tests/typedd-book/chapter01/run @@ -1,5 +1,4 @@ -rm -rf build - -$1 --no-color --console-width 0 All.idr --check -$1 --no-color --console-width 0 --no-banner HoleFix.idr < input +. ../../testutils.sh +check All.idr +idris2 HoleFix.idr < input diff --git a/tests/typedd-book/chapter02/run b/tests/typedd-book/chapter02/run index 3a9f0e952b..3304749548 100755 --- a/tests/typedd-book/chapter02/run +++ b/tests/typedd-book/chapter02/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 All.idr --check +. ../../testutils.sh +check All.idr diff --git a/tests/typedd-book/chapter03/run b/tests/typedd-book/chapter03/run index 3a9f0e952b..3304749548 100755 --- a/tests/typedd-book/chapter03/run +++ b/tests/typedd-book/chapter03/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 All.idr --check +. ../../testutils.sh +check All.idr diff --git a/tests/typedd-book/chapter04/run b/tests/typedd-book/chapter04/run index 3a9f0e952b..3304749548 100755 --- a/tests/typedd-book/chapter04/run +++ b/tests/typedd-book/chapter04/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 All.idr --check +. ../../testutils.sh +check All.idr diff --git a/tests/typedd-book/chapter05/run b/tests/typedd-book/chapter05/run index 3a9f0e952b..3304749548 100755 --- a/tests/typedd-book/chapter05/run +++ b/tests/typedd-book/chapter05/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 All.idr --check +. ../../testutils.sh +check All.idr diff --git a/tests/typedd-book/chapter06/run b/tests/typedd-book/chapter06/run index 3a9f0e952b..3304749548 100755 --- a/tests/typedd-book/chapter06/run +++ b/tests/typedd-book/chapter06/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 All.idr --check +. ../../testutils.sh +check All.idr diff --git a/tests/typedd-book/chapter07/run b/tests/typedd-book/chapter07/run index 3a9f0e952b..3304749548 100755 --- a/tests/typedd-book/chapter07/run +++ b/tests/typedd-book/chapter07/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 All.idr --check +. ../../testutils.sh +check All.idr diff --git a/tests/typedd-book/chapter08/run b/tests/typedd-book/chapter08/run index 3a9f0e952b..3304749548 100755 --- a/tests/typedd-book/chapter08/run +++ b/tests/typedd-book/chapter08/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 All.idr --check +. ../../testutils.sh +check All.idr diff --git a/tests/typedd-book/chapter09/run b/tests/typedd-book/chapter09/run index 3a9f0e952b..3304749548 100755 --- a/tests/typedd-book/chapter09/run +++ b/tests/typedd-book/chapter09/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 All.idr --check +. ../../testutils.sh +check All.idr diff --git a/tests/typedd-book/chapter10/run b/tests/typedd-book/chapter10/run index db7b2ff58c..f8b9cfcf61 100755 --- a/tests/typedd-book/chapter10/run +++ b/tests/typedd-book/chapter10/run @@ -1,5 +1,4 @@ -rm -rf build - -$1 --no-color --console-width 0 DLFail.idr --check -$1 --no-color --console-width 0 All.idr --check +. ../../testutils.sh +check DLFail.idr +check All.idr diff --git a/tests/typedd-book/chapter11/run b/tests/typedd-book/chapter11/run index 3a9f0e952b..3304749548 100755 --- a/tests/typedd-book/chapter11/run +++ b/tests/typedd-book/chapter11/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 All.idr --check +. ../../testutils.sh +check All.idr diff --git a/tests/typedd-book/chapter12/run b/tests/typedd-book/chapter12/run index 3a9f0e952b..3304749548 100755 --- a/tests/typedd-book/chapter12/run +++ b/tests/typedd-book/chapter12/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 All.idr --check +. ../../testutils.sh +check All.idr diff --git a/tests/typedd-book/chapter13/run b/tests/typedd-book/chapter13/run index 3a9f0e952b..3304749548 100755 --- a/tests/typedd-book/chapter13/run +++ b/tests/typedd-book/chapter13/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 All.idr --check +. ../../testutils.sh +check All.idr diff --git a/tests/typedd-book/chapter14/run b/tests/typedd-book/chapter14/run index 3a9f0e952b..3304749548 100755 --- a/tests/typedd-book/chapter14/run +++ b/tests/typedd-book/chapter14/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 All.idr --check +. ../../testutils.sh +check All.idr diff --git a/tests/vmcode/basic001/run b/tests/vmcode/basic001/run index 7d28a16320..951db64d81 100755 --- a/tests/vmcode/basic001/run +++ b/tests/vmcode/basic001/run @@ -1,4 +1,3 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner Test.idr -x main --cg vmcode-interp +. ../../testutils.sh +run Test.idr --cg vmcode-interp From 61878ef1b65500637cfa19414254cc33e9c99163 Mon Sep 17 00:00:00 2001 From: Robert Wright Date: Tue, 29 Aug 2023 17:02:01 +0100 Subject: [PATCH 04/43] Remove trivial testing input files --- tests/allschemes/scheme001/input | 2 -- tests/base/control_app001/expected | 2 -- tests/base/control_app001/input | 0 tests/base/control_app001/run | 2 +- tests/base/data_bits001/expected | 4 +--- tests/base/data_bits001/input | 2 -- tests/base/data_bits001/run | 2 +- tests/base/data_bits002/expected | 4 +--- tests/base/data_bits002/input | 2 -- tests/base/data_bits002/run | 2 +- tests/base/data_integral/expected | 4 +--- tests/base/data_integral/input | 2 -- tests/base/data_integral/run | 2 +- tests/base/data_list001/expected | 4 +--- tests/base/data_list001/input | 2 -- tests/base/data_list001/run | 2 +- tests/base/data_string_lines001/expected | 4 +--- tests/base/data_string_lines001/input | 2 -- tests/base/data_string_lines001/run | 2 +- tests/base/data_string_unlines001/expected | 4 +--- tests/base/data_string_unlines001/input | 2 -- tests/base/data_string_unlines001/run | 2 +- tests/base/data_vect001/expected | 4 +--- tests/base/data_vect001/input | 2 -- tests/base/data_vect001/run | 2 +- tests/base/system_directory/expected | 2 -- tests/base/system_directory/input | 2 -- tests/base/system_directory/run | 2 +- tests/base/system_errno/expected | 2 -- tests/base/system_errno/input | 2 -- tests/base/system_errno/run | 2 +- tests/base/system_file001/expected | 7 ++----- tests/base/system_file001/input | 2 -- tests/base/system_file001/run | 10 +++++----- tests/base/system_file_copyFile/expected | 2 -- tests/base/system_file_copyFile/input | 2 -- tests/base/system_file_copyFile/run | 2 +- tests/base/system_file_fRead/expected | 4 +--- tests/base/system_file_fRead/input | 2 -- tests/base/system_file_fRead/run | 2 +- tests/base/system_file_popen/expected | 4 +--- tests/base/system_file_popen/input | 2 -- tests/base/system_file_popen/run | 2 +- tests/base/system_run/expected | 4 +--- tests/base/system_run/input | 2 -- tests/base/system_run/run | 2 +- tests/base/system_signal001/expected | 5 +---- tests/base/system_signal001/input | 1 - tests/base/system_signal001/run | 2 +- tests/base/system_signal002/expected | 5 +---- tests/base/system_signal002/input | 1 - tests/base/system_signal002/run | 2 +- tests/base/system_signal003/expected | 3 --- tests/base/system_signal003/input | 1 - tests/base/system_signal003/run | 2 +- tests/base/system_signal004/expected | 5 +---- tests/base/system_signal004/input | 1 - tests/base/system_signal004/run | 2 +- tests/base/system_time001/expected | 7 ++----- tests/base/system_time001/input | 2 -- tests/base/system_time001/run | 10 +++++----- tests/chez/bitops/expected | 2 -- tests/chez/bitops/input | 2 -- tests/chez/bitops/run | 2 +- tests/chez/casts/expected | 2 -- tests/chez/casts/input | 2 -- tests/chez/casts/run | 2 +- tests/chez/chez001/expected | 4 +--- tests/chez/chez001/input | 2 -- tests/chez/chez001/run | 2 +- tests/chez/chez003/expected | 4 +--- tests/chez/chez003/input | 2 -- tests/chez/chez003/run | 2 +- tests/chez/chez004/expected | 4 +--- tests/chez/chez004/input | 2 -- tests/chez/chez004/run | 2 +- tests/chez/chez007/expected | 4 +--- tests/chez/chez007/input | 2 -- tests/chez/chez007/run | 2 +- tests/chez/chez008/expected | 4 +--- tests/chez/chez008/input | 2 -- tests/chez/chez008/run | 2 +- tests/chez/chez009/expected | 4 +--- tests/chez/chez009/input | 2 -- tests/chez/chez009/run | 2 +- tests/chez/chez010/expected | 4 +--- tests/chez/chez010/input | 2 -- tests/chez/chez010/run | 2 +- tests/chez/chez012/expected | 4 +--- tests/chez/chez012/input | 2 -- tests/chez/chez012/run | 2 +- tests/chez/chez013/expected | 4 +--- tests/chez/chez013/input | 2 -- tests/chez/chez013/run | 2 +- tests/chez/chez014/expected | 4 +--- tests/chez/chez014/input | 2 -- tests/chez/chez014/run | 2 +- tests/chez/chez015/expected | 4 +--- tests/chez/chez015/input | 2 -- tests/chez/chez015/run | 2 +- tests/chez/chez016/expected | 4 +--- tests/chez/chez016/input | 2 -- tests/chez/chez016/run | 2 +- tests/chez/chez017/expected.in | 4 +--- tests/chez/chez017/input | 2 -- tests/chez/chez017/run | 2 +- tests/chez/chez018/expected | 4 +--- tests/chez/chez018/input | 2 -- tests/chez/chez018/run | 2 +- tests/chez/chez020/expected | 4 +--- tests/chez/chez020/input | 2 -- tests/chez/chez020/run | 2 +- tests/chez/chez022/expected | 4 +--- tests/chez/chez022/input | 2 -- tests/chez/chez022/run | 2 +- tests/chez/chez023/expected | 4 +--- tests/chez/chez023/input | 2 -- tests/chez/chez023/run | 2 +- tests/chez/chez024/expected | 4 +--- tests/chez/chez024/input | 2 -- tests/chez/chez024/run | 2 +- tests/chez/chez027/expected | 4 +--- tests/chez/chez027/input | 2 -- tests/chez/chez027/run | 2 +- tests/chez/chez028/expected | 4 +--- tests/chez/chez028/input | 2 -- tests/chez/chez028/run | 2 +- tests/chez/chez029/expected | 4 +--- tests/chez/chez029/input | 2 -- tests/chez/chez029/run | 2 +- tests/chez/chez032/expected | 4 +--- tests/chez/chez032/input | 2 -- tests/chez/chez032/run | 2 +- tests/chez/chez033/expected | 5 +---- tests/chez/chez033/input | 2 -- tests/chez/chez033/run | 2 +- tests/chez/chez034/expected | 4 +--- tests/chez/chez034/input | 2 -- tests/chez/chez034/run | 2 +- tests/chez/constfold/expected | 4 +--- tests/chez/constfold/input | 2 -- tests/chez/constfold/run | 2 +- tests/chez/constfold2/expected | 4 +--- tests/chez/constfold2/input | 2 -- tests/chez/constfold2/run | 2 +- tests/chez/constfold3/expected | 4 +--- tests/chez/constfold3/input | 2 -- tests/chez/constfold3/run | 2 +- tests/chez/memo/expected | 4 +--- tests/chez/memo/input | 2 -- tests/chez/memo/run | 2 +- tests/chez/nat2fin/expected | 4 +--- tests/chez/nat2fin/input | 2 -- tests/chez/nat2fin/run | 2 +- tests/chez/newints/expected | 2 -- tests/chez/newints/input | 2 -- tests/chez/newints/run | 2 +- tests/codegen/enum/expected | 4 +--- tests/codegen/enum/input | 2 -- tests/codegen/enum/run | 2 +- tests/contrib/json_001/expected | 4 +--- tests/contrib/json_001/input | 2 -- tests/contrib/json_001/run | 2 +- tests/contrib/json_002/expected | 4 +--- tests/contrib/json_002/input | 2 -- tests/contrib/json_002/run | 2 +- tests/contrib/json_003/expected | 4 +--- tests/contrib/json_003/input | 2 -- tests/contrib/json_003/run | 2 +- tests/contrib/list_alternating/expected | 4 +--- tests/contrib/list_alternating/input | 2 -- tests/contrib/list_alternating/run | 2 +- tests/contrib/sortedmap_001/expected | 4 +--- tests/contrib/sortedmap_001/input | 2 -- tests/contrib/sortedmap_001/run | 2 +- tests/contrib/system_directory_tree_copyDir/expected | 2 -- tests/contrib/system_directory_tree_copyDir/input | 2 -- tests/contrib/system_directory_tree_copyDir/run | 2 +- tests/idris2/basic055/expected | 4 +--- tests/idris2/basic055/input | 2 -- tests/idris2/basic055/run | 2 +- tests/idris2/basic056/expected | 4 +--- tests/idris2/basic056/input | 2 -- tests/idris2/basic056/run | 2 +- tests/idris2/idiom001/expected | 4 +--- tests/idris2/idiom001/input | 2 -- tests/idris2/idiom001/run | 2 +- tests/idris2/primloop/expected | 4 +--- tests/idris2/primloop/input | 2 -- tests/idris2/primloop/run | 2 +- tests/idris2/reflection014/expected | 4 +--- tests/idris2/reflection014/input | 2 -- tests/idris2/reflection014/run | 2 +- tests/node/bitops/expected | 2 -- tests/node/bitops/input | 2 -- tests/node/bitops/run | 2 +- tests/node/casts/expected | 2 -- tests/node/casts/input | 2 -- tests/node/casts/run | 2 +- tests/node/fastConcat/expected | 4 +--- tests/node/fastConcat/input | 2 -- tests/node/fastConcat/run | 2 +- tests/node/fix1839/expected | 4 +--- tests/node/fix1839/input | 2 -- tests/node/fix1839/run | 2 +- tests/node/idiom001/expected | 2 -- tests/node/idiom001/input | 2 -- tests/node/idiom001/run | 2 +- tests/node/integer_array/expected | 4 +--- tests/node/integer_array/input | 2 -- tests/node/integer_array/run | 2 +- tests/node/memo/expected | 4 +--- tests/node/memo/input | 2 -- tests/node/memo/run | 2 +- tests/node/newints/expected | 2 -- tests/node/newints/input | 2 -- tests/node/newints/run | 2 +- tests/node/node001/expected | 4 +--- tests/node/node001/input | 2 -- tests/node/node001/run | 2 +- tests/node/node003/expected | 4 +--- tests/node/node003/input | 2 -- tests/node/node003/run | 2 +- tests/node/node004/expected | 4 +--- tests/node/node004/input | 2 -- tests/node/node004/run | 2 +- tests/node/node007/expected | 4 +--- tests/node/node007/input | 2 -- tests/node/node007/run | 2 +- tests/node/node008/expected | 4 +--- tests/node/node008/input | 2 -- tests/node/node008/run | 2 +- tests/node/node009/expected | 4 +--- tests/node/node009/input | 2 -- tests/node/node009/run | 2 +- tests/node/node012/expected | 4 +--- tests/node/node012/input | 2 -- tests/node/node012/run | 2 +- tests/node/node014/expected | 2 -- tests/node/node014/input | 2 -- tests/node/node014/run | 2 +- tests/node/node015/expected | 4 +--- tests/node/node015/input | 2 -- tests/node/node015/run | 2 +- tests/node/node017/expected.in | 4 +--- tests/node/node017/input | 2 -- tests/node/node017/run | 2 +- tests/node/node018/expected | 4 +--- tests/node/node018/input | 2 -- tests/node/node018/run | 2 +- tests/node/node020/expected | 4 +--- tests/node/node020/input | 2 -- tests/node/node020/run | 2 +- tests/node/node022/expected | 4 +--- tests/node/node022/input | 2 -- tests/node/node022/run | 2 +- tests/node/node023/expected | 4 +--- tests/node/node023/input | 2 -- tests/node/node023/run | 2 +- tests/node/node024/expected | 4 +--- tests/node/node024/input | 2 -- tests/node/node024/run | 2 +- tests/node/node025/expected | 4 +--- tests/node/node025/input | 2 -- tests/node/node025/run | 2 +- tests/node/node026/expected | 4 +--- tests/node/node026/input | 2 -- tests/node/node026/run | 2 +- tests/node/node027/expected | 2 -- tests/node/node027/input | 2 -- tests/node/node027/run | 2 +- tests/node/perf001/expected | 4 +--- tests/node/perf001/input | 2 -- tests/node/perf001/run | 2 +- tests/node/stringcast/expected | 4 +--- tests/node/stringcast/input | 2 -- tests/node/stringcast/run | 2 +- tests/prelude/bind001/expected | 4 +--- tests/prelude/bind001/input | 2 -- tests/prelude/bind001/run | 2 +- tests/prelude/double001/expected | 4 +--- tests/prelude/double001/input | 2 -- tests/prelude/double001/run | 2 +- 283 files changed, 183 insertions(+), 562 deletions(-) delete mode 100644 tests/allschemes/scheme001/input delete mode 100644 tests/base/control_app001/input delete mode 100644 tests/base/data_bits001/input delete mode 100644 tests/base/data_bits002/input delete mode 100644 tests/base/data_integral/input delete mode 100644 tests/base/data_list001/input delete mode 100644 tests/base/data_string_lines001/input delete mode 100644 tests/base/data_string_unlines001/input delete mode 100644 tests/base/data_vect001/input delete mode 100644 tests/base/system_directory/input delete mode 100644 tests/base/system_errno/input delete mode 100644 tests/base/system_file001/input delete mode 100644 tests/base/system_file_copyFile/input delete mode 100644 tests/base/system_file_fRead/input delete mode 100644 tests/base/system_file_popen/input delete mode 100644 tests/base/system_run/input delete mode 100644 tests/base/system_signal001/input delete mode 100644 tests/base/system_signal002/input delete mode 100644 tests/base/system_signal003/input delete mode 100644 tests/base/system_signal004/input delete mode 100644 tests/base/system_time001/input delete mode 100644 tests/chez/bitops/input delete mode 100644 tests/chez/casts/input delete mode 100644 tests/chez/chez001/input delete mode 100644 tests/chez/chez003/input delete mode 100644 tests/chez/chez004/input delete mode 100644 tests/chez/chez007/input delete mode 100644 tests/chez/chez008/input delete mode 100644 tests/chez/chez009/input delete mode 100644 tests/chez/chez010/input delete mode 100644 tests/chez/chez012/input delete mode 100644 tests/chez/chez013/input delete mode 100644 tests/chez/chez014/input delete mode 100644 tests/chez/chez015/input delete mode 100644 tests/chez/chez016/input delete mode 100644 tests/chez/chez017/input delete mode 100644 tests/chez/chez018/input delete mode 100644 tests/chez/chez020/input delete mode 100644 tests/chez/chez022/input delete mode 100644 tests/chez/chez023/input delete mode 100644 tests/chez/chez024/input delete mode 100644 tests/chez/chez027/input delete mode 100644 tests/chez/chez028/input delete mode 100644 tests/chez/chez029/input delete mode 100644 tests/chez/chez032/input delete mode 100644 tests/chez/chez033/input delete mode 100644 tests/chez/chez034/input delete mode 100644 tests/chez/constfold/input delete mode 100644 tests/chez/constfold2/input delete mode 100644 tests/chez/constfold3/input delete mode 100644 tests/chez/memo/input delete mode 100644 tests/chez/nat2fin/input delete mode 100644 tests/chez/newints/input delete mode 100644 tests/codegen/enum/input delete mode 100644 tests/contrib/json_001/input delete mode 100644 tests/contrib/json_002/input delete mode 100644 tests/contrib/json_003/input delete mode 100644 tests/contrib/list_alternating/input delete mode 100644 tests/contrib/sortedmap_001/input delete mode 100644 tests/contrib/system_directory_tree_copyDir/input delete mode 100644 tests/idris2/basic055/input delete mode 100644 tests/idris2/basic056/input delete mode 100644 tests/idris2/idiom001/input delete mode 100644 tests/idris2/primloop/input delete mode 100644 tests/idris2/reflection014/input delete mode 100644 tests/node/bitops/input delete mode 100644 tests/node/casts/input delete mode 100644 tests/node/fastConcat/input delete mode 100644 tests/node/fix1839/input delete mode 100644 tests/node/idiom001/input delete mode 100644 tests/node/integer_array/input delete mode 100644 tests/node/memo/input delete mode 100644 tests/node/newints/input delete mode 100644 tests/node/node001/input delete mode 100644 tests/node/node003/input delete mode 100644 tests/node/node004/input delete mode 100644 tests/node/node007/input delete mode 100644 tests/node/node008/input delete mode 100644 tests/node/node009/input delete mode 100644 tests/node/node012/input delete mode 100644 tests/node/node014/input delete mode 100644 tests/node/node015/input delete mode 100644 tests/node/node017/input delete mode 100644 tests/node/node018/input delete mode 100644 tests/node/node020/input delete mode 100644 tests/node/node022/input delete mode 100644 tests/node/node023/input delete mode 100644 tests/node/node024/input delete mode 100644 tests/node/node025/input delete mode 100644 tests/node/node026/input delete mode 100644 tests/node/node027/input delete mode 100644 tests/node/perf001/input delete mode 100644 tests/node/stringcast/input delete mode 100644 tests/prelude/bind001/input delete mode 100644 tests/prelude/double001/input diff --git a/tests/allschemes/scheme001/input b/tests/allschemes/scheme001/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/allschemes/scheme001/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/base/control_app001/expected b/tests/base/control_app001/expected index 9b92a3f3a6..9d00464596 100644 --- a/tests/base/control_app001/expected +++ b/tests/base/control_app001/expected @@ -1,3 +1 @@ 1/1: Building TestException (TestException.idr) -TestException> -Bye for now! diff --git a/tests/base/control_app001/input b/tests/base/control_app001/input deleted file mode 100644 index e69de29bb2..0000000000 diff --git a/tests/base/control_app001/run b/tests/base/control_app001/run index 163b71b8ed..16a2df7b09 100755 --- a/tests/base/control_app001/run +++ b/tests/base/control_app001/run @@ -1,3 +1,3 @@ . ../../testutils.sh -idris2 TestException.idr < input +check TestException.idr diff --git a/tests/base/data_bits001/expected b/tests/base/data_bits001/expected index 714459593f..b83e64927e 100644 --- a/tests/base/data_bits001/expected +++ b/tests/base/data_bits001/expected @@ -1,5 +1,4 @@ -1/1: Building BitOps (BitOps.idr) -Main> [0, 1, 2, 4, 8, 16, 32, 64, 127] +[0, 1, 2, 4, 8, 16, 32, 64, 127] [0, 1, 2, 4, 8, 16, 32, 64, 128, 256, 512, 1024, 2048, 4096, 8192, 16384, 32767] [0, 1, 2, 4, 8, 16, 32, 64, 128, 256, 512, 1024, 2048, 4096, 8192, 16384, 32768, 65536, 131072, 262144, 524288, 1048576, 2097152, 4194304, 8388608, 16777216, 33554432, 67108864, 134217728, 268435456, 536870912, 1073741824, 2147483647] [0, 1, 2, 4, 8, 16, 32, 64, 128, 256, 512, 1024, 2048, 4096, 8192, 16384, 32768, 65536, 131072, 262144, 524288, 1048576, 2097152, 4194304, 8388608, 16777216, 33554432, 67108864, 134217728, 268435456, 536870912, 1073741824, 2147483648, 4294967296, 8589934592, 17179869184, 34359738368, 68719476736, 137438953472, 274877906944, 549755813888, 1099511627776, 2199023255552, 4398046511104, 8796093022208, 17592186044416, 35184372088832, 70368744177664, 140737488355328, 281474976710656, 562949953421312, 1125899906842624, 2251799813685248, 4503599627370496, 9007199254740992, 18014398509481984, 36028797018963968, 72057594037927936, 144115188075855872, 288230376151711744, 576460752303423488, 1152921504606846976, 2305843009213693952, 4611686018427387903] @@ -59,4 +58,3 @@ Main> [0, 1, 2, 4, 8, 16, 32, 64, 127] [0, 8, 16] [0, 16, 32] [0, 32, 64, 33] -Main> Bye for now! diff --git a/tests/base/data_bits001/input b/tests/base/data_bits001/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/base/data_bits001/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/base/data_bits001/run b/tests/base/data_bits001/run index 14e998613d..5bb842ecfb 100755 --- a/tests/base/data_bits001/run +++ b/tests/base/data_bits001/run @@ -1,3 +1,3 @@ . ../../testutils.sh -idris2 BitOps.idr < input +run BitOps.idr diff --git a/tests/base/data_bits002/expected b/tests/base/data_bits002/expected index 0e00f29cd0..68daf4da82 100644 --- a/tests/base/data_bits002/expected +++ b/tests/base/data_bits002/expected @@ -1,5 +1,4 @@ -1/1: Building BitRotate (BitRotate.idr) -Main> ------------------------------------------------------------------------ +------------------------------------------------------------------------ -- RotR Bits8 00000000 00000000 00000001 10000000 @@ -413,4 +412,3 @@ Main> ------------------------------------------------------------------------ 1110000000000000000000000000000000000000000000000000000000000000 1100000000000000000000000000000000000000000000000000000000000001 1000000000000000000000000000000000000000000000000000000000000000 0000000000000000000000000000000000000000000000000000000000000001 -Main> Bye for now! diff --git a/tests/base/data_bits002/input b/tests/base/data_bits002/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/base/data_bits002/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/base/data_bits002/run b/tests/base/data_bits002/run index 4397887ef9..b0e6763b21 100755 --- a/tests/base/data_bits002/run +++ b/tests/base/data_bits002/run @@ -1,3 +1,3 @@ . ../../testutils.sh -idris2 BitRotate.idr < input +run BitRotate.idr diff --git a/tests/base/data_integral/expected b/tests/base/data_integral/expected index 39ffccc9e2..f091af7898 100644 --- a/tests/base/data_integral/expected +++ b/tests/base/data_integral/expected @@ -1,5 +1,4 @@ -1/1: Building Integral (Integral.idr) -Main> "Nat Even" +"Nat Even" [True, True, True, True, True] "Nat Odd" [True, True, True, True, True] @@ -7,4 +6,3 @@ Main> "Nat Even" [True, True, True, True, True] "Integral Odd" [True, True, True, True, True, True] -Main> Bye for now! diff --git a/tests/base/data_integral/input b/tests/base/data_integral/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/base/data_integral/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/base/data_integral/run b/tests/base/data_integral/run index 2a3399f8bf..6e0c0a57d1 100644 --- a/tests/base/data_integral/run +++ b/tests/base/data_integral/run @@ -1,3 +1,3 @@ . ../../testutils.sh -idris2 Integral.idr < input +run Integral.idr diff --git a/tests/base/data_list001/expected b/tests/base/data_list001/expected index ca7e4caea3..9523a51059 100644 --- a/tests/base/data_list001/expected +++ b/tests/base/data_list001/expected @@ -1,5 +1,3 @@ -1/1: Building List (List.idr) -Main> [[9, 6, 7, 8], [6, 9, 7, 8], [6, 7, 8, 9], [9, 6], [6, 9], [9]] +[[9, 6, 7, 8], [6, 9, 7, 8], [6, 7, 8, 9], [9, 6], [6, 9], [9]] [[], [4], [3], [4, 5], [3, 5], [3, 4]] [[6], [6, 4], [3, 6], [6, 4, 5], [3, 6, 5], [3, 4, 6]] -Main> Bye for now! diff --git a/tests/base/data_list001/input b/tests/base/data_list001/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/base/data_list001/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/base/data_list001/run b/tests/base/data_list001/run index 8d9753ad57..1744025ece 100755 --- a/tests/base/data_list001/run +++ b/tests/base/data_list001/run @@ -1,3 +1,3 @@ . ../../testutils.sh -idris2 List.idr < input +run List.idr diff --git a/tests/base/data_string_lines001/expected b/tests/base/data_string_lines001/expected index abc1ec7bc6..22a29cb419 100644 --- a/tests/base/data_string_lines001/expected +++ b/tests/base/data_string_lines001/expected @@ -1,5 +1,4 @@ -1/1: Building Lines (Lines.idr) -Main> [] +[] ["ab"] ["ab"] ["ab", "cd"] @@ -7,4 +6,3 @@ Main> [] ["a", "b"] ["a", "b"] ["", "", ""] -Main> Bye for now! diff --git a/tests/base/data_string_lines001/input b/tests/base/data_string_lines001/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/base/data_string_lines001/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/base/data_string_lines001/run b/tests/base/data_string_lines001/run index dbdc30013b..5453ad8925 100755 --- a/tests/base/data_string_lines001/run +++ b/tests/base/data_string_lines001/run @@ -1,3 +1,3 @@ . ../../testutils.sh -idris2 Lines.idr < input +run Lines.idr diff --git a/tests/base/data_string_unlines001/expected b/tests/base/data_string_unlines001/expected index 3bd2d1e7bf..2c855faeaf 100644 --- a/tests/base/data_string_unlines001/expected +++ b/tests/base/data_string_unlines001/expected @@ -1,8 +1,6 @@ -1/1: Building Unlines (Unlines.idr) -Main> "<>" +"<>" "" "" "<>" "" "" -Main> Bye for now! diff --git a/tests/base/data_string_unlines001/input b/tests/base/data_string_unlines001/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/base/data_string_unlines001/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/base/data_string_unlines001/run b/tests/base/data_string_unlines001/run index 802aa5baf3..40d40db87f 100755 --- a/tests/base/data_string_unlines001/run +++ b/tests/base/data_string_unlines001/run @@ -1,3 +1,3 @@ . ../../testutils.sh -idris2 Unlines.idr < input +run Unlines.idr diff --git a/tests/base/data_vect001/expected b/tests/base/data_vect001/expected index 928cda3dc8..9556ec6de6 100644 --- a/tests/base/data_vect001/expected +++ b/tests/base/data_vect001/expected @@ -1,5 +1,4 @@ -1/1: Building Vect (Vect.idr) -Main> [(0 ** []), (1 ** [3]), (5 ** [3, 4, 5, 6, 7])] +[(0 ** []), (1 ** [3]), (5 ** [3, 4, 5, 6, 7])] [(0 ** ([], 3)), (1 ** ([3], 4)), (2 ** ([3, 4], 5))] [(1 ** [2]), (2 ** [6, 2]), (4 ** [120, 24, 6, 2])] [[2, 3, 4], [3, 4], [4], []] @@ -9,4 +8,3 @@ Main> [(0 ** []), (1 ** [3]), (5 ** [3, 4, 5, 6, 7])] [[1, 5], [2, 6], [3, 7], [4, 8]] [] [0, 1, 2] -Main> Bye for now! diff --git a/tests/base/data_vect001/input b/tests/base/data_vect001/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/base/data_vect001/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/base/data_vect001/run b/tests/base/data_vect001/run index fcf96b59e0..a7fd08dd77 100755 --- a/tests/base/data_vect001/run +++ b/tests/base/data_vect001/run @@ -1,3 +1,3 @@ . ../../testutils.sh -idris2 Vect.idr < input +run Vect.idr diff --git a/tests/base/system_directory/expected b/tests/base/system_directory/expected index 2f0f68c043..e69de29bb2 100644 --- a/tests/base/system_directory/expected +++ b/tests/base/system_directory/expected @@ -1,2 +0,0 @@ -1/1: Building ReadDir (ReadDir.idr) -Main> Main> Bye for now! diff --git a/tests/base/system_directory/input b/tests/base/system_directory/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/base/system_directory/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/base/system_directory/run b/tests/base/system_directory/run index e98dbda6f3..cf17aac1cd 100755 --- a/tests/base/system_directory/run +++ b/tests/base/system_directory/run @@ -1,3 +1,3 @@ . ../../testutils.sh -idris2 ReadDir.idr < input +run ReadDir.idr diff --git a/tests/base/system_errno/expected b/tests/base/system_errno/expected index 21ce34f617..e69de29bb2 100644 --- a/tests/base/system_errno/expected +++ b/tests/base/system_errno/expected @@ -1,2 +0,0 @@ -1/1: Building Test (Test.idr) -Main> Main> Bye for now! diff --git a/tests/base/system_errno/input b/tests/base/system_errno/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/base/system_errno/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/base/system_errno/run b/tests/base/system_errno/run index 441b64ba30..ceb556ba04 100755 --- a/tests/base/system_errno/run +++ b/tests/base/system_errno/run @@ -1,3 +1,3 @@ . ../../testutils.sh -idris2 Test.idr < input +run Test.idr diff --git a/tests/base/system_file001/expected b/tests/base/system_file001/expected index 5d81e356ba..4d7364c166 100644 --- a/tests/base/system_file001/expected +++ b/tests/base/system_file001/expected @@ -1,5 +1,4 @@ -1/1: Building ReadFilePage (ReadFilePage.idr) -Main> empty: [] +empty: [] one two @@ -28,8 +27,7 @@ three lines in four five lines total -Main> Bye for now! -Main> empty: [] +empty: [] one two @@ -58,4 +56,3 @@ three lines in four five lines total -Main> Bye for now! diff --git a/tests/base/system_file001/input b/tests/base/system_file001/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/base/system_file001/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/base/system_file001/run b/tests/base/system_file001/run index cb3d7cbdc5..b4f65d4d45 100755 --- a/tests/base/system_file001/run +++ b/tests/base/system_file001/run @@ -2,14 +2,14 @@ # @readFilePage@ uses primitive functions with definitions for both # C (supported by most backends) and Node. -idris2 --cg chez ReadFilePage.idr < input -idris2 --cg node ReadFilePage.idr < input +run --cg chez ReadFilePage.idr +run --cg node ReadFilePage.idr # The following backends failed for reasons unrelated to this test. They can be # uncommented at a future date. # RACKET failed to find the builtin idris support library -# idris2 --cg racket ReadFilePage.idr < input +# run --cg racket ReadFilePage.idr # REFC doesn't support :exec yet -# idris2 --cg refc ReadFilePage.idr < input +# run --cg refc ReadFilePage.idr # GAMBIT hung seemingly indefinitely -# idris2 --cg gambit ReadFilePage.idr < input +# run --cg gambit ReadFilePage.idr diff --git a/tests/base/system_file_copyFile/expected b/tests/base/system_file_copyFile/expected index 450e0e78c3..efd35673a0 100644 --- a/tests/base/system_file_copyFile/expected +++ b/tests/base/system_file_copyFile/expected @@ -1,3 +1 @@ -1/1: Building CopyFile (CopyFile.idr) -Main> Main> Bye for now! AQIDBA== diff --git a/tests/base/system_file_copyFile/input b/tests/base/system_file_copyFile/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/base/system_file_copyFile/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/base/system_file_copyFile/run b/tests/base/system_file_copyFile/run index 50b1e24b10..c324eb67e6 100644 --- a/tests/base/system_file_copyFile/run +++ b/tests/base/system_file_copyFile/run @@ -1,6 +1,6 @@ . ../../testutils.sh -idris2 CopyFile.idr < input +run CopyFile.idr base64 -i dest.bin rm dest.bin diff --git a/tests/base/system_file_fRead/expected b/tests/base/system_file_fRead/expected index 747a1e5e55..04aa614823 100644 --- a/tests/base/system_file_fRead/expected +++ b/tests/base/system_file_fRead/expected @@ -1,3 +1 @@ -1/1: Building ReadFile (ReadFile.idr) -Main> "Hello, world\nLorem ipsum\n" -Main> Bye for now! +"Hello, world\nLorem ipsum\n" diff --git a/tests/base/system_file_fRead/input b/tests/base/system_file_fRead/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/base/system_file_fRead/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/base/system_file_fRead/run b/tests/base/system_file_fRead/run index d302d0f29d..f1c990b352 100644 --- a/tests/base/system_file_fRead/run +++ b/tests/base/system_file_fRead/run @@ -1,3 +1,3 @@ . ../../testutils.sh -idris2 ReadFile.idr < input +run ReadFile.idr diff --git a/tests/base/system_file_popen/expected b/tests/base/system_file_popen/expected index 03342c71d7..7cd8b3d00e 100644 --- a/tests/base/system_file_popen/expected +++ b/tests/base/system_file_popen/expected @@ -1,5 +1,3 @@ -1/1: Building Popen (Popen.idr) -Main> "Hello, world" +"Hello, world" "Hello, $PATH" 17 -Main> Bye for now! diff --git a/tests/base/system_file_popen/input b/tests/base/system_file_popen/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/base/system_file_popen/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/base/system_file_popen/run b/tests/base/system_file_popen/run index a096627a9e..c72547808b 100644 --- a/tests/base/system_file_popen/run +++ b/tests/base/system_file_popen/run @@ -1,3 +1,3 @@ . ../../testutils.sh -idris2 Popen.idr < input +run Popen.idr diff --git a/tests/base/system_run/expected b/tests/base/system_run/expected index 96a4fbd9a3..f8a7b70618 100644 --- a/tests/base/system_run/expected +++ b/tests/base/system_run/expected @@ -1,5 +1,3 @@ -1/1: Building Run (Run.idr) -Main> "Hello, world" +"Hello, world" "Hello, $PATH" ("", 17) -Main> Bye for now! diff --git a/tests/base/system_run/input b/tests/base/system_run/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/base/system_run/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/base/system_run/run b/tests/base/system_run/run index 125d488562..02d226f5a7 100644 --- a/tests/base/system_run/run +++ b/tests/base/system_run/run @@ -1,3 +1,3 @@ . ../../testutils.sh -idris2 Run.idr < input +run Run.idr diff --git a/tests/base/system_signal001/expected b/tests/base/system_signal001/expected index 19472fb6a1..1af30b5d47 100644 --- a/tests/base/system_signal001/expected +++ b/tests/base/system_signal001/expected @@ -1,6 +1,3 @@ -1/1: Building IgnoreSignal (IgnoreSignal.idr) -Main> before +before after done. -Main> -Bye for now! diff --git a/tests/base/system_signal001/input b/tests/base/system_signal001/input deleted file mode 100644 index 6c0a230c0b..0000000000 --- a/tests/base/system_signal001/input +++ /dev/null @@ -1 +0,0 @@ -:exec main diff --git a/tests/base/system_signal001/run b/tests/base/system_signal001/run index 8a317bcdcb..81a62bbe91 100755 --- a/tests/base/system_signal001/run +++ b/tests/base/system_signal001/run @@ -1,3 +1,3 @@ . ../../testutils.sh -idris2 IgnoreSignal.idr < input +run IgnoreSignal.idr diff --git a/tests/base/system_signal002/expected b/tests/base/system_signal002/expected index 3fdc4d9999..1af30b5d47 100644 --- a/tests/base/system_signal002/expected +++ b/tests/base/system_signal002/expected @@ -1,6 +1,3 @@ -1/1: Building HandleSignal (HandleSignal.idr) -Main> before +before after done. -Main> -Bye for now! diff --git a/tests/base/system_signal002/input b/tests/base/system_signal002/input deleted file mode 100644 index 6c0a230c0b..0000000000 --- a/tests/base/system_signal002/input +++ /dev/null @@ -1 +0,0 @@ -:exec main diff --git a/tests/base/system_signal002/run b/tests/base/system_signal002/run index 8ea51c0059..33d7917950 100755 --- a/tests/base/system_signal002/run +++ b/tests/base/system_signal002/run @@ -1,3 +1,3 @@ . ../../testutils.sh -idris2 HandleSignal.idr < input +run HandleSignal.idr diff --git a/tests/base/system_signal003/expected b/tests/base/system_signal003/expected index 4870b854bb..e69de29bb2 100644 --- a/tests/base/system_signal003/expected +++ b/tests/base/system_signal003/expected @@ -1,3 +0,0 @@ -1/1: Building DefaultSignal (DefaultSignal.idr) -Main> Main> -Bye for now! diff --git a/tests/base/system_signal003/input b/tests/base/system_signal003/input deleted file mode 100644 index 6c0a230c0b..0000000000 --- a/tests/base/system_signal003/input +++ /dev/null @@ -1 +0,0 @@ -:exec main diff --git a/tests/base/system_signal003/run b/tests/base/system_signal003/run index e527a9082e..6a37971a33 100755 --- a/tests/base/system_signal003/run +++ b/tests/base/system_signal003/run @@ -1,3 +1,3 @@ . ../../testutils.sh -idris2 DefaultSignal.idr < input 2> /dev/null +run DefaultSignal.idr 2> /dev/null diff --git a/tests/base/system_signal004/expected b/tests/base/system_signal004/expected index d56ae30ce4..1af30b5d47 100644 --- a/tests/base/system_signal004/expected +++ b/tests/base/system_signal004/expected @@ -1,6 +1,3 @@ -1/1: Building HandleManySignals (HandleManySignals.idr) -Main> before +before after done. -Main> -Bye for now! diff --git a/tests/base/system_signal004/input b/tests/base/system_signal004/input deleted file mode 100644 index 6c0a230c0b..0000000000 --- a/tests/base/system_signal004/input +++ /dev/null @@ -1 +0,0 @@ -:exec main diff --git a/tests/base/system_signal004/run b/tests/base/system_signal004/run index 7f4d97c77d..4212ea53b5 100755 --- a/tests/base/system_signal004/run +++ b/tests/base/system_signal004/run @@ -1,3 +1,3 @@ . ../../testutils.sh -idris2 HandleManySignals.idr < input +run HandleManySignals.idr diff --git a/tests/base/system_time001/expected b/tests/base/system_time001/expected index c1eef5de2a..95786ec64e 100644 --- a/tests/base/system_time001/expected +++ b/tests/base/system_time001/expected @@ -1,5 +1,2 @@ -1/1: Building Time (Time.idr) -Main> Retrieved unix timestamp from time function. -Main> Bye for now! -Main> Retrieved unix timestamp from time function. -Main> Bye for now! +Retrieved unix timestamp from time function. +Retrieved unix timestamp from time function. diff --git a/tests/base/system_time001/input b/tests/base/system_time001/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/base/system_time001/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/base/system_time001/run b/tests/base/system_time001/run index 34d15f1121..c99fdbc149 100755 --- a/tests/base/system_time001/run +++ b/tests/base/system_time001/run @@ -2,14 +2,14 @@ # @time@ uses a primitive function with definitions for both # C (supported by most backends) and Javascript (Node & Browsers). -idris2 --cg chez Time.idr < input -idris2 --cg node Time.idr < input +run --cg chez Time.idr +run --cg node Time.idr # The following backends failed for reasons unrelated to this test. They can be # uncommented at a future date. # RACKET failed to find the builtin idris support library -# idris2 --cg racket ReadFilePage.idr < input +# run --cg racket ReadFilePage.idr # REFC doesn't support :exec yet -# idris2 --cg refc ReadFilePage.idr < input +# run --cg refc ReadFilePage.idr # GAMBIT hung seemingly indefinitely -# idris2 --cg gambit ReadFilePage.idr < input +# run --cg gambit ReadFilePage.idr diff --git a/tests/chez/bitops/expected b/tests/chez/bitops/expected index 8848f88cd5..e69de29bb2 100644 --- a/tests/chez/bitops/expected +++ b/tests/chez/bitops/expected @@ -1,2 +0,0 @@ -1/1: Building BitOps (BitOps.idr) -Main> Main> Bye for now! diff --git a/tests/chez/bitops/input b/tests/chez/bitops/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/chez/bitops/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/chez/bitops/run b/tests/chez/bitops/run index 14e998613d..5bb842ecfb 100644 --- a/tests/chez/bitops/run +++ b/tests/chez/bitops/run @@ -1,3 +1,3 @@ . ../../testutils.sh -idris2 BitOps.idr < input +run BitOps.idr diff --git a/tests/chez/casts/expected b/tests/chez/casts/expected index c11635236b..e69de29bb2 100644 --- a/tests/chez/casts/expected +++ b/tests/chez/casts/expected @@ -1,2 +0,0 @@ -1/1: Building Casts (Casts.idr) -Main> Main> Bye for now! diff --git a/tests/chez/casts/input b/tests/chez/casts/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/chez/casts/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/chez/casts/run b/tests/chez/casts/run index 5265730eaa..2cb7ef45a2 100644 --- a/tests/chez/casts/run +++ b/tests/chez/casts/run @@ -1,3 +1,3 @@ . ../../testutils.sh -idris2 Casts.idr < input +run Casts.idr diff --git a/tests/chez/chez001/expected b/tests/chez/chez001/expected index cc9e3bcec6..746a47c3ef 100644 --- a/tests/chez/chez001/expected +++ b/tests/chez/chez001/expected @@ -1,3 +1 @@ -1/1: Building Total (Total.idr) -Main> [1, 2, 2, 4, 3, 6, 4, 8, 5, 10] -Main> Bye for now! +[1, 2, 2, 4, 3, 6, 4, 8, 5, 10] diff --git a/tests/chez/chez001/input b/tests/chez/chez001/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/chez/chez001/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/chez/chez001/run b/tests/chez/chez001/run index c1204f255d..b1513928dc 100755 --- a/tests/chez/chez001/run +++ b/tests/chez/chez001/run @@ -1,3 +1,3 @@ . ../../testutils.sh -idris2 Total.idr < input +run Total.idr diff --git a/tests/chez/chez003/expected b/tests/chez/chez003/expected index d1dbda2627..e37df034b3 100644 --- a/tests/chez/chez003/expected +++ b/tests/chez/chez003/expected @@ -1,6 +1,4 @@ -1/1: Building IORef (IORef.idr) -Main> 94 +94 94 188 188 -Main> Bye for now! diff --git a/tests/chez/chez003/input b/tests/chez/chez003/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/chez/chez003/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/chez/chez003/run b/tests/chez/chez003/run index 58796816d0..61ce178fdd 100755 --- a/tests/chez/chez003/run +++ b/tests/chez/chez003/run @@ -1,3 +1,3 @@ . ../../testutils.sh -idris2 IORef.idr < input +run IORef.idr diff --git a/tests/chez/chez004/expected b/tests/chez/chez004/expected index 8a952b321e..ca7483f2d2 100644 --- a/tests/chez/chez004/expected +++ b/tests/chez/chez004/expected @@ -1,5 +1,4 @@ -1/1: Building Buffer (Buffer.idr) -Main> 100 +100 94 94.42 bytes: 27 @@ -21,4 +20,3 @@ total size = 20 00 00 00 00 00 00 00 00 00 00 00 00 00 00 00 00 ................ total size = 80 -Main> Bye for now! diff --git a/tests/chez/chez004/input b/tests/chez/chez004/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/chez/chez004/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/chez/chez004/run b/tests/chez/chez004/run index 87a9948fb4..f99a1f7750 100755 --- a/tests/chez/chez004/run +++ b/tests/chez/chez004/run @@ -1,4 +1,4 @@ . ../../testutils.sh -idris2 -p contrib Buffer.idr < input +run -p contrib Buffer.idr rm -rf build test.buf diff --git a/tests/chez/chez007/expected b/tests/chez/chez007/expected index 3361839df9..70719619ce 100644 --- a/tests/chez/chez007/expected +++ b/tests/chez/chez007/expected @@ -1,5 +1,3 @@ -1/1: Building TypeCase (TypeCase.idr) -Main> "Function from Nat to Nat" +"Function from Nat to Nat" "Function from Nat to Vector of 0 Int" "Function on Type" -Main> Bye for now! diff --git a/tests/chez/chez007/input b/tests/chez/chez007/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/chez/chez007/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/chez/chez007/run b/tests/chez/chez007/run index 786567d611..8a1cb2e944 100755 --- a/tests/chez/chez007/run +++ b/tests/chez/chez007/run @@ -1,3 +1,3 @@ . ../../testutils.sh -idris2 TypeCase.idr < input +run TypeCase.idr diff --git a/tests/chez/chez008/expected b/tests/chez/chez008/expected index 9500505475..e8183f05f5 100644 --- a/tests/chez/chez008/expected +++ b/tests/chez/chez008/expected @@ -1,5 +1,3 @@ -1/1: Building Nat (Nat.idr) -Main> 1 1 1 -Main> Bye for now! +1 diff --git a/tests/chez/chez008/input b/tests/chez/chez008/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/chez/chez008/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/chez/chez008/run b/tests/chez/chez008/run index 3c4ab0b826..bb2e1d58b2 100755 --- a/tests/chez/chez008/run +++ b/tests/chez/chez008/run @@ -1,3 +1,3 @@ . ../../testutils.sh -idris2 Nat.idr < input +run Nat.idr diff --git a/tests/chez/chez009/expected b/tests/chez/chez009/expected index 759037d743..42d27acd6e 100644 --- a/tests/chez/chez009/expected +++ b/tests/chez/chez009/expected @@ -1,4 +1,2 @@ -1/1: Building uni (uni.idr) -Main> 42 +42 ällo -Main> Bye for now! diff --git a/tests/chez/chez009/input b/tests/chez/chez009/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/chez/chez009/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/chez/chez009/run b/tests/chez/chez009/run index 36867b3057..34bb6f6e73 100755 --- a/tests/chez/chez009/run +++ b/tests/chez/chez009/run @@ -1,3 +1,3 @@ . ../../testutils.sh -idris2 uni.idr < input +run uni.idr diff --git a/tests/chez/chez010/expected b/tests/chez/chez010/expected index dfa5665b0b..9a17052b79 100644 --- a/tests/chez/chez010/expected +++ b/tests/chez/chez010/expected @@ -1,5 +1,4 @@ -1/1: Building CB (CB.idr) -Main> 9 +9 Callback coming In callback 24 @@ -8,4 +7,3 @@ In callback with (1, 2) 3 9 'k' -Main> Bye for now! diff --git a/tests/chez/chez010/input b/tests/chez/chez010/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/chez/chez010/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/chez/chez010/run b/tests/chez/chez010/run index edca4eb608..ff08c6191f 100755 --- a/tests/chez/chez010/run +++ b/tests/chez/chez010/run @@ -11,5 +11,5 @@ case $(uname -s) in esac ${MAKE} all > /dev/null -idris2 CB.idr < input +run CB.idr ${MAKE} clean > /dev/null diff --git a/tests/chez/chez012/expected b/tests/chez/chez012/expected index defaf6349b..8ef81328a3 100644 --- a/tests/chez/chez012/expected +++ b/tests/chez/chez012/expected @@ -1,4 +1,2 @@ -1/1: Building array (array.idr) -Main> [Nothing, Nothing, Nothing, Nothing, Nothing, Nothing, Nothing, Nothing, Nothing, Nothing, Just "Hello", Just "World", Nothing, Nothing, Nothing, Nothing, Nothing, Nothing, Nothing, Nothing] +[Nothing, Nothing, Nothing, Nothing, Nothing, Nothing, Nothing, Nothing, Nothing, Nothing, Just "Hello", Just "World", Nothing, Nothing, Nothing, Nothing, Nothing, Nothing, Nothing, Nothing] [Just 1, Just 2, Just 3, Just 4, Just 5] -Main> Bye for now! diff --git a/tests/chez/chez012/input b/tests/chez/chez012/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/chez/chez012/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/chez/chez012/run b/tests/chez/chez012/run index 7f12edfdfe..a69e093bbb 100755 --- a/tests/chez/chez012/run +++ b/tests/chez/chez012/run @@ -1,3 +1,3 @@ . ../../testutils.sh -idris2 array.idr < input +run array.idr diff --git a/tests/chez/chez013/expected b/tests/chez/chez013/expected index 86daab1c1a..ce62263bb6 100644 --- a/tests/chez/chez013/expected +++ b/tests/chez/chez013/expected @@ -1,5 +1,3 @@ -1/1: Building Struct (Struct.idr) -Main> Made it! +Made it! (40, 30) "Here": (40, 30) -Main> Bye for now! diff --git a/tests/chez/chez013/input b/tests/chez/chez013/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/chez/chez013/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/chez/chez013/run b/tests/chez/chez013/run index 2199a3b818..5f02e3e6c2 100755 --- a/tests/chez/chez013/run +++ b/tests/chez/chez013/run @@ -11,5 +11,5 @@ case $(uname -s) in esac ${MAKE} all > /dev/null -idris2 Struct.idr < input +run Struct.idr ${MAKE} clean > /dev/null diff --git a/tests/chez/chez014/expected b/tests/chez/chez014/expected index 5c4e82d4ab..d7c97490fb 100644 --- a/tests/chez/chez014/expected +++ b/tests/chez/chez014/expected @@ -1,6 +1,4 @@ -1/1: Building Echo (Echo.idr) -Main> Received: hello world from a ipv4 socket! +Received: hello world from a ipv4 socket! Received: echo: hello world from a ipv4 socket! Received: hello world from a unix socket! Received: echo: hello world from a unix socket! -Main> Bye for now! diff --git a/tests/chez/chez014/input b/tests/chez/chez014/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/chez/chez014/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/chez/chez014/run b/tests/chez/chez014/run index c2e5aa740d..908c5d280e 100755 --- a/tests/chez/chez014/run +++ b/tests/chez/chez014/run @@ -1,3 +1,3 @@ . ../../testutils.sh -idris2 -p network Echo.idr < input +run -p network Echo.idr diff --git a/tests/chez/chez015/expected b/tests/chez/chez015/expected index 478fb8fa70..72800451f9 100644 --- a/tests/chez/chez015/expected +++ b/tests/chez/chez015/expected @@ -1,4 +1,2 @@ -1/1: Building Numbers (Numbers.idr) -Main> [3518437212345678901234567890560, 1537557061795061679839506167983751, 3518437212345678901234567889686, 8051343735344802977653473432, 339] +[3518437212345678901234567890560, 1537557061795061679839506167983751, 3518437212345678901234567889686, 8051343735344802977653473432, 339] [-9223372036854775439, 9223372036854772792, 9223372036854775423, 24465177816590917, 91] -Main> Bye for now! diff --git a/tests/chez/chez015/input b/tests/chez/chez015/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/chez/chez015/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/chez/chez015/run b/tests/chez/chez015/run index 3f99939085..b2b292332d 100755 --- a/tests/chez/chez015/run +++ b/tests/chez/chez015/run @@ -1,3 +1,3 @@ . ../../testutils.sh -idris2 Numbers.idr < input +run Numbers.idr diff --git a/tests/chez/chez016/expected b/tests/chez/chez016/expected index 5fba03089d..88a0340072 100644 --- a/tests/chez/chez016/expected +++ b/tests/chez/chez016/expected @@ -1,3 +1 @@ -1/1: Building Main (Main.idr) -Main> Running Chez program located in folder with spaces -Main> Bye for now! +Running Chez program located in folder with spaces diff --git a/tests/chez/chez016/input b/tests/chez/chez016/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/chez/chez016/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/chez/chez016/run b/tests/chez/chez016/run index 003a15d3dd..57da701b09 100755 --- a/tests/chez/chez016/run +++ b/tests/chez/chez016/run @@ -2,4 +2,4 @@ cd "folder with spaces" || exit -idris2 Main.idr < ../input +run Main.idr diff --git a/tests/chez/chez017/expected.in b/tests/chez/chez017/expected.in index 3096c91912..916d21080b 100644 --- a/tests/chez/chez017/expected.in +++ b/tests/chez/chez017/expected.in @@ -1,7 +1,5 @@ -1/1: Building dir (dir.idr) -Main> File Exists +File Exists False True Just "__PWD__testdir" -Main> Bye for now! hello diff --git a/tests/chez/chez017/input b/tests/chez/chez017/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/chez/chez017/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/chez/chez017/run b/tests/chez/chez017/run index a657ed7f39..f15d224d6c 100755 --- a/tests/chez/chez017/run +++ b/tests/chez/chez017/run @@ -1,6 +1,6 @@ . ../../testutils.sh ./gen_expected.sh -idris2 dir.idr < input +run dir.idr cat testdir/test.txt rm -rf testdir diff --git a/tests/chez/chez018/expected b/tests/chez/chez018/expected index 0600e9e6bc..561ce89710 100644 --- a/tests/chez/chez018/expected +++ b/tests/chez/chez018/expected @@ -1,6 +1,4 @@ -1/1: Building File (File.idr) -Main> test test +test test unfinished lineabc def File Not Found -Main> Bye for now! diff --git a/tests/chez/chez018/input b/tests/chez/chez018/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/chez/chez018/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/chez/chez018/run b/tests/chez/chez018/run index 2017333d63..901015ecae 100755 --- a/tests/chez/chez018/run +++ b/tests/chez/chez018/run @@ -1,5 +1,5 @@ . ../../testutils.sh -idris2 File.idr < input +run File.idr rm -rf build testout.txt diff --git a/tests/chez/chez020/expected b/tests/chez/chez020/expected index 0bdb9d17d3..580712d465 100644 --- a/tests/chez/chez020/expected +++ b/tests/chez/chez020/expected @@ -1,5 +1,3 @@ -1/1: Building Popen (Popen.idr) -Main> opened +opened closed Idris 2 -Main> Bye for now! diff --git a/tests/chez/chez020/input b/tests/chez/chez020/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/chez/chez020/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/chez/chez020/run b/tests/chez/chez020/run index 0f522a9265..7a47b290f1 100644 --- a/tests/chez/chez020/run +++ b/tests/chez/chez020/run @@ -1,3 +1,3 @@ . ../../testutils.sh -POPEN_CMD="$idris2 --version" idris2 Popen.idr < input +POPEN_CMD="$idris2 --version" run Popen.idr diff --git a/tests/chez/chez022/expected b/tests/chez/chez022/expected index c8c910f6c9..4e7ab721f6 100644 --- a/tests/chez/chez022/expected +++ b/tests/chez/chez022/expected @@ -1,9 +1,7 @@ -1/1: Building usealloc (usealloc.idr) -Main> Hello +Hello Hello Done Free X Freeing 0 Hello Free Y Freeing 1 Hello -Main> Bye for now! diff --git a/tests/chez/chez022/input b/tests/chez/chez022/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/chez/chez022/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/chez/chez022/run b/tests/chez/chez022/run index 45ad78bf70..1d90c9ff9c 100755 --- a/tests/chez/chez022/run +++ b/tests/chez/chez022/run @@ -11,6 +11,6 @@ case $(uname -s) in esac ${MAKE} all > /dev/null -idris2 usealloc.idr < input +run usealloc.idr . ../../testutils.sh ${MAKE} clean > /dev/null diff --git a/tests/chez/chez023/expected b/tests/chez/chez023/expected index 8f255d5735..b44c01edf0 100644 --- a/tests/chez/chez023/expected +++ b/tests/chez/chez023/expected @@ -1,7 +1,5 @@ -1/1: Building File (File.idr) -Main> Hello +Hello 'I' dris! No exceptions occurred -Main> Bye for now! diff --git a/tests/chez/chez023/input b/tests/chez/chez023/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/chez/chez023/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/chez/chez023/run b/tests/chez/chez023/run index c70689a7a7..cee160bcc9 100644 --- a/tests/chez/chez023/run +++ b/tests/chez/chez023/run @@ -1,5 +1,5 @@ . ../../testutils.sh -idris2 File.idr < input +run File.idr rm -rf build test.txt diff --git a/tests/chez/chez024/expected b/tests/chez/chez024/expected index 8602153ff2..7a67560649 100644 --- a/tests/chez/chez024/expected +++ b/tests/chez/chez024/expected @@ -1,5 +1,4 @@ -1/1: Building Envy (Envy.idr) -Main> True +True HI True HI @@ -7,4 +6,3 @@ True EH True Nothing there -Main> Bye for now! diff --git a/tests/chez/chez024/input b/tests/chez/chez024/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/chez/chez024/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/chez/chez024/run b/tests/chez/chez024/run index d718d5caee..973bba48a7 100644 --- a/tests/chez/chez024/run +++ b/tests/chez/chez024/run @@ -1,3 +1,3 @@ . ../../testutils.sh -idris2 Envy.idr < input +run Envy.idr diff --git a/tests/chez/chez027/expected b/tests/chez/chez027/expected index bc72c01bc1..c1df5a00d7 100644 --- a/tests/chez/chez027/expected +++ b/tests/chez/chez027/expected @@ -1,5 +1,4 @@ -1/1: Building StringParser (StringParser.idr) -Main> hiya +hiya 2 Parse failed at position 0-0: could not satisfy predicate Parse failed at position 0-0: Not good @@ -18,4 +17,3 @@ False ['a', 'b', 'c', 'd'] () Parse failed at position 0-0: Purposefully changed OK to Fail -Main> Bye for now! diff --git a/tests/chez/chez027/input b/tests/chez/chez027/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/chez/chez027/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/chez/chez027/run b/tests/chez/chez027/run index 99b2ea84f7..500c1d920f 100644 --- a/tests/chez/chez027/run +++ b/tests/chez/chez027/run @@ -1,3 +1,3 @@ . ../../testutils.sh -idris2 -p contrib StringParser.idr < input +run -p contrib StringParser.idr diff --git a/tests/chez/chez028/expected b/tests/chez/chez028/expected index f9f44d722c..e09d96bab0 100644 --- a/tests/chez/chez028/expected +++ b/tests/chez/chez028/expected @@ -1,7 +1,5 @@ -1/1: Building ExpressionParser (ExpressionParser.idr) -Main> 5678 +5678 -3 262145 10 7 -Main> Bye for now! diff --git a/tests/chez/chez028/input b/tests/chez/chez028/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/chez/chez028/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/chez/chez028/run b/tests/chez/chez028/run index 7adaecccc2..f8a9bea0c2 100644 --- a/tests/chez/chez028/run +++ b/tests/chez/chez028/run @@ -1,3 +1,3 @@ . ../../testutils.sh -idris2 -p contrib ExpressionParser.idr < input +run -p contrib ExpressionParser.idr diff --git a/tests/chez/chez029/expected b/tests/chez/chez029/expected index aee2314b79..f9260cb774 100644 --- a/tests/chez/chez029/expected +++ b/tests/chez/chez029/expected @@ -1,5 +1,4 @@ -1/1: Building BitCasts (BitCasts.idr) -Main> ["123", "123", "123", "123", "123"] +["123", "123", "123", "123", "123"] ["1234", "1234", "1234", "1234"] ["1234567", "1234567", "1234567"] ["134", "134", "134", "134"] @@ -8,4 +7,3 @@ Main> ["123", "123", "123", "123", "123"] ["134", "134"] ["134"] ["237", "65517", "4294967277", "18446744073709551597"] -Main> Bye for now! diff --git a/tests/chez/chez029/input b/tests/chez/chez029/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/chez/chez029/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/chez/chez029/run b/tests/chez/chez029/run index e4a0b380e2..9b07e12aef 100644 --- a/tests/chez/chez029/run +++ b/tests/chez/chez029/run @@ -1,3 +1,3 @@ . ../../testutils.sh -idris2 BitCasts.idr < input +run BitCasts.idr diff --git a/tests/chez/chez032/expected b/tests/chez/chez032/expected index ea5c96346f..477719b4f7 100644 --- a/tests/chez/chez032/expected +++ b/tests/chez/chez032/expected @@ -1,5 +1,4 @@ -1/1: Building BitOps (BitOps.idr) -Main> -9223372036854775808 +-9223372036854775808 [0, 1, 2, 4, 8, 16, 32, 64, 127] [0, 1, 2, 4, 8, 16, 32, 64, 128, 256, 512, 1024, 2048, 4096, 8192, 16384, 32767] @@ -40,4 +39,3 @@ Main> -9223372036854775808 [340282366920938463463374607431768211444, 11, 10, 9, 15, 0] [9223372036854775796, 11, 10, 9, 15, 0] [-9223372036854775798, -11, -12, -9, -15, -2, 9223372036854775797, 10, 11, 9, 0] -Main> Bye for now! diff --git a/tests/chez/chez032/input b/tests/chez/chez032/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/chez/chez032/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/chez/chez032/run b/tests/chez/chez032/run index 14e998613d..5bb842ecfb 100644 --- a/tests/chez/chez032/run +++ b/tests/chez/chez032/run @@ -1,3 +1,3 @@ . ../../testutils.sh -idris2 BitOps.idr < input +run BitOps.idr diff --git a/tests/chez/chez033/expected b/tests/chez/chez033/expected index 298c4f1c9d..cc32087ec6 100644 --- a/tests/chez/chez033/expected +++ b/tests/chez/chez033/expected @@ -1,4 +1 @@ -1/2: Building Mod (Mod.idr) -2/2: Building Main (Main.idr) -Main> 10001 -Main> Bye for now! +10001 diff --git a/tests/chez/chez033/input b/tests/chez/chez033/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/chez/chez033/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/chez/chez033/run b/tests/chez/chez033/run index e6679e9548..33e48535bb 100644 --- a/tests/chez/chez033/run +++ b/tests/chez/chez033/run @@ -1,3 +1,3 @@ . ../../testutils.sh -idris2 Main.idr --inc chez < input +run Main.idr --inc chez diff --git a/tests/chez/chez034/expected b/tests/chez/chez034/expected index efcff77081..918b88786d 100644 --- a/tests/chez/chez034/expected +++ b/tests/chez/chez034/expected @@ -1,9 +1,7 @@ -1/1: Building ThreadData (ThreadData.idr) -Main> parent data initialized to: 13 +parent data initialized to: 13 parent data now: "hello" child data: "hello" child data now: 17 parent data still: "hello" child exiting parent exiting -Main> Bye for now! diff --git a/tests/chez/chez034/input b/tests/chez/chez034/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/chez/chez034/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/chez/chez034/run b/tests/chez/chez034/run index 990c8486c0..466124edcf 100644 --- a/tests/chez/chez034/run +++ b/tests/chez/chez034/run @@ -1,3 +1,3 @@ . ../../testutils.sh -idris2 ThreadData.idr < input +run ThreadData.idr diff --git a/tests/chez/constfold/expected b/tests/chez/constfold/expected index 049d278288..6bd61a4e24 100644 --- a/tests/chez/constfold/expected +++ b/tests/chez/constfold/expected @@ -1,3 +1 @@ -1/1: Building Check (Check.idr) -Main> Constant expression correctly folded -Main> Bye for now! +Constant expression correctly folded diff --git a/tests/chez/constfold/input b/tests/chez/constfold/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/chez/constfold/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/chez/constfold/run b/tests/chez/constfold/run index 8c605a6262..7c4cb71293 100644 --- a/tests/chez/constfold/run +++ b/tests/chez/constfold/run @@ -1,4 +1,4 @@ . ../../testutils.sh idris2 --quiet -o fold Fold.idr -idris2 Check.idr < input +run Check.idr diff --git a/tests/chez/constfold2/expected b/tests/chez/constfold2/expected index 049d278288..6bd61a4e24 100644 --- a/tests/chez/constfold2/expected +++ b/tests/chez/constfold2/expected @@ -1,3 +1 @@ -1/1: Building Check (Check.idr) -Main> Constant expression correctly folded -Main> Bye for now! +Constant expression correctly folded diff --git a/tests/chez/constfold2/input b/tests/chez/constfold2/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/chez/constfold2/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/chez/constfold2/run b/tests/chez/constfold2/run index 8c605a6262..7c4cb71293 100644 --- a/tests/chez/constfold2/run +++ b/tests/chez/constfold2/run @@ -1,4 +1,4 @@ . ../../testutils.sh idris2 --quiet -o fold Fold.idr -idris2 Check.idr < input +run Check.idr diff --git a/tests/chez/constfold3/expected b/tests/chez/constfold3/expected index 049d278288..6bd61a4e24 100644 --- a/tests/chez/constfold3/expected +++ b/tests/chez/constfold3/expected @@ -1,3 +1 @@ -1/1: Building Check (Check.idr) -Main> Constant expression correctly folded -Main> Bye for now! +Constant expression correctly folded diff --git a/tests/chez/constfold3/input b/tests/chez/constfold3/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/chez/constfold3/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/chez/constfold3/run b/tests/chez/constfold3/run index 8c605a6262..7c4cb71293 100644 --- a/tests/chez/constfold3/run +++ b/tests/chez/constfold3/run @@ -1,4 +1,4 @@ . ../../testutils.sh idris2 --quiet -o fold Fold.idr -idris2 Check.idr < input +run Check.idr diff --git a/tests/chez/memo/expected b/tests/chez/memo/expected index 341a02ebf0..493d3b8b14 100644 --- a/tests/chez/memo/expected +++ b/tests/chez/memo/expected @@ -1,3 +1 @@ -1/1: Building Memo (Memo.idr) -Main> 100000000000000000000 -Main> Bye for now! +100000000000000000000 diff --git a/tests/chez/memo/input b/tests/chez/memo/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/chez/memo/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/chez/memo/run b/tests/chez/memo/run index 2a10ada75f..f8843a026d 100644 --- a/tests/chez/memo/run +++ b/tests/chez/memo/run @@ -1,3 +1,3 @@ . ../../testutils.sh -idris2 Memo.idr < input +run Memo.idr diff --git a/tests/chez/nat2fin/expected b/tests/chez/nat2fin/expected index 5d19b98b40..721a48ef6d 100644 --- a/tests/chez/nat2fin/expected +++ b/tests/chez/nat2fin/expected @@ -1,3 +1 @@ -1/1: Building Check (Check.idr) -Main> natToFinLt optimized away -Main> Bye for now! +natToFinLt optimized away diff --git a/tests/chez/nat2fin/input b/tests/chez/nat2fin/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/chez/nat2fin/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/chez/nat2fin/run b/tests/chez/nat2fin/run index 5c619e2494..65c495ceb8 100644 --- a/tests/chez/nat2fin/run +++ b/tests/chez/nat2fin/run @@ -1,4 +1,4 @@ . ../../testutils.sh idris2 --quiet -o test Test.idr -idris2 Check.idr < input +run Check.idr diff --git a/tests/chez/newints/expected b/tests/chez/newints/expected index 80ec897d1a..e69de29bb2 100644 --- a/tests/chez/newints/expected +++ b/tests/chez/newints/expected @@ -1,2 +0,0 @@ -1/1: Building IntOps (IntOps.idr) -Main> Main> Bye for now! diff --git a/tests/chez/newints/input b/tests/chez/newints/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/chez/newints/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/chez/newints/run b/tests/chez/newints/run index de239c9f52..977371e71e 100644 --- a/tests/chez/newints/run +++ b/tests/chez/newints/run @@ -1,3 +1,3 @@ . ../../testutils.sh -idris2 IntOps.idr < input +run IntOps.idr diff --git a/tests/codegen/enum/expected b/tests/codegen/enum/expected index c5cf3dab1c..2076ed42d6 100644 --- a/tests/codegen/enum/expected +++ b/tests/codegen/enum/expected @@ -1,3 +1 @@ -1/1: Building Check (Check.idr) -Main> Enum conversion optimized away -Main> Bye for now! +Enum conversion optimized away diff --git a/tests/codegen/enum/input b/tests/codegen/enum/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/codegen/enum/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/codegen/enum/run b/tests/codegen/enum/run index 2cdf4dd5c3..887c6b4052 100644 --- a/tests/codegen/enum/run +++ b/tests/codegen/enum/run @@ -1,4 +1,4 @@ . ../../testutils.sh idris2 --quiet -o enum Enum.idr -idris2 Check.idr < input +run Check.idr diff --git a/tests/contrib/json_001/expected b/tests/contrib/json_001/expected index e266789b07..fe51488c70 100644 --- a/tests/contrib/json_001/expected +++ b/tests/contrib/json_001/expected @@ -1,3 +1 @@ -1/1: Building CharEncoding (CharEncoding.idr) -Main> [] -Main> Bye for now! +[] diff --git a/tests/contrib/json_001/input b/tests/contrib/json_001/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/contrib/json_001/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/contrib/json_001/run b/tests/contrib/json_001/run index 592466380e..5880029aa5 100755 --- a/tests/contrib/json_001/run +++ b/tests/contrib/json_001/run @@ -1,3 +1,3 @@ . ../../testutils.sh -idris2 -p contrib CharEncoding.idr < input +run -p contrib CharEncoding.idr diff --git a/tests/contrib/json_002/expected b/tests/contrib/json_002/expected index af98e8c2dc..0256160d08 100644 --- a/tests/contrib/json_002/expected +++ b/tests/contrib/json_002/expected @@ -1,5 +1,4 @@ -1/1: Building ShowJSON (ShowJSON.idr) -Main> {"a":null,"b":true,"c":1.0,"d":"Hello, world","e":[null,"Lorem ipsum"],"f":{"key":"value"}} +{"a":null,"b":true,"c":1.0,"d":"Hello, world","e":[null,"Lorem ipsum"],"f":{"key":"value"}} JObject [("a", JNull), ("b", JBoolean True), ("c", JNumber 1.0), ("d", JString "Hello, world"), ("e", JArray [JNull, JString "Lorem ipsum"]), ("f", JObject [("key", JString "value")])] { "a": null, @@ -14,4 +13,3 @@ JObject [("a", JNull), ("b", JBoolean True), ("c", JNumber 1.0), ("d", JString " "key": "value" } } -Main> Bye for now! diff --git a/tests/contrib/json_002/input b/tests/contrib/json_002/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/contrib/json_002/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/contrib/json_002/run b/tests/contrib/json_002/run index 49902be073..52864ced24 100644 --- a/tests/contrib/json_002/run +++ b/tests/contrib/json_002/run @@ -1,3 +1,3 @@ . ../../testutils.sh -idris2 -p contrib ShowJSON.idr < input +run -p contrib ShowJSON.idr diff --git a/tests/contrib/json_003/expected b/tests/contrib/json_003/expected index d9ada41319..f879442730 100644 --- a/tests/contrib/json_003/expected +++ b/tests/contrib/json_003/expected @@ -1,3 +1 @@ -1/1: Building CastJSON (CastJSON.idr) -Main> {"a":null,"b":true,"c":1.0,"d":"Hello, world","e":["Lorem","ipsum"]} -Main> Bye for now! +{"a":null,"b":true,"c":1.0,"d":"Hello, world","e":["Lorem","ipsum"]} diff --git a/tests/contrib/json_003/input b/tests/contrib/json_003/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/contrib/json_003/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/contrib/json_003/run b/tests/contrib/json_003/run index fe94f3e04d..291d3fd279 100644 --- a/tests/contrib/json_003/run +++ b/tests/contrib/json_003/run @@ -1,3 +1,3 @@ . ../../testutils.sh -idris2 -p contrib CastJSON.idr < input +run -p contrib CastJSON.idr diff --git a/tests/contrib/list_alternating/expected b/tests/contrib/list_alternating/expected index 2eb07eaf49..8b4e5bbca0 100644 --- a/tests/contrib/list_alternating/expected +++ b/tests/contrib/list_alternating/expected @@ -1,5 +1,4 @@ -1/1: Building AlternatingList (AlternatingList.idr) -Main> [1.0, "Hello", 2.0, "world", 3.0] +[1.0, "Hello", 2.0, "world", 3.0] True False True @@ -32,4 +31,3 @@ True [1.0, 2.0, 3.0] ["Hello", "world"] ["1.0", "Hello", "2.0", "world", "3.0"] -Main> Bye for now! diff --git a/tests/contrib/list_alternating/input b/tests/contrib/list_alternating/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/contrib/list_alternating/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/contrib/list_alternating/run b/tests/contrib/list_alternating/run index ed570b6195..6215febb95 100644 --- a/tests/contrib/list_alternating/run +++ b/tests/contrib/list_alternating/run @@ -1,3 +1,3 @@ . ../../testutils.sh -idris2 -p contrib AlternatingList.idr < input +run -p contrib AlternatingList.idr diff --git a/tests/contrib/sortedmap_001/expected b/tests/contrib/sortedmap_001/expected index f31645a288..fec12764d3 100644 --- a/tests/contrib/sortedmap_001/expected +++ b/tests/contrib/sortedmap_001/expected @@ -1,5 +1,3 @@ -1/1: Building SortedMapTest (SortedMapTest.idr) -Main> Nothing Nothing Nothing Nothing @@ -16,4 +14,4 @@ Nothing Nothing Nothing Nothing -Main> Bye for now! +Nothing diff --git a/tests/contrib/sortedmap_001/input b/tests/contrib/sortedmap_001/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/contrib/sortedmap_001/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/contrib/sortedmap_001/run b/tests/contrib/sortedmap_001/run index cc69869397..69738df4bf 100644 --- a/tests/contrib/sortedmap_001/run +++ b/tests/contrib/sortedmap_001/run @@ -1,3 +1,3 @@ . ../../testutils.sh -idris2 -p contrib SortedMapTest.idr < input +run -p contrib SortedMapTest.idr diff --git a/tests/contrib/system_directory_tree_copyDir/expected b/tests/contrib/system_directory_tree_copyDir/expected index 87fe57772c..ad562d0eb1 100644 --- a/tests/contrib/system_directory_tree_copyDir/expected +++ b/tests/contrib/system_directory_tree_copyDir/expected @@ -1,5 +1,3 @@ -1/1: Building CopyDir (CopyDir.idr) -Main> Main> Bye for now! nestedDir source.bin diff --git a/tests/contrib/system_directory_tree_copyDir/input b/tests/contrib/system_directory_tree_copyDir/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/contrib/system_directory_tree_copyDir/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/contrib/system_directory_tree_copyDir/run b/tests/contrib/system_directory_tree_copyDir/run index 557e7f40fa..3828044f33 100644 --- a/tests/contrib/system_directory_tree_copyDir/run +++ b/tests/contrib/system_directory_tree_copyDir/run @@ -1,6 +1,6 @@ . ../../testutils.sh -idris2 -p contrib CopyDir.idr < input +run -p contrib CopyDir.idr ls -R resultDir | sed '/^resultDir:$/d' base64 -i resultDir/source.bin cat resultDir/nestedDir/anotherFile.txt diff --git a/tests/idris2/basic055/expected b/tests/idris2/basic055/expected index ea5c96346f..477719b4f7 100644 --- a/tests/idris2/basic055/expected +++ b/tests/idris2/basic055/expected @@ -1,5 +1,4 @@ -1/1: Building BitOps (BitOps.idr) -Main> -9223372036854775808 +-9223372036854775808 [0, 1, 2, 4, 8, 16, 32, 64, 127] [0, 1, 2, 4, 8, 16, 32, 64, 128, 256, 512, 1024, 2048, 4096, 8192, 16384, 32767] @@ -40,4 +39,3 @@ Main> -9223372036854775808 [340282366920938463463374607431768211444, 11, 10, 9, 15, 0] [9223372036854775796, 11, 10, 9, 15, 0] [-9223372036854775798, -11, -12, -9, -15, -2, 9223372036854775797, 10, 11, 9, 0] -Main> Bye for now! diff --git a/tests/idris2/basic055/input b/tests/idris2/basic055/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/idris2/basic055/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/idris2/basic055/run b/tests/idris2/basic055/run index 14e998613d..5bb842ecfb 100644 --- a/tests/idris2/basic055/run +++ b/tests/idris2/basic055/run @@ -1,3 +1,3 @@ . ../../testutils.sh -idris2 BitOps.idr < input +run BitOps.idr diff --git a/tests/idris2/basic056/expected b/tests/idris2/basic056/expected index aff24fb1ea..296ba3d0d0 100644 --- a/tests/idris2/basic056/expected +++ b/tests/idris2/basic056/expected @@ -1,4 +1,2 @@ -1/1: Building DoubleLit (DoubleLit.idr) -Main> MkInUnit 0.25 _ +MkInUnit 0.25 _ MkNewtype 123.456 -Main> Bye for now! diff --git a/tests/idris2/basic056/input b/tests/idris2/basic056/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/idris2/basic056/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/idris2/basic056/run b/tests/idris2/basic056/run index 1b277263ad..f64488f035 100644 --- a/tests/idris2/basic056/run +++ b/tests/idris2/basic056/run @@ -1,3 +1,3 @@ . ../../testutils.sh -idris2 DoubleLit.idr < input +run DoubleLit.idr diff --git a/tests/idris2/idiom001/expected b/tests/idris2/idiom001/expected index fed011a4c5..72c8262b6b 100644 --- a/tests/idris2/idiom001/expected +++ b/tests/idris2/idiom001/expected @@ -1,4 +1,2 @@ -1/1: Building Main (Main.idr) -Main> [Just 6, Just 10, Nothing, Just 3, Just 10, Just 14, Nothing, Just 7, Nothing, Nothing, Nothing, Nothing, Just 3, Just 7, Nothing, Just 0] +[Just 6, Just 10, Nothing, Just 3, Just 10, Just 14, Nothing, Just 7, Nothing, Nothing, Nothing, Nothing, Just 3, Just 7, Nothing, Just 0] Right [Just 6, Just 10, Nothing, Just 3, Just 10, Just 14, Nothing, Just 7, Nothing, Nothing, Nothing, Nothing, Just 3, Just 7, Nothing, Just 0] -Main> Bye for now! diff --git a/tests/idris2/idiom001/input b/tests/idris2/idiom001/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/idris2/idiom001/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/idris2/idiom001/run b/tests/idris2/idiom001/run index 37beba4789..c6e6f2ab79 100644 --- a/tests/idris2/idiom001/run +++ b/tests/idris2/idiom001/run @@ -1,3 +1,3 @@ . ../../testutils.sh -idris2 Main.idr < input +run Main.idr diff --git a/tests/idris2/primloop/expected b/tests/idris2/primloop/expected index 8e1e4ffe62..f4c0c86b6f 100644 --- a/tests/idris2/primloop/expected +++ b/tests/idris2/primloop/expected @@ -1,5 +1,4 @@ -1/1: Building PrimLoop (PrimLoop.idr) -PrimLoop> 10 +10 9 8 7 @@ -9,4 +8,3 @@ PrimLoop> 10 3 2 1 -PrimLoop> Bye for now! diff --git a/tests/idris2/primloop/input b/tests/idris2/primloop/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/idris2/primloop/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/idris2/primloop/run b/tests/idris2/primloop/run index 33107d5859..63df8b24f7 100644 --- a/tests/idris2/primloop/run +++ b/tests/idris2/primloop/run @@ -1,3 +1,3 @@ . ../../testutils.sh -idris2 PrimLoop.idr < input +run PrimLoop.idr diff --git a/tests/idris2/reflection014/expected b/tests/idris2/reflection014/expected index 481d52bb4a..1cc8b5e10d 100644 --- a/tests/idris2/reflection014/expected +++ b/tests/idris2/reflection014/expected @@ -1,4 +1,2 @@ -1/1: Building refdecl (refdecl.idr) -Main> True +True False -Main> Bye for now! diff --git a/tests/idris2/reflection014/input b/tests/idris2/reflection014/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/idris2/reflection014/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/idris2/reflection014/run b/tests/idris2/reflection014/run index e86f7b9b4f..5c104e274c 100755 --- a/tests/idris2/reflection014/run +++ b/tests/idris2/reflection014/run @@ -1,3 +1,3 @@ . ../../testutils.sh -idris2 refdecl.idr < input +run refdecl.idr diff --git a/tests/node/bitops/expected b/tests/node/bitops/expected index 8848f88cd5..e69de29bb2 100644 --- a/tests/node/bitops/expected +++ b/tests/node/bitops/expected @@ -1,2 +0,0 @@ -1/1: Building BitOps (BitOps.idr) -Main> Main> Bye for now! diff --git a/tests/node/bitops/input b/tests/node/bitops/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/node/bitops/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/node/bitops/run b/tests/node/bitops/run index d5069e73e0..80bfdbf7fd 100644 --- a/tests/node/bitops/run +++ b/tests/node/bitops/run @@ -1,3 +1,3 @@ . ../../testutils.sh -idris2 --cg node BitOps.idr < input +run --cg node BitOps.idr diff --git a/tests/node/casts/expected b/tests/node/casts/expected index c11635236b..e69de29bb2 100644 --- a/tests/node/casts/expected +++ b/tests/node/casts/expected @@ -1,2 +0,0 @@ -1/1: Building Casts (Casts.idr) -Main> Main> Bye for now! diff --git a/tests/node/casts/input b/tests/node/casts/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/node/casts/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/node/casts/run b/tests/node/casts/run index 0f3c79c675..f848177774 100644 --- a/tests/node/casts/run +++ b/tests/node/casts/run @@ -1,3 +1,3 @@ . ../../testutils.sh -idris2 --cg node Casts.idr < input +run --cg node Casts.idr diff --git a/tests/node/fastConcat/expected b/tests/node/fastConcat/expected index 06e7c89862..9127656afb 100644 --- a/tests/node/fastConcat/expected +++ b/tests/node/fastConcat/expected @@ -1,4 +1,2 @@ -1/1: Building FastConcat (FastConcat.idr) -Main> 1000000 +1000000 -Main> Bye for now! diff --git a/tests/node/fastConcat/input b/tests/node/fastConcat/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/node/fastConcat/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/node/fastConcat/run b/tests/node/fastConcat/run index de7425502b..fa8a84ed86 100644 --- a/tests/node/fastConcat/run +++ b/tests/node/fastConcat/run @@ -1,3 +1,3 @@ . ../../testutils.sh -idris2 --cg node FastConcat.idr < input +run --cg node FastConcat.idr diff --git a/tests/node/fix1839/expected b/tests/node/fix1839/expected index 516cd3ed3b..448dfaa715 100644 --- a/tests/node/fix1839/expected +++ b/tests/node/fix1839/expected @@ -1,4 +1,2 @@ -1/1: Building OS (OS.idr) -Main> True -Main> Bye for now! +True Error: INTERNAL ERROR: prim not implemented: prim__os diff --git a/tests/node/fix1839/input b/tests/node/fix1839/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/node/fix1839/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/node/fix1839/run b/tests/node/fix1839/run index 60f3f7f7d7..f5d5960045 100644 --- a/tests/node/fix1839/run +++ b/tests/node/fix1839/run @@ -1,4 +1,4 @@ . ../../testutils.sh -idris2 --cg node OS.idr < input +run --cg node OS.idr idris2 --cg javascript -o os.js OS.idr diff --git a/tests/node/idiom001/expected b/tests/node/idiom001/expected index bea2e2ebc2..e69de29bb2 100644 --- a/tests/node/idiom001/expected +++ b/tests/node/idiom001/expected @@ -1,2 +0,0 @@ -1/1: Building Main (Main.idr) -Main> Main> Bye for now! diff --git a/tests/node/idiom001/input b/tests/node/idiom001/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/node/idiom001/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/node/idiom001/run b/tests/node/idiom001/run index 9e47f7b816..71cd612937 100644 --- a/tests/node/idiom001/run +++ b/tests/node/idiom001/run @@ -1,3 +1,3 @@ . ../../testutils.sh -idris2 --cg node Main.idr < input +run --cg node Main.idr diff --git a/tests/node/integer_array/expected b/tests/node/integer_array/expected index 1919dae03f..dfed9c9b1c 100644 --- a/tests/node/integer_array/expected +++ b/tests/node/integer_array/expected @@ -1,3 +1 @@ -1/1: Building array (array.idr) -Main> ["", "", "", "", "", "", "", "", "", "", "Hello", "World", "", "", "", "", "", "", "", ""] -Main> Bye for now! +["", "", "", "", "", "", "", "", "", "", "Hello", "World", "", "", "", "", "", "", "", ""] diff --git a/tests/node/integer_array/input b/tests/node/integer_array/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/node/integer_array/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/node/integer_array/run b/tests/node/integer_array/run index a67f8abb85..a8b3a176e3 100755 --- a/tests/node/integer_array/run +++ b/tests/node/integer_array/run @@ -1,3 +1,3 @@ . ../../testutils.sh -idris2 --cg node array.idr < input +run --cg node array.idr diff --git a/tests/node/memo/expected b/tests/node/memo/expected index 341a02ebf0..493d3b8b14 100644 --- a/tests/node/memo/expected +++ b/tests/node/memo/expected @@ -1,3 +1 @@ -1/1: Building Memo (Memo.idr) -Main> 100000000000000000000 -Main> Bye for now! +100000000000000000000 diff --git a/tests/node/memo/input b/tests/node/memo/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/node/memo/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/node/memo/run b/tests/node/memo/run index 32e8075391..0a41019e38 100644 --- a/tests/node/memo/run +++ b/tests/node/memo/run @@ -1,3 +1,3 @@ . ../../testutils.sh -idris2 --cg node Memo.idr < input +run --cg node Memo.idr diff --git a/tests/node/newints/expected b/tests/node/newints/expected index 80ec897d1a..e69de29bb2 100644 --- a/tests/node/newints/expected +++ b/tests/node/newints/expected @@ -1,2 +0,0 @@ -1/1: Building IntOps (IntOps.idr) -Main> Main> Bye for now! diff --git a/tests/node/newints/input b/tests/node/newints/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/node/newints/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/node/newints/run b/tests/node/newints/run index 16900feb15..0b4c2e7ebb 100644 --- a/tests/node/newints/run +++ b/tests/node/newints/run @@ -1,3 +1,3 @@ . ../../testutils.sh -idris2 --cg node IntOps.idr < input +run --cg node IntOps.idr diff --git a/tests/node/node001/expected b/tests/node/node001/expected index cc9e3bcec6..746a47c3ef 100644 --- a/tests/node/node001/expected +++ b/tests/node/node001/expected @@ -1,3 +1 @@ -1/1: Building Total (Total.idr) -Main> [1, 2, 2, 4, 3, 6, 4, 8, 5, 10] -Main> Bye for now! +[1, 2, 2, 4, 3, 6, 4, 8, 5, 10] diff --git a/tests/node/node001/input b/tests/node/node001/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/node/node001/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/node/node001/run b/tests/node/node001/run index a044ce7c09..b938f77a74 100755 --- a/tests/node/node001/run +++ b/tests/node/node001/run @@ -1,3 +1,3 @@ . ../../testutils.sh -idris2 --cg node Total.idr < input +run --cg node Total.idr diff --git a/tests/node/node003/expected b/tests/node/node003/expected index d1dbda2627..e37df034b3 100644 --- a/tests/node/node003/expected +++ b/tests/node/node003/expected @@ -1,6 +1,4 @@ -1/1: Building IORef (IORef.idr) -Main> 94 +94 94 188 188 -Main> Bye for now! diff --git a/tests/node/node003/input b/tests/node/node003/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/node/node003/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/node/node003/run b/tests/node/node003/run index 20afe1fe11..be60514c31 100755 --- a/tests/node/node003/run +++ b/tests/node/node003/run @@ -1,3 +1,3 @@ . ../../testutils.sh -idris2 --cg node IORef.idr < input +run --cg node IORef.idr diff --git a/tests/node/node004/expected b/tests/node/node004/expected index c0c7582126..72922f1f8d 100644 --- a/tests/node/node004/expected +++ b/tests/node/node004/expected @@ -1,5 +1,4 @@ -1/1: Building Buffer (Buffer.idr) -Main> 100 +100 94 94.42 "Hello" @@ -7,4 +6,3 @@ Main> 100 65535 [0, 94, 0, 0, 0, 65, 65, 65, 65, 0, 123, 20, 174, 71, 225, 154, 87, 64, 0, 0, 72, 101, 108, 108, 111, 32, 116, 104, 101, 114, 101, 33, 255, 255, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0] [0, 94, 0, 0, 0, 65, 65, 65, 65, 0, 123, 20, 174, 71, 225, 154, 87, 64, 0, 0, 72, 101, 108, 108, 111, 32, 116, 104, 101, 114, 101, 33, 255, 255, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0] -Main> Bye for now! diff --git a/tests/node/node004/input b/tests/node/node004/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/node/node004/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/node/node004/run b/tests/node/node004/run index 4731309fd7..218079face 100755 --- a/tests/node/node004/run +++ b/tests/node/node004/run @@ -2,4 +2,4 @@ rm test.buf -idris2 --cg node Buffer.idr < input +run --cg node Buffer.idr diff --git a/tests/node/node007/expected b/tests/node/node007/expected index 3361839df9..70719619ce 100644 --- a/tests/node/node007/expected +++ b/tests/node/node007/expected @@ -1,5 +1,3 @@ -1/1: Building TypeCase (TypeCase.idr) -Main> "Function from Nat to Nat" +"Function from Nat to Nat" "Function from Nat to Vector of 0 Int" "Function on Type" -Main> Bye for now! diff --git a/tests/node/node007/input b/tests/node/node007/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/node/node007/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/node/node007/run b/tests/node/node007/run index 7d554a8d79..52afecedbb 100755 --- a/tests/node/node007/run +++ b/tests/node/node007/run @@ -1,3 +1,3 @@ . ../../testutils.sh -idris2 --cg node TypeCase.idr < input +run --cg node TypeCase.idr diff --git a/tests/node/node008/expected b/tests/node/node008/expected index 9500505475..e8183f05f5 100644 --- a/tests/node/node008/expected +++ b/tests/node/node008/expected @@ -1,5 +1,3 @@ -1/1: Building Nat (Nat.idr) -Main> 1 1 1 -Main> Bye for now! +1 diff --git a/tests/node/node008/input b/tests/node/node008/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/node/node008/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/node/node008/run b/tests/node/node008/run index d7524b7005..f350dc2299 100755 --- a/tests/node/node008/run +++ b/tests/node/node008/run @@ -1,3 +1,3 @@ . ../../testutils.sh -idris2 --cg node Nat.idr < input +run --cg node Nat.idr diff --git a/tests/node/node009/expected b/tests/node/node009/expected index 759037d743..42d27acd6e 100644 --- a/tests/node/node009/expected +++ b/tests/node/node009/expected @@ -1,4 +1,2 @@ -1/1: Building uni (uni.idr) -Main> 42 +42 ällo -Main> Bye for now! diff --git a/tests/node/node009/input b/tests/node/node009/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/node/node009/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/node/node009/run b/tests/node/node009/run index e153ebec0e..a4bce0a9a7 100755 --- a/tests/node/node009/run +++ b/tests/node/node009/run @@ -1,3 +1,3 @@ . ../../testutils.sh -idris2 --cg node uni.idr < input +run --cg node uni.idr diff --git a/tests/node/node012/expected b/tests/node/node012/expected index defaf6349b..8ef81328a3 100644 --- a/tests/node/node012/expected +++ b/tests/node/node012/expected @@ -1,4 +1,2 @@ -1/1: Building array (array.idr) -Main> [Nothing, Nothing, Nothing, Nothing, Nothing, Nothing, Nothing, Nothing, Nothing, Nothing, Just "Hello", Just "World", Nothing, Nothing, Nothing, Nothing, Nothing, Nothing, Nothing, Nothing] +[Nothing, Nothing, Nothing, Nothing, Nothing, Nothing, Nothing, Nothing, Nothing, Nothing, Just "Hello", Just "World", Nothing, Nothing, Nothing, Nothing, Nothing, Nothing, Nothing, Nothing] [Just 1, Just 2, Just 3, Just 4, Just 5] -Main> Bye for now! diff --git a/tests/node/node012/input b/tests/node/node012/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/node/node012/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/node/node012/run b/tests/node/node012/run index a67f8abb85..a8b3a176e3 100755 --- a/tests/node/node012/run +++ b/tests/node/node012/run @@ -1,3 +1,3 @@ . ../../testutils.sh -idris2 --cg node array.idr < input +run --cg node array.idr diff --git a/tests/node/node014/expected b/tests/node/node014/expected index 4d6c212bbd..92aad5a55a 100644 --- a/tests/node/node014/expected +++ b/tests/node/node014/expected @@ -1,4 +1,2 @@ Received: hello world! Received: echo: hello world! -1/1: Building Echo (Echo.idr) -Main> Main> Bye for now! diff --git a/tests/node/node014/input b/tests/node/node014/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/node/node014/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/node/node014/run b/tests/node/node014/run index cbe7920e1c..f29a1de4ad 100755 --- a/tests/node/node014/run +++ b/tests/node/node014/run @@ -1,3 +1,3 @@ . ../../testutils.sh -idris2 --cg node -p network Echo.idr < input +run --cg node -p network Echo.idr diff --git a/tests/node/node015/expected b/tests/node/node015/expected index 49162bd515..a16478ddf4 100644 --- a/tests/node/node015/expected +++ b/tests/node/node015/expected @@ -1,4 +1,2 @@ -1/1: Building Numbers (Numbers.idr) -Main> [3518437212345678901234567890560, 1537557061795061679839506167983751, 3518437212345678901234567889686, 8051343735344802977653473432, 339] +[3518437212345678901234567890560, 1537557061795061679839506167983751, 3518437212345678901234567889686, 8051343735344802977653473432, 339] [912198731, 301395778, 912197977, 2419624, 106] -Main> Bye for now! diff --git a/tests/node/node015/input b/tests/node/node015/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/node/node015/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/node/node015/run b/tests/node/node015/run index cf34c951c8..a91652ce9b 100755 --- a/tests/node/node015/run +++ b/tests/node/node015/run @@ -1,3 +1,3 @@ . ../../testutils.sh -idris2 --cg node Numbers.idr < input +run --cg node Numbers.idr diff --git a/tests/node/node017/expected.in b/tests/node/node017/expected.in index 3096c91912..916d21080b 100644 --- a/tests/node/node017/expected.in +++ b/tests/node/node017/expected.in @@ -1,7 +1,5 @@ -1/1: Building dir (dir.idr) -Main> File Exists +File Exists False True Just "__PWD__testdir" -Main> Bye for now! hello diff --git a/tests/node/node017/input b/tests/node/node017/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/node/node017/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/node/node017/run b/tests/node/node017/run index 05c005b8ae..45e683248b 100755 --- a/tests/node/node017/run +++ b/tests/node/node017/run @@ -3,5 +3,5 @@ rm -rf testdir ./gen_expected.sh -idris2 --cg node dir.idr < input +run --cg node dir.idr cat testdir/test.txt diff --git a/tests/node/node018/expected b/tests/node/node018/expected index 0600e9e6bc..561ce89710 100644 --- a/tests/node/node018/expected +++ b/tests/node/node018/expected @@ -1,6 +1,4 @@ -1/1: Building File (File.idr) -Main> test test +test test unfinished lineabc def File Not Found -Main> Bye for now! diff --git a/tests/node/node018/input b/tests/node/node018/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/node/node018/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/node/node018/run b/tests/node/node018/run index f5cc69a816..bc1a5021cd 100755 --- a/tests/node/node018/run +++ b/tests/node/node018/run @@ -2,4 +2,4 @@ rm testout.txt -idris2 --cg node File.idr < input +run --cg node File.idr diff --git a/tests/node/node020/expected b/tests/node/node020/expected index 0bdb9d17d3..580712d465 100644 --- a/tests/node/node020/expected +++ b/tests/node/node020/expected @@ -1,5 +1,3 @@ -1/1: Building Popen (Popen.idr) -Main> opened +opened closed Idris 2 -Main> Bye for now! diff --git a/tests/node/node020/input b/tests/node/node020/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/node/node020/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/node/node020/run b/tests/node/node020/run index 809dc5fc59..f5d9e3c9b1 100644 --- a/tests/node/node020/run +++ b/tests/node/node020/run @@ -1,3 +1,3 @@ . ../../testutils.sh -POPEN_CMD="$idris2 --version" idris2 --cg node Popen.idr < input +POPEN_CMD="$idris2 --version" run --cg node Popen.idr diff --git a/tests/node/node022/expected b/tests/node/node022/expected index aee2314b79..f9260cb774 100644 --- a/tests/node/node022/expected +++ b/tests/node/node022/expected @@ -1,5 +1,4 @@ -1/1: Building BitCasts (BitCasts.idr) -Main> ["123", "123", "123", "123", "123"] +["123", "123", "123", "123", "123"] ["1234", "1234", "1234", "1234"] ["1234567", "1234567", "1234567"] ["134", "134", "134", "134"] @@ -8,4 +7,3 @@ Main> ["123", "123", "123", "123", "123"] ["134", "134"] ["134"] ["237", "65517", "4294967277", "18446744073709551597"] -Main> Bye for now! diff --git a/tests/node/node022/input b/tests/node/node022/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/node/node022/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/node/node022/run b/tests/node/node022/run index 624f06658f..997c6b8b01 100644 --- a/tests/node/node022/run +++ b/tests/node/node022/run @@ -1,3 +1,3 @@ . ../../testutils.sh -idris2 --cg node BitCasts.idr < input +run --cg node BitCasts.idr diff --git a/tests/node/node023/expected b/tests/node/node023/expected index 45a062228b..b130552e51 100644 --- a/tests/node/node023/expected +++ b/tests/node/node023/expected @@ -1,7 +1,5 @@ -1/1: Building Casts (Casts.idr) -Main> ok ok ok ok ok -Main> Bye for now! +ok diff --git a/tests/node/node023/input b/tests/node/node023/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/node/node023/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/node/node023/run b/tests/node/node023/run index 0f3c79c675..f848177774 100644 --- a/tests/node/node023/run +++ b/tests/node/node023/run @@ -1,3 +1,3 @@ . ../../testutils.sh -idris2 --cg node Casts.idr < input +run --cg node Casts.idr diff --git a/tests/node/node024/expected b/tests/node/node024/expected index cc59ca1e98..4f2c60daf8 100644 --- a/tests/node/node024/expected +++ b/tests/node/node024/expected @@ -1,5 +1,4 @@ -1/1: Building BitOps (BitOps.idr) -Main> -134217728 +-134217728 [0, 1, 2, 4, 8, 16, 32, 64, 127] [0, 1, 2, 4, 8, 16, 32, 64, 128, 256, 512, 1024, 2048, 4096, 8192, 16384, 32767] @@ -40,4 +39,3 @@ Main> -134217728 [340282366920938463463374607431768211444, 11, 10, 9, 15, 0] [2147483636, 11, 10, 9, 15, 0] [-2147483638, -11, -12, -9, -15, -2, 134217717, 10, 11, 9, 0] -Main> Bye for now! diff --git a/tests/node/node024/input b/tests/node/node024/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/node/node024/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/node/node024/run b/tests/node/node024/run index d5069e73e0..80bfdbf7fd 100644 --- a/tests/node/node024/run +++ b/tests/node/node024/run @@ -1,3 +1,3 @@ . ../../testutils.sh -idris2 --cg node BitOps.idr < input +run --cg node BitOps.idr diff --git a/tests/node/node025/expected b/tests/node/node025/expected index 9d6d63053f..44e0be8e35 100644 --- a/tests/node/node025/expected +++ b/tests/node/node025/expected @@ -1,6 +1,4 @@ -1/1: Building Fix1037 (Fix1037.idr) -Main> 0 0 0 0 -Main> Bye for now! +0 diff --git a/tests/node/node025/input b/tests/node/node025/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/node/node025/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/node/node025/run b/tests/node/node025/run index 0fb0d30f9f..f0d56b8b2b 100644 --- a/tests/node/node025/run +++ b/tests/node/node025/run @@ -1,3 +1,3 @@ . ../../testutils.sh -idris2 --cg node Fix1037.idr < input +run --cg node Fix1037.idr diff --git a/tests/node/node026/expected b/tests/node/node026/expected index 40ec7d5377..b3fb92c767 100644 --- a/tests/node/node026/expected +++ b/tests/node/node026/expected @@ -1,3 +1 @@ -1/1: Building Fix1795 (Fix1795.idr) -Main> Right ('1', 1) -Main> Bye for now! +Right ('1', 1) diff --git a/tests/node/node026/input b/tests/node/node026/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/node/node026/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/node/node026/run b/tests/node/node026/run index bf7de0183e..7e20c37ba3 100644 --- a/tests/node/node026/run +++ b/tests/node/node026/run @@ -1,3 +1,3 @@ . ../../testutils.sh -idris2 -p contrib --cg node Fix1795.idr < input +run -p contrib --cg node Fix1795.idr diff --git a/tests/node/node027/expected b/tests/node/node027/expected index c1247394a7..e69de29bb2 100644 --- a/tests/node/node027/expected +++ b/tests/node/node027/expected @@ -1,2 +0,0 @@ -1/1: Building PID (PID.idr) -Main> Main> Bye for now! diff --git a/tests/node/node027/input b/tests/node/node027/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/node/node027/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/node/node027/run b/tests/node/node027/run index 57df102cf8..0aadf2f22b 100644 --- a/tests/node/node027/run +++ b/tests/node/node027/run @@ -1,3 +1,3 @@ . ../../testutils.sh -idris2 --cg node PID.idr < input +run --cg node PID.idr diff --git a/tests/node/perf001/expected b/tests/node/perf001/expected index 893b70a22a..e9063aec63 100644 --- a/tests/node/perf001/expected +++ b/tests/node/perf001/expected @@ -1,3 +1 @@ -1/1: Building Span (Span.idr) -Main> ([1, 2, 3, 4, 5, 6, 7, 8, 9, 10, 11, 12, 13, 14, 15, 16, 17, 18, 19, 20, 21, 22, 23, 24, 25, 26, 27, 28, 29, 30, 31, 32, 33, 34, 35, 36, 37, 38, 39, 40, 41, 42, 43, 44, 45, 46, 47, 48, 49, 50, 51, 52, 53, 54, 55, 56, 57, 58, 59, 60, 61, 62, 63, 64, 65, 66, 67, 68, 69, 70, 71, 72, 73, 74, 75, 76, 77, 78, 79, 80], [81, 82, 83, 84, 85, 86, 87, 88, 89, 90, 91, 92, 93, 94, 95, 96, 97, 98, 99, 100]) -Main> Bye for now! +([1, 2, 3, 4, 5, 6, 7, 8, 9, 10, 11, 12, 13, 14, 15, 16, 17, 18, 19, 20, 21, 22, 23, 24, 25, 26, 27, 28, 29, 30, 31, 32, 33, 34, 35, 36, 37, 38, 39, 40, 41, 42, 43, 44, 45, 46, 47, 48, 49, 50, 51, 52, 53, 54, 55, 56, 57, 58, 59, 60, 61, 62, 63, 64, 65, 66, 67, 68, 69, 70, 71, 72, 73, 74, 75, 76, 77, 78, 79, 80], [81, 82, 83, 84, 85, 86, 87, 88, 89, 90, 91, 92, 93, 94, 95, 96, 97, 98, 99, 100]) diff --git a/tests/node/perf001/input b/tests/node/perf001/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/node/perf001/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/node/perf001/run b/tests/node/perf001/run index 09892f812e..ff95006be7 100644 --- a/tests/node/perf001/run +++ b/tests/node/perf001/run @@ -1,3 +1,3 @@ . ../../testutils.sh -idris2 --cg node Span.idr < input +run --cg node Span.idr diff --git a/tests/node/stringcast/expected b/tests/node/stringcast/expected index 01a178e3b1..66f4ca4a5a 100644 --- a/tests/node/stringcast/expected +++ b/tests/node/stringcast/expected @@ -1,5 +1,3 @@ -1/1: Building StringCast (StringCast.idr) -Main> 0 0 0 0 @@ -11,4 +9,4 @@ Main> 0 0 0 0 -Main> Bye for now! +0 diff --git a/tests/node/stringcast/input b/tests/node/stringcast/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/node/stringcast/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/node/stringcast/run b/tests/node/stringcast/run index d6166759d3..1e92367f46 100644 --- a/tests/node/stringcast/run +++ b/tests/node/stringcast/run @@ -1,3 +1,3 @@ . ../../testutils.sh -idris2 --cg node StringCast.idr < input +run --cg node StringCast.idr diff --git a/tests/prelude/bind001/expected b/tests/prelude/bind001/expected index 25d747481c..451bd091e4 100644 --- a/tests/prelude/bind001/expected +++ b/tests/prelude/bind001/expected @@ -1,4 +1,2 @@ -1/1: Building bind (bind.idr) -Main> [1, 1, 2, 1, 2, 3, 1, 2, 3, 4] +[1, 1, 2, 1, 2, 3, 1, 2, 3, 4] 12502500 -Main> Bye for now! diff --git a/tests/prelude/bind001/input b/tests/prelude/bind001/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/prelude/bind001/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/prelude/bind001/run b/tests/prelude/bind001/run index bcaf1b9b4f..63db08c646 100755 --- a/tests/prelude/bind001/run +++ b/tests/prelude/bind001/run @@ -1,3 +1,3 @@ . ../../testutils.sh -idris2 bind.idr < input +run bind.idr diff --git a/tests/prelude/double001/expected b/tests/prelude/double001/expected index ecb63744af..18f1589834 100644 --- a/tests/prelude/double001/expected +++ b/tests/prelude/double001/expected @@ -1,5 +1,4 @@ -1/1: Building Types (Types.idr) -Main> 0.0 +0.0 0.0 0.0 1.0 @@ -71,4 +70,3 @@ Main> 0.0 +nan.0 +nan.0 +nan.0 -Main> Bye for now! diff --git a/tests/prelude/double001/input b/tests/prelude/double001/input deleted file mode 100644 index fc5992c298..0000000000 --- a/tests/prelude/double001/input +++ /dev/null @@ -1,2 +0,0 @@ -:exec main -:q diff --git a/tests/prelude/double001/run b/tests/prelude/double001/run index 34962defeb..55f3f7186d 100755 --- a/tests/prelude/double001/run +++ b/tests/prelude/double001/run @@ -1,3 +1,3 @@ . ../../testutils.sh -idris2 Types.idr < input +run Types.idr From 2bde9104e0285df46d077b8070a33e3853eb9e24 Mon Sep 17 00:00:00 2001 From: Robert Wright Date: Wed, 30 Aug 2023 14:46:49 +0100 Subject: [PATCH 05/43] Add filter_test_dir test utility --- tests/chez/chez017/.gitignore | 2 -- .../expected.in => chez/chez017/expected} | 2 +- tests/chez/chez017/gen_expected.sh | 8 ------- tests/chez/chez017/run | 4 ++-- tests/idris2/pkg010/.gitignore | 2 -- tests/idris2/pkg010/expected | 6 +++++ tests/idris2/pkg010/expected.in | 6 ----- tests/idris2/pkg010/gen_expected.sh | 3 --- tests/idris2/pkg010/run | 15 +++---------- tests/idris2/pkg013/.gitignore | 2 -- tests/idris2/pkg013/expected | 1 + tests/idris2/pkg013/expected.in | 1 - tests/idris2/pkg013/gen_expected.sh | 3 --- tests/idris2/pkg013/run | 10 +-------- tests/idris2/pkg014/.gitignore | 2 -- tests/idris2/pkg014/expected | 1 + tests/idris2/pkg014/expected.in | 1 - tests/idris2/pkg014/gen_expected.sh | 3 --- tests/idris2/pkg014/run | 10 +-------- tests/idris2/pkg017/.gitignore | 2 -- tests/idris2/pkg017/expected | 22 +++++++++++++++++++ tests/idris2/pkg017/expected.in | 22 ------------------- tests/idris2/pkg017/gen_expected.sh | 3 --- tests/idris2/pkg017/run | 16 ++++---------- tests/node/node017/.gitignore | 2 -- .../expected.in => node/node017/expected} | 2 +- tests/node/node017/gen_expected.sh | 8 ------- tests/node/node017/run | 3 +-- tests/testutils.sh | 22 +++++++++++++++++++ 29 files changed, 66 insertions(+), 118 deletions(-) delete mode 100644 tests/chez/chez017/.gitignore rename tests/{node/node017/expected.in => chez/chez017/expected} (50%) delete mode 100755 tests/chez/chez017/gen_expected.sh delete mode 100644 tests/idris2/pkg010/.gitignore create mode 100644 tests/idris2/pkg010/expected delete mode 100644 tests/idris2/pkg010/expected.in delete mode 100755 tests/idris2/pkg010/gen_expected.sh delete mode 100644 tests/idris2/pkg013/.gitignore create mode 100644 tests/idris2/pkg013/expected delete mode 100644 tests/idris2/pkg013/expected.in delete mode 100755 tests/idris2/pkg013/gen_expected.sh delete mode 100644 tests/idris2/pkg014/.gitignore create mode 100644 tests/idris2/pkg014/expected delete mode 100644 tests/idris2/pkg014/expected.in delete mode 100755 tests/idris2/pkg014/gen_expected.sh delete mode 100644 tests/idris2/pkg017/.gitignore create mode 100644 tests/idris2/pkg017/expected delete mode 100644 tests/idris2/pkg017/expected.in delete mode 100755 tests/idris2/pkg017/gen_expected.sh delete mode 100644 tests/node/node017/.gitignore rename tests/{chez/chez017/expected.in => node/node017/expected} (50%) delete mode 100755 tests/node/node017/gen_expected.sh diff --git a/tests/chez/chez017/.gitignore b/tests/chez/chez017/.gitignore deleted file mode 100644 index 062cbdb0b5..0000000000 --- a/tests/chez/chez017/.gitignore +++ /dev/null @@ -1,2 +0,0 @@ -/expected - diff --git a/tests/node/node017/expected.in b/tests/chez/chez017/expected similarity index 50% rename from tests/node/node017/expected.in rename to tests/chez/chez017/expected index 916d21080b..6704aadaac 100644 --- a/tests/node/node017/expected.in +++ b/tests/chez/chez017/expected @@ -1,5 +1,5 @@ File Exists False True -Just "__PWD__testdir" +Just "__TEST_DIR__/testdir" hello diff --git a/tests/chez/chez017/gen_expected.sh b/tests/chez/chez017/gen_expected.sh deleted file mode 100755 index f670861294..0000000000 --- a/tests/chez/chez017/gen_expected.sh +++ /dev/null @@ -1,8 +0,0 @@ -#!/usr/bin/env sh -if [ "$OS" = "windows" ]; then - MY_PWD="$(cygpath -m "$(pwd)")\\\\" -else - MY_PWD="$(pwd)/" -fi - -sed -e "s|__PWD__|${MY_PWD}|g" expected.in >expected diff --git a/tests/chez/chez017/run b/tests/chez/chez017/run index f15d224d6c..de86932a8b 100755 --- a/tests/chez/chez017/run +++ b/tests/chez/chez017/run @@ -1,6 +1,6 @@ . ../../testutils.sh -./gen_expected.sh -run dir.idr +run dir.idr | filter_test_dir + cat testdir/test.txt rm -rf testdir diff --git a/tests/idris2/pkg010/.gitignore b/tests/idris2/pkg010/.gitignore deleted file mode 100644 index 062cbdb0b5..0000000000 --- a/tests/idris2/pkg010/.gitignore +++ /dev/null @@ -1,2 +0,0 @@ -/expected - diff --git a/tests/idris2/pkg010/expected b/tests/idris2/pkg010/expected new file mode 100644 index 0000000000..8b26eec408 --- /dev/null +++ b/tests/idris2/pkg010/expected @@ -0,0 +1,6 @@ +1/1: Building Main (Main.idr) +Installing __TEST_DIR__/build/ttc/Main.ttc to __TEST_DIR__/currently/nonexistent/dir/idris2-0.6.0/testpkg-0 +Installing __TEST_DIR__/build/ttc/Main.ttm to __TEST_DIR__/currently/nonexistent/dir/idris2-0.6.0/testpkg-0 +Installing __TEST_DIR__/build/ttc/Main.ttc to __TEST_DIR__/currently/nonexistent/dir/idris2-0.6.0/testpkg-0 +Installing __TEST_DIR__/build/ttc/Main.ttm to __TEST_DIR__/currently/nonexistent/dir/idris2-0.6.0/testpkg-0 +Installing package file for testpkg to __TEST_DIR__/currently/nonexistent/dir/idris2-0.6.0/testpkg-0 diff --git a/tests/idris2/pkg010/expected.in b/tests/idris2/pkg010/expected.in deleted file mode 100644 index 87fb1dc779..0000000000 --- a/tests/idris2/pkg010/expected.in +++ /dev/null @@ -1,6 +0,0 @@ -1/1: Building Main (Main.idr) -Installing __PWD__build/ttc/Main.ttc to __PWD__currently/nonexistent/dir/idris2-0.6.0/testpkg-0 -Installing __PWD__build/ttc/Main.ttm to __PWD__currently/nonexistent/dir/idris2-0.6.0/testpkg-0 -Installing __PWD__build/ttc/Main.ttc to __PWD__currently/nonexistent/dir/idris2-0.6.0/testpkg-0 -Installing __PWD__build/ttc/Main.ttm to __PWD__currently/nonexistent/dir/idris2-0.6.0/testpkg-0 -Installing package file for testpkg to __PWD__currently/nonexistent/dir/idris2-0.6.0/testpkg-0 diff --git a/tests/idris2/pkg010/gen_expected.sh b/tests/idris2/pkg010/gen_expected.sh deleted file mode 100755 index 193a8ad315..0000000000 --- a/tests/idris2/pkg010/gen_expected.sh +++ /dev/null @@ -1,3 +0,0 @@ -#!/usr/bin/env sh - -sed -e "s|__PWD__|${MY_PWD}|g" expected.in >expected diff --git a/tests/idris2/pkg010/run b/tests/idris2/pkg010/run index 32df9792a3..9277c11787 100755 --- a/tests/idris2/pkg010/run +++ b/tests/idris2/pkg010/run @@ -1,19 +1,10 @@ . ../../testutils.sh -if [ "$OS" = "windows" ]; then - MY_PWD="$(cygpath -m "$(pwd)")\\\\" -else - MY_PWD="$(pwd)/" -fi - -MY_PWD="${MY_PWD}" ./gen_expected.sh - export IDRIS2_PACKAGE_PATH=$IDRIS2_PREFIX/$NAME_VERSION -export IDRIS2_PREFIX=${MY_PWD}currently/nonexistent/dir/ +export IDRIS2_PREFIX=$test_dir/currently/nonexistent/dir/ export IDRIS2_INC_CGS= -idris2 --install ./testpkg.ipkg | sed -r "s/.([0-9]){10}//g" +idris2 --install ./testpkg.ipkg | sed -r "s/.([0-9]){10}//g" | filter_test_dir -# ../ is there for some extra safety for using rm -rf -rm -rf ../pkg010/build ../pkg010/currently +rm -rf currently diff --git a/tests/idris2/pkg013/.gitignore b/tests/idris2/pkg013/.gitignore deleted file mode 100644 index 062cbdb0b5..0000000000 --- a/tests/idris2/pkg013/.gitignore +++ /dev/null @@ -1,2 +0,0 @@ -/expected - diff --git a/tests/idris2/pkg013/expected b/tests/idris2/pkg013/expected new file mode 100644 index 0000000000..5a1648c6d1 --- /dev/null +++ b/tests/idris2/pkg013/expected @@ -0,0 +1 @@ +LOG package.depends:10: all depends: ["__TEST_DIR__/depends/bar-0.7.2", "__TEST_DIR__/depends/foo-0.1.0"] diff --git a/tests/idris2/pkg013/expected.in b/tests/idris2/pkg013/expected.in deleted file mode 100644 index c8780af900..0000000000 --- a/tests/idris2/pkg013/expected.in +++ /dev/null @@ -1 +0,0 @@ -LOG package.depends:10: all depends: ["__PWD__depends/bar-0.7.2", "__PWD__depends/foo-0.1.0"] diff --git a/tests/idris2/pkg013/gen_expected.sh b/tests/idris2/pkg013/gen_expected.sh deleted file mode 100755 index 193a8ad315..0000000000 --- a/tests/idris2/pkg013/gen_expected.sh +++ /dev/null @@ -1,3 +0,0 @@ -#!/usr/bin/env sh - -sed -e "s|__PWD__|${MY_PWD}|g" expected.in >expected diff --git a/tests/idris2/pkg013/run b/tests/idris2/pkg013/run index 84b6a3e573..4a39ba369b 100644 --- a/tests/idris2/pkg013/run +++ b/tests/idris2/pkg013/run @@ -1,11 +1,3 @@ . ../../testutils.sh -if [ "$OS" = "windows" ]; then - MY_PWD="$(cygpath -m "$(pwd)")\\\\" -else - MY_PWD="$(pwd)/" -fi - -MY_PWD="${MY_PWD}" ./gen_expected.sh - -idris2 --build test.ipkg --log package.depends:10 +idris2 --build test.ipkg --log package.depends:10 | filter_test_dir diff --git a/tests/idris2/pkg014/.gitignore b/tests/idris2/pkg014/.gitignore deleted file mode 100644 index 062cbdb0b5..0000000000 --- a/tests/idris2/pkg014/.gitignore +++ /dev/null @@ -1,2 +0,0 @@ -/expected - diff --git a/tests/idris2/pkg014/expected b/tests/idris2/pkg014/expected new file mode 100644 index 0000000000..a3766016f6 --- /dev/null +++ b/tests/idris2/pkg014/expected @@ -0,0 +1 @@ +LOG package.depends:10: all depends: ["__TEST_DIR__/depends/bar-0.1.0", "__TEST_DIR__/depends/baz-0.2.0", "__TEST_DIR__/depends/foo-0.2.0"] diff --git a/tests/idris2/pkg014/expected.in b/tests/idris2/pkg014/expected.in deleted file mode 100644 index 8b1b4576e7..0000000000 --- a/tests/idris2/pkg014/expected.in +++ /dev/null @@ -1 +0,0 @@ -LOG package.depends:10: all depends: ["__PWD__depends/bar-0.1.0", "__PWD__depends/baz-0.2.0", "__PWD__depends/foo-0.2.0"] diff --git a/tests/idris2/pkg014/gen_expected.sh b/tests/idris2/pkg014/gen_expected.sh deleted file mode 100755 index 193a8ad315..0000000000 --- a/tests/idris2/pkg014/gen_expected.sh +++ /dev/null @@ -1,3 +0,0 @@ -#!/usr/bin/env sh - -sed -e "s|__PWD__|${MY_PWD}|g" expected.in >expected diff --git a/tests/idris2/pkg014/run b/tests/idris2/pkg014/run index 84b6a3e573..4a39ba369b 100644 --- a/tests/idris2/pkg014/run +++ b/tests/idris2/pkg014/run @@ -1,11 +1,3 @@ . ../../testutils.sh -if [ "$OS" = "windows" ]; then - MY_PWD="$(cygpath -m "$(pwd)")\\\\" -else - MY_PWD="$(pwd)/" -fi - -MY_PWD="${MY_PWD}" ./gen_expected.sh - -idris2 --build test.ipkg --log package.depends:10 +idris2 --build test.ipkg --log package.depends:10 | filter_test_dir diff --git a/tests/idris2/pkg017/.gitignore b/tests/idris2/pkg017/.gitignore deleted file mode 100644 index 062cbdb0b5..0000000000 --- a/tests/idris2/pkg017/.gitignore +++ /dev/null @@ -1,2 +0,0 @@ -/expected - diff --git a/tests/idris2/pkg017/expected b/tests/idris2/pkg017/expected new file mode 100644 index 0000000000..9bf1e56120 --- /dev/null +++ b/tests/idris2/pkg017/expected @@ -0,0 +1,22 @@ +(:protocol-version 2 1) +(:return (:ok "Current working directory is \"__TEST_DIR__/b1\"") 1) +(:write-string "1/1: Building B1 (src/B1.idr)" 2) +(:return (:ok ()) 2) +(:return (:ok "1" ((0 1 ((:decor :data))))) 4) +(:return (:ok (("A.i" (:filename "__TEST_DIR__/prefix/idris2-0.6.0/a1-0/A.idr") (:start 2 0) (:end 3 7)))) 5) +(:return (:ok "Current working directory is \"__TEST_DIR__/b2\"") 6) +(:write-string "1/1: Building B2 (src/B2.idr)" 7) +(:return (:ok ()) 7) +(:return (:ok "2" ((0 1 ((:decor :data))))) 8) +(:return (:ok (("A.i" (:filename "__TEST_DIR__/prefix/idris2-0.6.0/a2-0/A.idr") (:start 2 0) (:end 3 7)))) 9) +he file is done, aborting +(:protocol-version 2 1) +(:return (:ok "Current working directory is \"__TEST_DIR__/b2\"") 1) +(:return (:ok ()) 2) +(:return (:ok "2" ((0 1 ((:decor :data))))) 4) +(:return (:ok (("A.i" (:filename "__TEST_DIR__/prefix/idris2-0.6.0/a2-0/A.idr") (:start 2 0) (:end 3 7)))) 5) +(:return (:ok "Current working directory is \"__TEST_DIR__/b1\"") 6) +(:return (:ok ()) 7) +(:return (:ok "1" ((0 1 ((:decor :data))))) 8) +(:return (:ok (("A.i" (:filename "__TEST_DIR__/prefix/idris2-0.6.0/a1-0/A.idr") (:start 2 0) (:end 3 7)))) 9) +he file is done, aborting diff --git a/tests/idris2/pkg017/expected.in b/tests/idris2/pkg017/expected.in deleted file mode 100644 index 9c54837642..0000000000 --- a/tests/idris2/pkg017/expected.in +++ /dev/null @@ -1,22 +0,0 @@ -(:protocol-version 2 1) -(:return (:ok "Current working directory is \"__PWD__b1\"") 1) -(:write-string "1/1: Building B1 (src/B1.idr)" 2) -(:return (:ok ()) 2) -(:return (:ok "1" ((0 1 ((:decor :data))))) 4) -(:return (:ok (("A.i" (:filename "__PWD__prefix/idris2-0.6.0/a1-0/A.idr") (:start 2 0) (:end 3 7)))) 5) -(:return (:ok "Current working directory is \"__PWD__b2\"") 6) -(:write-string "1/1: Building B2 (src/B2.idr)" 7) -(:return (:ok ()) 7) -(:return (:ok "2" ((0 1 ((:decor :data))))) 8) -(:return (:ok (("A.i" (:filename "__PWD__prefix/idris2-0.6.0/a2-0/A.idr") (:start 2 0) (:end 3 7)))) 9) -he file is done, aborting -(:protocol-version 2 1) -(:return (:ok "Current working directory is \"__PWD__b2\"") 1) -(:return (:ok ()) 2) -(:return (:ok "2" ((0 1 ((:decor :data))))) 4) -(:return (:ok (("A.i" (:filename "__PWD__prefix/idris2-0.6.0/a2-0/A.idr") (:start 2 0) (:end 3 7)))) 5) -(:return (:ok "Current working directory is \"__PWD__b1\"") 6) -(:return (:ok ()) 7) -(:return (:ok "1" ((0 1 ((:decor :data))))) 8) -(:return (:ok (("A.i" (:filename "__PWD__prefix/idris2-0.6.0/a1-0/A.idr") (:start 2 0) (:end 3 7)))) 9) -he file is done, aborting diff --git a/tests/idris2/pkg017/gen_expected.sh b/tests/idris2/pkg017/gen_expected.sh deleted file mode 100755 index 193a8ad315..0000000000 --- a/tests/idris2/pkg017/gen_expected.sh +++ /dev/null @@ -1,3 +0,0 @@ -#!/usr/bin/env sh - -sed -e "s|__PWD__|${MY_PWD}|g" expected.in >expected diff --git a/tests/idris2/pkg017/run b/tests/idris2/pkg017/run index 3cbf19eb7e..b99a65f8ad 100644 --- a/tests/idris2/pkg017/run +++ b/tests/idris2/pkg017/run @@ -1,20 +1,12 @@ . ../../testutils.sh -if [ "$OS" = "windows" ]; then - MY_PWD="$(cygpath -m "$(pwd)")\\\\" -else - MY_PWD="$(pwd)/" -fi - -MY_PWD="${MY_PWD}" ./gen_expected.sh - mkdir prefix -IDRIS2_PREFIX="${MY_PWD}/prefix" idris2 --install-with-src a1/a1.ipkg > /dev/null -IDRIS2_PREFIX="${MY_PWD}/prefix" idris2 --install-with-src a2/a2.ipkg > /dev/null +IDRIS2_PREFIX="$test_dir/prefix" idris2 --install-with-src a1/a1.ipkg > /dev/null +IDRIS2_PREFIX="$test_dir/prefix" idris2 --install-with-src a2/a2.ipkg > /dev/null -IDRIS2_PREFIX="${MY_PWD}/prefix" idris2 --no-prelude --ide-mode < input1 | grep -v ":highlight-source" | cut -c 7- -IDRIS2_PREFIX="${MY_PWD}/prefix" idris2 --no-prelude --ide-mode < input2 | grep -v ":highlight-source" | cut -c 7- +IDRIS2_PREFIX="$test_dir/prefix" idris2 --no-prelude --ide-mode < input1 | grep -v ":highlight-source" | cut -c 7- | filter_test_dir +IDRIS2_PREFIX="$test_dir/prefix" idris2 --no-prelude --ide-mode < input2 | grep -v ":highlight-source" | cut -c 7- | filter_test_dir rm -r a1/build a2/build b1/build b2/build rm -rf prefix diff --git a/tests/node/node017/.gitignore b/tests/node/node017/.gitignore deleted file mode 100644 index 062cbdb0b5..0000000000 --- a/tests/node/node017/.gitignore +++ /dev/null @@ -1,2 +0,0 @@ -/expected - diff --git a/tests/chez/chez017/expected.in b/tests/node/node017/expected similarity index 50% rename from tests/chez/chez017/expected.in rename to tests/node/node017/expected index 916d21080b..6704aadaac 100644 --- a/tests/chez/chez017/expected.in +++ b/tests/node/node017/expected @@ -1,5 +1,5 @@ File Exists False True -Just "__PWD__testdir" +Just "__TEST_DIR__/testdir" hello diff --git a/tests/node/node017/gen_expected.sh b/tests/node/node017/gen_expected.sh deleted file mode 100755 index f670861294..0000000000 --- a/tests/node/node017/gen_expected.sh +++ /dev/null @@ -1,8 +0,0 @@ -#!/usr/bin/env sh -if [ "$OS" = "windows" ]; then - MY_PWD="$(cygpath -m "$(pwd)")\\\\" -else - MY_PWD="$(pwd)/" -fi - -sed -e "s|__PWD__|${MY_PWD}|g" expected.in >expected diff --git a/tests/node/node017/run b/tests/node/node017/run index 45e683248b..beb8b73c58 100755 --- a/tests/node/node017/run +++ b/tests/node/node017/run @@ -2,6 +2,5 @@ rm -rf testdir -./gen_expected.sh -run --cg node dir.idr +run --cg node dir.idr | filter_test_dir cat testdir/test.txt diff --git a/tests/testutils.sh b/tests/testutils.sh index 24951eb143..6dcb4179b9 100755 --- a/tests/testutils.sh +++ b/tests/testutils.sh @@ -21,3 +21,25 @@ check() { run() { idris2 --exec main "$@" } + +# Escape a string as a sed pattern literal +# Usage: sed -e "s/$(sed_literal 'some\literal/')/some replacement/g" +sed_literal() { + # Using printf instead of echo, as echo interprets backslashes as escape sequences on some platforms + printf '%s\n' "$1" | sed -e 's/[]\/$*.^[]/\\&/g' +} + +# Folder containing the currently running test +if [ "$OS" = "windows" ]; then + test_dir="$(cygpath -m "$(pwd)")" +else + test_dir="$(pwd)" +fi + +# Remove test directory from output +# Useful for consistency of output between machines +# Usage: run SomeTest.idr | filter_test_dir +filter_test_dir() { + sed -e 's/\\\{1,2\}\b/\//g' | # Guess at where Windows \ need to be replaced by / + sed -e "s/$(sed_literal "$test_dir")/__TEST_DIR__/g" +} From 424bc20c27180c7feb212c28b9d707fb44356f71 Mon Sep 17 00:00:00 2001 From: Robert Wright Date: Fri, 1 Sep 2023 11:10:49 +0100 Subject: [PATCH 06/43] Group tests by directory --- .github/workflows/ci-idris2.yml | 2 +- tests/Main.idr | 439 +++++------------- tests/chez/barrier001/Main.idr | 2 +- tests/idris2/{api001 => api}/README | 0 tests/idris2/{ => api}/api001/Hello.idr | 0 tests/idris2/{ => api}/api001/LazyCodegen.idr | 0 tests/idris2/{ => api}/api001/expected | 0 tests/idris2/{ => api}/api001/run | 2 +- tests/idris2/{ => basic}/basic001/Vect.idr | 0 tests/idris2/{ => basic}/basic001/expected | 0 tests/idris2/{ => basic}/basic001/input | 0 tests/idris2/{ => basic}/basic001/run | 2 +- tests/idris2/{ => basic}/basic002/Do.idr | 0 tests/idris2/{ => basic}/basic002/expected | 0 tests/idris2/{ => basic}/basic002/input | 0 tests/idris2/{ => basic}/basic002/run | 2 +- tests/idris2/{ => basic}/basic003/Ambig1.idr | 0 tests/idris2/{ => basic}/basic003/Ambig2.idr | 0 tests/idris2/{ => basic}/basic003/expected | 0 tests/idris2/{ => basic}/basic003/run | 2 +- tests/idris2/{ => basic}/basic004/Stuff.idr | 0 tests/idris2/{ => basic}/basic004/Wheres.idr | 0 tests/idris2/{ => basic}/basic004/expected | 0 tests/idris2/{ => basic}/basic004/input | 0 tests/idris2/{ => basic}/basic004/run | 2 +- tests/idris2/{ => basic}/basic005/NoInfer.idr | 0 tests/idris2/{ => basic}/basic005/expected | 0 tests/idris2/{ => basic}/basic005/run | 2 +- tests/idris2/{ => basic}/basic006/PMLet.idr | 0 tests/idris2/{ => basic}/basic006/Stuff.idr | 0 tests/idris2/{ => basic}/basic006/expected | 0 tests/idris2/{ => basic}/basic006/input | 0 tests/idris2/{ => basic}/basic006/run | 2 +- tests/idris2/{ => basic}/basic007/DoLocal.idr | 0 tests/idris2/{ => basic}/basic007/Stuff.idr | 0 tests/idris2/{ => basic}/basic007/expected | 0 tests/idris2/{ => basic}/basic007/input | 0 tests/idris2/{ => basic}/basic007/run | 2 +- tests/idris2/{ => basic}/basic008/If.idr | 0 tests/idris2/{ => basic}/basic008/Stuff.idr | 0 tests/idris2/{ => basic}/basic008/expected | 0 tests/idris2/{ => basic}/basic008/input | 0 tests/idris2/{ => basic}/basic008/run | 2 +- tests/idris2/{ => basic}/basic009/LetCase.idr | 0 tests/idris2/{ => basic}/basic009/Stuff.idr | 0 tests/idris2/{ => basic}/basic009/expected | 0 tests/idris2/{ => basic}/basic009/input | 0 tests/idris2/{ => basic}/basic009/run | 2 +- tests/idris2/{ => basic}/basic010/Comp.idr | 0 tests/idris2/{ => basic}/basic010/expected | 0 tests/idris2/{ => basic}/basic010/input | 0 tests/idris2/{ => basic}/basic010/run | 2 +- tests/idris2/{ => basic}/basic011/Dots1.idr | 0 tests/idris2/{ => basic}/basic011/Dots2.idr | 0 tests/idris2/{ => basic}/basic011/Dots3.idr | 0 tests/idris2/{ => basic}/basic011/expected | 0 tests/idris2/{ => basic}/basic011/run | 2 +- tests/idris2/{ => basic}/basic012/VIndex.idr | 0 tests/idris2/{ => basic}/basic012/expected | 0 tests/idris2/basic/basic012/run | 3 + .../idris2/{ => basic}/basic013/Implicits.idr | 0 tests/idris2/{ => basic}/basic013/expected | 0 tests/idris2/basic/basic013/run | 3 + tests/idris2/{ => basic}/basic014/Rewrite.idr | 0 tests/idris2/{ => basic}/basic014/expected | 0 tests/idris2/basic/basic014/run | 3 + tests/idris2/{ => basic}/basic015/George.idr | 0 tests/idris2/{ => basic}/basic015/expected | 0 tests/idris2/basic/basic015/run | 3 + tests/idris2/{ => basic}/basic016/Eta.idr | 0 tests/idris2/{ => basic}/basic016/Eta2.idr | 0 tests/idris2/{ => basic}/basic016/expected | 0 tests/idris2/{ => basic}/basic016/run | 2 +- tests/idris2/{ => basic}/basic017/CaseInf.idr | 0 tests/idris2/{ => basic}/basic017/expected | 0 tests/idris2/{ => basic}/basic017/input | 0 tests/idris2/{ => basic}/basic017/run | 2 +- tests/idris2/{ => basic}/basic018/Fin.idr | 0 tests/idris2/{ => basic}/basic018/expected | 0 tests/idris2/basic/basic018/run | 3 + .../idris2/{ => basic}/basic019/CaseBlock.idr | 0 tests/idris2/{ => basic}/basic019/expected | 0 tests/idris2/{ => basic}/basic019/input | 0 tests/idris2/{ => basic}/basic019/run | 2 +- tests/idris2/{ => basic}/basic020/Mut.idr | 0 tests/idris2/{ => basic}/basic020/expected | 0 tests/idris2/{ => basic}/basic020/input | 0 tests/idris2/{ => basic}/basic020/run | 2 +- tests/idris2/{ => basic}/basic021/CaseDep.idr | 0 tests/idris2/{ => basic}/basic021/expected | 0 tests/idris2/{ => basic}/basic021/input | 0 tests/idris2/{ => basic}/basic021/run | 2 +- tests/idris2/{ => basic}/basic022/Erase.idr | 0 tests/idris2/{ => basic}/basic022/expected | 0 tests/idris2/{ => basic}/basic022/input | 0 tests/idris2/{ => basic}/basic022/run | 2 +- tests/idris2/{ => basic}/basic023/Params.idr | 0 tests/idris2/{ => basic}/basic023/expected | 0 tests/idris2/{ => basic}/basic023/input | 0 tests/idris2/{ => basic}/basic023/run | 2 +- tests/idris2/{ => basic}/basic024/PatLam.idr | 0 tests/idris2/{ => basic}/basic024/expected | 0 tests/idris2/{ => basic}/basic024/input | 0 tests/idris2/{ => basic}/basic024/run | 2 +- tests/idris2/{ => basic}/basic025/expected | 0 tests/idris2/{ => basic}/basic025/input | 0 tests/idris2/basic/basic025/run | 3 + tests/idris2/{ => basic}/basic026/Erl.idr | 0 tests/idris2/{ => basic}/basic026/expected | 0 tests/idris2/basic/basic026/run | 3 + tests/idris2/{ => basic}/basic027/Temp.idr | 0 tests/idris2/{ => basic}/basic027/expected | 0 tests/idris2/basic/basic027/run | 3 + tests/idris2/{ => basic}/basic028/Do.idr | 0 tests/idris2/{ => basic}/basic028/expected | 0 tests/idris2/{ => basic}/basic028/input | 0 tests/idris2/{ => basic}/basic028/run | 2 +- tests/idris2/{ => basic}/basic029/Params.idr | 0 tests/idris2/{ => basic}/basic029/expected | 0 tests/idris2/{ => basic}/basic029/input | 0 tests/idris2/{ => basic}/basic029/run | 2 +- tests/idris2/{ => basic}/basic030/arity.idr | 0 tests/idris2/{ => basic}/basic030/expected | 0 tests/idris2/basic/basic030/run | 3 + tests/idris2/{ => basic}/basic031/erased.idr | 0 tests/idris2/{ => basic}/basic031/expected | 0 tests/idris2/basic/basic031/run | 3 + tests/idris2/{ => basic}/basic032/Idiom.idr | 0 tests/idris2/{ => basic}/basic032/Idiom2.idr | 0 tests/idris2/{ => basic}/basic032/expected | 0 tests/idris2/{ => basic}/basic032/input | 0 tests/idris2/{ => basic}/basic032/run | 2 +- tests/idris2/{ => basic}/basic033/expected | 0 tests/idris2/basic/basic033/run | 3 + .../{ => basic}/basic033/unboundimps.idr | 0 tests/idris2/{ => basic}/basic034/expected | 0 tests/idris2/{ => basic}/basic034/lets.idr | 0 tests/idris2/basic/basic034/run | 3 + tests/idris2/{ => basic}/basic035/expected | 0 tests/idris2/{ => basic}/basic035/input | 0 tests/idris2/{ => basic}/basic035/run | 2 +- tests/idris2/{ => basic}/basic035/using.idr | 0 tests/idris2/{ => basic}/basic036/defimp.idr | 0 tests/idris2/{ => basic}/basic036/expected | 0 tests/idris2/{ => basic}/basic036/input | 0 tests/idris2/{ => basic}/basic036/run | 2 +- .../idris2/{ => basic}/basic037/Comments.idr | 0 .../idris2/{ => basic}/basic037/Issue279.idr | 0 tests/idris2/{ => basic}/basic037/expected | 0 tests/idris2/{ => basic}/basic037/run | 2 +- tests/idris2/{ => basic}/basic038/Resugar.idr | 0 tests/idris2/{ => basic}/basic038/expected | 0 tests/idris2/{ => basic}/basic038/input | 0 tests/idris2/{ => basic}/basic038/run | 2 +- tests/idris2/{ => basic}/basic039/Main.idr | 0 tests/idris2/{ => basic}/basic039/expected | 0 tests/idris2/{ => basic}/basic039/input | 0 tests/idris2/{ => basic}/basic039/run | 2 +- tests/idris2/{ => basic}/basic040/Default.idr | 0 tests/idris2/{ => basic}/basic040/expected | 0 tests/idris2/{ => basic}/basic040/run | 2 +- tests/idris2/{ => basic}/basic041/QDo.idr | 0 tests/idris2/{ => basic}/basic041/expected | 0 tests/idris2/basic/basic041/run | 3 + .../{ => basic}/basic042/LiteralsInteger.idr | 0 .../{ => basic}/basic042/LiteralsString.idr | 0 tests/idris2/{ => basic}/basic042/expected | 0 tests/idris2/{ => basic}/basic042/input | 0 tests/idris2/{ => basic}/basic042/input2 | 0 tests/idris2/{ => basic}/basic042/run | 2 +- .../idris2/{ => basic}/basic043/BitCasts.idr | 0 tests/idris2/{ => basic}/basic043/expected | 0 tests/idris2/{ => basic}/basic043/input | 0 tests/idris2/{ => basic}/basic043/run | 2 +- tests/idris2/{ => basic}/basic044/Term.idr | 0 tests/idris2/{ => basic}/basic044/Vec.idr | 0 tests/idris2/{ => basic}/basic044/expected | 0 tests/idris2/{ => basic}/basic044/run | 2 +- tests/idris2/{ => basic}/basic045/Main.idr | 0 tests/idris2/{ => basic}/basic045/expected | 0 tests/idris2/basic/basic045/run | 3 + .../{ => basic}/basic046/TupleSections.idr | 0 tests/idris2/{ => basic}/basic046/expected | 0 tests/idris2/{ => basic}/basic046/input | 0 tests/idris2/{ => basic}/basic046/run | 2 +- .../{ => basic}/basic047/InterleavingLets.idr | 0 tests/idris2/{ => basic}/basic047/expected | 0 tests/idris2/{ => basic}/basic047/input | 0 tests/idris2/{ => basic}/basic047/run | 2 +- tests/idris2/{ => basic}/basic049/Fld.idr | 0 tests/idris2/{ => basic}/basic049/expected | 0 tests/idris2/{ => basic}/basic049/input | 0 tests/idris2/{ => basic}/basic049/run | 2 +- tests/idris2/{ => basic}/basic050/Ilc.idr | 0 tests/idris2/{ => basic}/basic050/expected | 0 tests/idris2/{ => basic}/basic050/input | 0 tests/idris2/{ => basic}/basic050/run | 2 +- .../idris2/{ => basic}/basic051/Issue833.idr | 0 tests/idris2/{ => basic}/basic051/expected | 0 tests/idris2/basic/basic051/run | 3 + .../{ => basic}/basic052/DoubleClBrace.idr | 0 tests/idris2/{ => basic}/basic052/expected | 0 tests/idris2/{ => basic}/basic052/input | 0 tests/idris2/{ => basic}/basic052/run | 2 +- .../basic053/UnderscoredIntegerLiterals.idr | 0 tests/idris2/{ => basic}/basic053/expected | 0 tests/idris2/{ => basic}/basic053/input | 0 tests/idris2/{ => basic}/basic053/run | 2 +- .../idris2/{ => basic}/basic054/Issue1023.idr | 0 tests/idris2/{ => basic}/basic054/expected | 0 tests/idris2/basic/basic054/run | 3 + tests/idris2/{ => basic}/basic055/BitOps.idr | 0 tests/idris2/{ => basic}/basic055/expected | 0 tests/idris2/basic/basic055/run | 3 + .../idris2/{ => basic}/basic056/DoubleLit.idr | 0 tests/idris2/{ => basic}/basic056/expected | 0 tests/idris2/basic/basic056/run | 3 + tests/idris2/{ => basic}/basic057/LetIn.idr | 0 tests/idris2/{ => basic}/basic057/expected | 0 tests/idris2/basic/basic057/run | 3 + .../{ => basic}/basic058/DataTypeOp.idr | 0 .../{ => basic}/basic058/DataTypeProj.idr | 0 tests/idris2/{ => basic}/basic058/expected | 0 tests/idris2/{ => basic}/basic058/run | 2 +- .../{ => basic}/basic059/MultiClaim.idr | 0 tests/idris2/{ => basic}/basic059/expected | 0 tests/idris2/basic/basic059/run | 3 + tests/idris2/{ => basic}/basic060/Snoc.idr | 0 tests/idris2/{ => basic}/basic060/expected | 0 tests/idris2/{ => basic}/basic060/input | 0 tests/idris2/{ => basic}/basic060/run | 2 +- .../idris2/{ => basic}/basic061/IgnoreDo.idr | 0 tests/idris2/{ => basic}/basic061/expected | 0 tests/idris2/{ => basic}/basic061/input | 0 tests/idris2/{ => basic}/basic061/run | 2 +- .../idris2/{ => basic}/basic062/Issue1943.idr | 0 tests/idris2/{ => basic}/basic062/expected | 0 tests/idris2/basic/basic062/run | 3 + .../{ => basic}/basic063/NoDeclaration.idr | 0 tests/idris2/{ => basic}/basic063/expected | 0 tests/idris2/{ => basic}/basic063/input | 0 tests/idris2/{ => basic}/basic063/run | 2 +- .../idris2/{ => basic}/basic064/Issue2072.idr | 0 tests/idris2/{ => basic}/basic064/expected | 0 tests/idris2/{ => basic}/basic064/input | 0 tests/idris2/{ => basic}/basic064/run | 2 +- .../idris2/{ => basic}/basic065/Issue215.idr | 0 tests/idris2/{ => basic}/basic065/expected | 0 tests/idris2/basic/basic065/run | 3 + tests/idris2/{ => basic}/basic066/comment.idr | 0 tests/idris2/{ => basic}/basic066/expected | 0 tests/idris2/basic/basic066/run | 3 + tests/idris2/{ => basic}/basic067/expected | 0 tests/idris2/{ => basic}/basic067/input | 0 tests/idris2/{ => basic}/basic067/run | 2 +- .../idris2/{ => basic}/basic067/unclosed1.idr | 0 .../idris2/{ => basic}/basic067/unclosed2.idr | 0 .../idris2/{ => basic}/basic067/unclosed3.idr | 0 .../idris2/{ => basic}/basic068/Issue2138.idr | 0 tests/idris2/{ => basic}/basic068/expected | 0 tests/idris2/basic/basic068/run | 3 + .../idris2/{ => basic}/basic069/DebugInfo.idr | 0 tests/idris2/{ => basic}/basic069/expected | 0 tests/idris2/basic/basic069/run | 3 + .../idris2/{ => basic}/basic070/Issue2592.idr | 0 .../idris2/{ => basic}/basic070/Issue2593.idr | 0 .../idris2/{ => basic}/basic070/Issue2782.idr | 0 .../idris2/{ => basic}/basic070/Issue3016.idr | 0 tests/idris2/{ => basic}/basic070/expected | 0 tests/idris2/{ => basic}/basic070/run | 2 +- tests/idris2/{ => basic}/basic071/A.idr | 0 tests/idris2/{ => basic}/basic071/B.idr | 0 tests/idris2/{ => basic}/basic071/expected | 0 tests/idris2/{ => basic}/basic071/run | 0 .../idris2/{ => basic}/case001/InlineCase.idr | 0 tests/idris2/{ => basic}/case001/expected | 0 tests/idris2/{ => basic}/case001/run | 0 .../{ => basic}/dotted001/Issue2726.idr | 0 tests/idris2/{ => basic}/dotted001/expected | 0 tests/idris2/{ => basic}/dotted001/run | 2 +- tests/idris2/{ => basic}/idiom001/Main.idr | 0 tests/idris2/{ => basic}/idiom001/expected | 0 tests/idris2/basic/idiom001/run | 3 + .../interpolation001/IfThenElse.idr | 0 .../{ => basic}/interpolation001/expected | 0 tests/idris2/{ => basic}/interpolation001/run | 2 +- .../interpolation002/StringLiteral.idr | 0 .../{ => basic}/interpolation002/expected | 0 tests/idris2/{ => basic}/interpolation002/run | 2 +- .../{ => basic}/interpolation003/Test.idr | 0 .../{ => basic}/interpolation003/expected | 0 .../idris2/{ => basic}/interpolation003/input | 0 .../interpolation003}/run | 2 +- .../interpolation004/StringLiteral.idr | 0 .../{ => basic}/interpolation004/expected | 0 tests/idris2/{ => basic}/interpolation004/run | 2 +- tests/idris2/{ => basic}/literals001/Test.idr | 0 tests/idris2/{ => basic}/literals001/expected | 0 tests/idris2/{ => basic}/literals001/run | 0 .../{ => basic}/rewrite001/Issue2573.idr | 0 tests/idris2/{ => basic}/rewrite001/expected | 0 tests/idris2/basic/rewrite001/run | 3 + tests/idris2/basic012/run | 3 - tests/idris2/basic013/run | 3 - tests/idris2/basic014/run | 3 - tests/idris2/basic015/run | 3 - tests/idris2/basic018/run | 3 - tests/idris2/basic025/run | 3 - tests/idris2/basic026/run | 3 - tests/idris2/basic027/run | 3 - tests/idris2/basic030/run | 3 - tests/idris2/basic031/run | 3 - tests/idris2/basic033/run | 3 - tests/idris2/basic034/run | 3 - tests/idris2/basic041/run | 3 - tests/idris2/basic045/run | 3 - tests/idris2/basic051/run | 3 - tests/idris2/basic054/run | 3 - tests/idris2/basic055/run | 3 - tests/idris2/basic056/run | 3 - tests/idris2/basic057/run | 3 - tests/idris2/basic059/run | 3 - tests/idris2/basic062/run | 3 - tests/idris2/basic065/run | 3 - tests/idris2/basic066/run | 3 - tests/idris2/basic068/run | 3 - tests/idris2/basic069/run | 3 - .../idris2/{ => builtin}/builtin001/Test.idr | 0 .../idris2/{ => builtin}/builtin001/expected | 0 tests/idris2/{ => builtin}/builtin001/input | 0 .../{builtin002 => builtin/builtin001}/run | 2 +- .../idris2/{ => builtin}/builtin002/Test.idr | 0 .../idris2/{ => builtin}/builtin002/expected | 0 tests/idris2/{ => builtin}/builtin002/input | 0 .../{builtin003 => builtin/builtin002}/run | 2 +- .../idris2/{ => builtin}/builtin003/Test.idr | 0 .../idris2/{ => builtin}/builtin003/expected | 0 tests/idris2/{ => builtin}/builtin003/input | 0 .../{builtin004 => builtin/builtin003}/run | 2 +- .../idris2/{ => builtin}/builtin004/Test.idr | 0 .../idris2/{ => builtin}/builtin004/expected | 0 tests/idris2/{ => builtin}/builtin004/input | 0 tests/idris2/builtin/builtin004/run | 3 + .../idris2/{ => builtin}/builtin005/Test.idr | 0 .../idris2/{ => builtin}/builtin005/expected | 0 tests/idris2/{ => builtin}/builtin005/input | 0 tests/idris2/builtin/builtin005/run | 3 + .../idris2/{ => builtin}/builtin006/Test.idr | 0 .../idris2/{ => builtin}/builtin006/expected | 0 tests/idris2/{ => builtin}/builtin006/input | 0 tests/idris2/builtin/builtin006/run | 3 + .../idris2/{ => builtin}/builtin007/Test.idr | 0 .../idris2/{ => builtin}/builtin007/expected | 0 tests/idris2/{ => builtin}/builtin007/input | 0 tests/idris2/builtin/builtin007/run | 3 + .../idris2/{ => builtin}/builtin008/Test.idr | 0 .../idris2/{ => builtin}/builtin008/expected | 0 tests/idris2/{ => builtin}/builtin008/input | 0 tests/idris2/builtin/builtin008/run | 3 + .../idris2/{ => builtin}/builtin009/Test.idr | 0 .../idris2/{ => builtin}/builtin009/expected | 0 tests/idris2/builtin/builtin009/run | 3 + .../idris2/{ => builtin}/builtin010/Test.idr | 0 .../idris2/{ => builtin}/builtin010/expected | 0 tests/idris2/{ => builtin}/builtin010/input | 0 tests/idris2/builtin/builtin010/run | 3 + .../idris2/{ => builtin}/builtin011/Test.idr | 0 .../idris2/{ => builtin}/builtin011/expected | 0 tests/idris2/builtin/builtin011/run | 3 + .../{ => builtin}/builtin012/Issue1799.idr | 0 .../idris2/{ => builtin}/builtin012/expected | 0 tests/idris2/builtin/builtin012/run | 3 + tests/idris2/builtin005/run | 3 - tests/idris2/builtin006/run | 3 - tests/idris2/builtin007/run | 3 - tests/idris2/builtin008/run | 3 - tests/idris2/builtin009/run | 3 - tests/idris2/builtin010/run | 3 - tests/idris2/builtin011/run | 3 - tests/idris2/builtin012/run | 3 - .../idris2/{ => casetree}/casetree001/IsS.idr | 0 .../{ => casetree}/casetree001/Issue762.idr | 0 .../{ => casetree}/casetree001/expected | 0 tests/idris2/{ => casetree}/casetree001/run | 2 +- .../casetree002/DefaultCases.idr | 0 .../{ => casetree}/casetree002/Issue1079.idr | 0 .../{ => casetree}/casetree002/expected | 0 tests/idris2/{ => casetree}/casetree002/run | 2 +- .../{ => casetree}/casetree003/ForcedPats.idr | 0 .../{ => casetree}/casetree003/expected | 0 tests/idris2/{ => casetree}/casetree003/input | 0 tests/idris2/{ => casetree}/casetree003/run | 2 +- .../{ => casetree}/casetree004/LocalArgs.idr | 0 .../{ => casetree}/casetree004/PiMatch.idr | 0 .../{ => casetree}/casetree004/expected | 0 tests/idris2/{ => casetree}/casetree004/run | 2 +- .../{ => coverage}/coverage001/Vect.idr | 0 .../{ => coverage}/coverage001/Vect2.idr | 0 .../{ => coverage}/coverage001/Vect3.idr | 0 .../{ => coverage}/coverage001/expected | 0 tests/idris2/{ => coverage}/coverage001/input | 0 tests/idris2/{ => coverage}/coverage001/run | 2 +- .../{ => coverage}/coverage002/Vect.idr | 0 .../{ => coverage}/coverage002/expected | 0 tests/idris2/{ => coverage}/coverage002/input | 0 tests/idris2/{ => coverage}/coverage002/run | 2 +- .../{ => coverage}/coverage003/Cover.idr | 0 .../{ => coverage}/coverage003/expected | 0 tests/idris2/{ => coverage}/coverage003/input | 0 tests/idris2/{ => coverage}/coverage003/run | 2 +- .../{ => coverage}/coverage004/Cover.idr | 0 .../{ => coverage}/coverage004/expected | 0 tests/idris2/{ => coverage}/coverage004/input | 0 tests/idris2/{ => coverage}/coverage004/run | 2 +- .../{ => coverage}/coverage005/Cover.idr | 0 .../{ => coverage}/coverage005/expected | 0 tests/idris2/{ => coverage}/coverage005/input | 0 tests/idris2/{ => coverage}/coverage005/run | 2 +- .../{ => coverage}/coverage006/expected | 0 .../{ => coverage}/coverage006/foobar.idr | 0 tests/idris2/{ => coverage}/coverage006/input | 0 tests/idris2/{ => coverage}/coverage006/run | 2 +- .../idris2/{ => coverage}/coverage007/eq.idr | 0 .../{ => coverage}/coverage007/expected | 0 tests/idris2/coverage/coverage007/run | 3 + .../{ => coverage}/coverage008/expected | 0 tests/idris2/{ => coverage}/coverage008/input | 0 tests/idris2/{ => coverage}/coverage008/run | 2 +- .../{ => coverage}/coverage008/wcov.idr | 0 .../{ => coverage}/coverage009/expected | 0 tests/idris2/coverage/coverage009/run | 3 + .../coverage009/unreachable.idr | 0 .../{ => coverage}/coverage010/casetot.idr | 0 .../{ => coverage}/coverage010/expected | 0 tests/idris2/coverage/coverage010/run | 3 + .../{ => coverage}/coverage011/Sing.idr | 0 .../{ => coverage}/coverage011/expected | 0 tests/idris2/coverage/coverage011/run | 3 + .../{ => coverage}/coverage012/Issue484.idr | 0 .../{ => coverage}/coverage012/Issue899.idr | 0 .../{ => coverage}/coverage012/expected | 0 tests/idris2/{ => coverage}/coverage012/run | 2 +- .../coverage013/Issue1022-Refl.idr | 0 .../{ => coverage}/coverage013/Issue1022.idr | 0 .../{ => coverage}/coverage013/expected | 0 tests/idris2/{ => coverage}/coverage013/run | 2 +- .../{ => coverage}/coverage014/Issue794.idr | 0 .../{ => coverage}/coverage014/expected | 0 tests/idris2/coverage/coverage014/run | 3 + .../{ => coverage}/coverage015/Issue1169.idr | 0 .../{ => coverage}/coverage015/Issue1366.idr | 0 .../{ => coverage}/coverage015/expected | 0 tests/idris2/{ => coverage}/coverage015/run | 2 +- .../{ => coverage}/coverage016/Issue633-2.idr | 0 .../{ => coverage}/coverage016/Issue633.idr | 0 .../{ => coverage}/coverage016/expected | 0 tests/idris2/{ => coverage}/coverage016/run | 2 +- .../{ => coverage}/coverage017/Issue1421.idr | 0 .../{ => coverage}/coverage017/expected | 0 tests/idris2/coverage/coverage017/run | 3 + .../coverage018/Issue1831_1.idr | 0 .../coverage018/Issue1831_2.idr | 0 .../{ => coverage}/coverage018/expected | 0 tests/idris2/{ => coverage}/coverage018/input | 0 tests/idris2/{ => coverage}/coverage018/run | 2 +- .../{ => coverage}/coverage019/Issue1632.idr | 0 .../{ => coverage}/coverage019/expected | 0 tests/idris2/coverage/coverage019/run | 3 + tests/idris2/coverage007/run | 3 - tests/idris2/coverage009/run | 3 - tests/idris2/coverage010/run | 3 - tests/idris2/coverage011/run | 3 - tests/idris2/coverage014/run | 3 - tests/idris2/coverage017/run | 3 - tests/idris2/coverage019/run | 3 - tests/idris2/{ => data}/data001/Test.idr | 0 tests/idris2/{ => data}/data001/TestImpl.idr | 0 tests/idris2/{ => data}/data001/expected | 0 tests/idris2/data/data001/run | 3 + tests/idris2/{ => data}/data002/Test.idr | 0 tests/idris2/{ => data}/data002/expected | 0 tests/idris2/{ => data}/data002/input | 0 tests/idris2/data/data002/run | 3 + tests/idris2/{ => data}/record001/Record.idr | 0 tests/idris2/{ => data}/record001/expected | 0 tests/idris2/{ => data}/record001/input | 0 tests/idris2/{ => data}/record001/run | 2 +- tests/idris2/{ => data}/record002/Record.idr | 0 tests/idris2/{ => data}/record002/expected | 0 tests/idris2/{ => data}/record002/input | 0 tests/idris2/{ => data}/record002/run | 2 +- tests/idris2/{ => data}/record003/Record.idr | 0 tests/idris2/{ => data}/record003/expected | 0 tests/idris2/data/record003/run | 3 + tests/idris2/{ => data}/record004/Main.idr | 0 tests/idris2/{ => data}/record004/expected | 0 tests/idris2/{ => data}/record004/input | 0 tests/idris2/{ => data}/record004/run | 2 +- tests/idris2/{ => data}/record005/Fld.idr | 0 tests/idris2/{ => data}/record005/expected | 0 tests/idris2/{ => data}/record005/input | 0 tests/idris2/{ => data}/record005/run | 2 +- tests/idris2/{ => data}/record006/Fld.idr | 0 tests/idris2/{ => data}/record006/expected | 0 tests/idris2/{ => data}/record006/input | 0 tests/idris2/{ => data}/record006/run | 2 +- tests/idris2/{ => data}/record007/Bond.idr | 0 tests/idris2/{ => data}/record007/expected | 0 .../idris2/{error008 => data/record007}/input | 0 tests/idris2/{ => data}/record007/run | 2 +- tests/idris2/{ => data}/record008/Postfix.idr | 0 tests/idris2/{ => data}/record008/expected | 0 tests/idris2/{ => data}/record008/input | 0 tests/idris2/{ => data}/record008/run | 2 +- tests/idris2/{ => data}/record009/expected | 0 tests/idris2/{ => data}/record009/input | 0 tests/idris2/{ => data}/record009/record.idr | 0 tests/idris2/{ => data}/record009/run | 2 +- tests/idris2/{ => data}/record010/expected | 0 tests/idris2/{ => data}/record010/record.idr | 0 tests/idris2/data/record010/run | 3 + .../idris2/{ => data}/record011/Issue2095.idr | 0 tests/idris2/{ => data}/record011/expected | 0 tests/idris2/data/record011/run | 3 + .../idris2/{ => data}/record012/Issue2065.idr | 0 tests/idris2/{ => data}/record012/expected | 0 tests/idris2/data/record012/run | 3 + .../idris2/{ => data}/record013/Issue1945.idr | 0 tests/idris2/{ => data}/record013/expected | 0 tests/idris2/{ => data}/record013/input | 0 tests/idris2/{ => data}/record013/run | 2 +- .../idris2/{ => data}/record014/Issue1404.idr | 0 tests/idris2/{ => data}/record014/expected | 0 tests/idris2/data/record014/run | 3 + .../idris2/{ => data}/record015/Issue2176.idr | 0 tests/idris2/{ => data}/record015/expected | 0 tests/idris2/data/record015/run | 3 + .../{ => data}/record016/HoleRecord.idr | 0 tests/idris2/{ => data}/record016/expected | 0 tests/idris2/{ => data}/record016/input | 0 tests/idris2/{ => data}/record016/run | 2 +- .../{ => data}/record017/RecordOptions.idr | 0 tests/idris2/{ => data}/record017/expected | 0 tests/idris2/{ => data}/record017/run | 2 +- tests/idris2/{ => data}/record018/expected | 0 tests/idris2/{ => data}/record018/mut.idr | 0 tests/idris2/data/record018/run | 3 + .../{ => data}/record019/BindParams.idr | 0 tests/idris2/{ => data}/record019/expected | 0 tests/idris2/{ => data}/record019/run | 2 +- tests/idris2/data001/run | 3 - tests/idris2/data002/run | 3 - tests/idris2/{ => debug}/debug001/TypePat.idr | 0 tests/idris2/{ => debug}/debug001/expected | 0 tests/idris2/{ => debug}/debug001/input | 0 tests/idris2/{ => debug}/debug001/run | 2 +- tests/idris2/docs001/run | 3 - tests/idris2/docs005/run | 3 - tests/idris2/{ => error}/error001/Error.idr | 0 tests/idris2/{ => error}/error001/expected | 0 tests/idris2/error/error001/run | 3 + tests/idris2/{ => error}/error002/Error.idr | 0 tests/idris2/{ => error}/error002/expected | 0 tests/idris2/error/error002/run | 3 + tests/idris2/{ => error}/error003/Error.idr | 0 tests/idris2/{ => error}/error003/expected | 0 tests/idris2/error/error003/run | 3 + tests/idris2/{ => error}/error004/Error1.idr | 0 tests/idris2/{ => error}/error004/Error2.idr | 0 tests/idris2/{ => error}/error004/expected | 0 tests/idris2/{ => error}/error004/run | 2 +- tests/idris2/{ => error}/error005/IfErr.idr | 0 tests/idris2/{ => error}/error005/expected | 0 tests/idris2/error/error005/run | 3 + tests/idris2/{ => error}/error006/IfErr.idr | 0 tests/idris2/{ => error}/error006/expected | 0 tests/idris2/error/error006/run | 3 + tests/idris2/{ => error}/error007/CongErr.idr | 0 tests/idris2/{ => error}/error007/expected | 0 tests/idris2/error/error007/run | 3 + tests/idris2/{ => error}/error008/expected | 0 .../idris2/{error009 => error/error008}/input | 0 tests/idris2/{ => error}/error008/run | 2 +- tests/idris2/{ => error}/error009/Exists.idr | 0 tests/idris2/{ => error}/error009/expected | 0 .../{interface025 => error/error009}/input | 0 tests/idris2/{ => error}/error009/run | 2 +- tests/idris2/{ => error}/error010/Loop.idr | 0 tests/idris2/{ => error}/error010/expected | 0 tests/idris2/error/error010/run | 3 + .../error011/ConstructorDuplicate.idr | 0 tests/idris2/{ => error}/error011/expected | 0 tests/idris2/{ => error}/error011/run | 2 +- tests/idris2/{ => error}/error012/expected | 0 tests/idris2/{ => error}/error012/run | 0 .../idris2/{ => error}/error013/Issue361.idr | 0 tests/idris2/{ => error}/error013/expected | 0 tests/idris2/error/error013/run | 3 + .../idris2/{ => error}/error014/Issue735.idr | 0 tests/idris2/{ => error}/error014/expected | 0 tests/idris2/error/error014/run | 3 + .../idris2/{ => error}/error015/Issue110.idr | 0 tests/idris2/{ => error}/error015/expected | 0 tests/idris2/error/error015/run | 3 + .../idris2/{ => error}/error016/Issue1230.idr | 0 tests/idris2/{ => error}/error016/expected | 0 tests/idris2/{ => error}/error016/input | 0 tests/idris2/{ => error}/error016/run | 2 +- .../{ => error}/error017/Issue962-case.idr | 0 .../idris2/{ => error}/error017/Issue962.idr | 0 tests/idris2/{ => error}/error017/expected | 0 tests/idris2/{ => error}/error017/run | 2 +- .../{ => error}/error018/Issue1031-2.idr | 0 .../{ => error}/error018/Issue1031-3.idr | 0 .../{ => error}/error018/Issue1031-4.idr | 0 .../idris2/{ => error}/error018/Issue1031.idr | 0 tests/idris2/{ => error}/error018/expected | 0 tests/idris2/{ => error}/error018/run | 2 +- tests/idris2/{ => error}/error019/Error.idr | 0 tests/idris2/{ => error}/error019/expected | 0 tests/idris2/{ => error}/error019/run | 2 +- tests/idris2/{ => error}/error020/Error.idr | 0 tests/idris2/{ => error}/error020/expected | 0 tests/idris2/{ => error}/error020/run | 2 +- .../idris2/{ => error}/error021/DeepAmbig.idr | 0 tests/idris2/{ => error}/error021/expected | 0 tests/idris2/{ => error}/error021/run | 2 +- .../idris2/{ => error}/error022/UpdateLoc.idr | 0 tests/idris2/{ => error}/error022/expected | 0 tests/idris2/error/error022/run | 3 + tests/idris2/{ => error}/error023/Error1.idr | 0 tests/idris2/{ => error}/error023/Error2.idr | 0 tests/idris2/{ => error}/error023/expected | 0 tests/idris2/{ => error}/error023/run | 2 +- tests/idris2/{ => error}/error024/Error1.idr | 0 tests/idris2/{ => error}/error024/expected | 0 tests/idris2/error/error024/run | 3 + .../error025/IAlternativePrints.idr | 0 tests/idris2/{ => error}/error025/expected | 0 tests/idris2/{ => error}/error025/run | 2 +- .../idris2/{ => error}/error026/DoBlockFC.idr | 0 tests/idris2/{ => error}/error026/expected | 0 tests/idris2/error/error026/run | 3 + .../idris2/{ => error}/error027/Issue2950.idr | 0 tests/idris2/{ => error}/error027/expected | 0 tests/idris2/error/error027/run | 3 + tests/idris2/{ => error}/perror001/PError.idr | 0 tests/idris2/{ => error}/perror001/expected | 0 tests/idris2/error/perror001/run | 3 + tests/idris2/{ => error}/perror002/PError.idr | 0 tests/idris2/{ => error}/perror002/expected | 0 tests/idris2/error/perror002/run | 3 + tests/idris2/{ => error}/perror003/PError.idr | 0 .../idris2/{ => error}/perror003/PError2.idr | 0 tests/idris2/{ => error}/perror003/expected | 0 tests/idris2/{ => error}/perror003/run | 2 +- tests/idris2/{ => error}/perror004/PError.idr | 0 tests/idris2/{ => error}/perror004/expected | 0 tests/idris2/error/perror004/run | 3 + tests/idris2/{ => error}/perror005/PError.idr | 0 tests/idris2/{ => error}/perror005/expected | 0 tests/idris2/error/perror005/run | 3 + tests/idris2/{ => error}/perror006/PError.idr | 0 tests/idris2/{ => error}/perror006/expected | 0 tests/idris2/error/perror006/run | 3 + .../{ => error}/perror007/StrError1.idr | 0 .../{ => error}/perror007/StrError10.idr | 0 .../{ => error}/perror007/StrError11.idr | 0 .../{ => error}/perror007/StrError12.idr | 0 .../{ => error}/perror007/StrError2.idr | 0 .../{ => error}/perror007/StrError3.idr | 0 .../{ => error}/perror007/StrError4.idr | 0 .../{ => error}/perror007/StrError5.idr | 0 .../{ => error}/perror007/StrError6.idr | 0 .../{ => error}/perror007/StrError7.idr | 0 .../{ => error}/perror007/StrError8.idr | 0 .../{ => error}/perror007/StrError9.idr | 0 tests/idris2/{ => error}/perror007/expected | 0 tests/idris2/{ => error}/perror007/run | 2 +- .../{ => error}/perror008/Issue1224a.idr | 0 .../{ => error}/perror008/Issue1224b.idr | 0 .../{ => error}/perror008/Issue710a.idr | 0 .../{ => error}/perror008/Issue710b.idr | 0 .../{ => error}/perror008/Issue710c.idr | 0 .../{ => error}/perror008/Issue710d.idr | 0 .../{ => error}/perror008/Issue710e.idr | 0 .../{ => error}/perror008/Issue710f.idr | 0 tests/idris2/{ => error}/perror008/expected | 0 tests/idris2/{ => error}/perror008/run | 2 +- tests/idris2/{ => error}/perror009/Error1.idr | 0 tests/idris2/{ => error}/perror009/expected | 0 tests/idris2/{ => error}/perror009/run | 2 +- .../{ => error}/perror010/NamedReturn1.idr | 0 .../{ => error}/perror010/NamedReturn2.idr | 0 .../{ => error}/perror010/NamedReturn3.idr | 0 .../{ => error}/perror010/NamedReturn4.idr | 0 .../{ => error}/perror010/TrailingLam.idr | 0 tests/idris2/{ => error}/perror010/expected | 0 tests/idris2/{ => error}/perror010/input | 0 tests/idris2/{ => error}/perror010/run | 2 +- .../{ => error}/perror011/Issue1345.idr | 0 .../{ => error}/perror011/Issue1496-1.idr | 0 .../{ => error}/perror011/Issue1496-2.idr | 0 tests/idris2/{ => error}/perror011/Main.idr | 0 tests/idris2/{ => error}/perror011/Pretty.idr | 0 tests/idris2/{ => error}/perror011/expected | 0 tests/idris2/{ => error}/perror011/foo.ipkg | 0 tests/idris2/{ => error}/perror011/run | 2 +- .../{ => error}/perror012/CaseParseError.idr | 0 .../{ => error}/perror012/LamParseError.idr | 0 tests/idris2/{ => error}/perror012/expected | 0 tests/idris2/{ => error}/perror012/run | 2 +- .../{ => error}/perror013/EmptyFailing.idr | 0 .../{ => error}/perror013/EmptyMutual.idr | 0 .../{ => error}/perror013/EmptyParameters.idr | 0 .../{ => error}/perror013/EmptyUsing.idr | 0 tests/idris2/{ => error}/perror013/expected | 0 tests/idris2/{ => error}/perror013/run | 2 +- tests/idris2/{ => error}/perror013/x.ipkg | 0 .../{ => error}/perror014/ParseList.idr | 0 tests/idris2/{ => error}/perror014/expected | 0 tests/idris2/{ => error}/perror014/run | 2 +- .../{ => error}/perror015/ParseWith.idr | 0 tests/idris2/{ => error}/perror015/expected | 0 tests/idris2/{ => error}/perror015/run | 2 +- .../idris2/{ => error}/perror016/ParseIf.idr | 0 .../idris2/{ => error}/perror016/ParseIf2.idr | 0 .../idris2/{ => error}/perror016/ParseIf3.idr | 0 .../idris2/{ => error}/perror016/ParseIf4.idr | 0 tests/idris2/{ => error}/perror016/expected | 0 tests/idris2/{ => error}/perror016/run | 2 +- .../{ => error}/perror017/ParseImpl.idr | 0 tests/idris2/{ => error}/perror017/expected | 0 tests/idris2/{ => error}/perror017/run | 2 +- .../{ => error}/perror018/ParseRecord.idr | 0 .../{ => error}/perror018/ParseRecord2.idr | 0 .../{ => error}/perror018/ParseRecord3.idr | 0 tests/idris2/{ => error}/perror018/expected | 0 tests/idris2/{ => error}/perror018/run | 2 +- .../{ => error}/perror019/ImplError.idr | 0 tests/idris2/{ => error}/perror019/expected | 0 tests/idris2/error/perror019/run | 3 + .../{ => error}/perror020/Issue2769.idr | 0 .../{ => error}/perror020/Issue2769b.idr | 0 tests/idris2/{ => error}/perror020/expected | 0 tests/idris2/{ => error}/perror020/run | 2 +- .../idris2/{ => error}/perror021/Implicit.idr | 0 tests/idris2/{ => error}/perror021/expected | 0 tests/idris2/error/perror021/run | 3 + tests/idris2/{ => error}/perror022/Indent.idr | 0 tests/idris2/{ => error}/perror022/expected | 0 tests/idris2/error/perror022/run | 3 + .../{ => error}/perror023/ParseError.idr | 0 tests/idris2/{ => error}/perror023/expected | 0 tests/idris2/error/perror023/run | 3 + .../{ => error}/perror024/ParseError.idr | 0 tests/idris2/{ => error}/perror024/expected | 0 tests/idris2/error/perror024/run | 3 + .../{ => error}/perror025/DataWhere.idr | 0 tests/idris2/{ => error}/perror025/expected | 0 tests/idris2/error/perror025/run | 3 + tests/idris2/{ => error}/perror026/Micro.idr | 0 tests/idris2/{ => error}/perror026/expected | 0 tests/idris2/error/perror026/run | 3 + .../idris2/{ => error}/perror027/Outdent.idr | 0 tests/idris2/{ => error}/perror027/expected | 0 tests/idris2/error/perror027/run | 3 + .../idris2/{ => error}/perror028/LetInDo.idr | 0 tests/idris2/{ => error}/perror028/expected | 0 tests/idris2/error/perror028/run | 3 + .../{ => error}/perror029/DelayParse.idr | 0 tests/idris2/{ => error}/perror029/expected | 0 tests/idris2/error/perror029/run | 3 + tests/idris2/error001/run | 3 - tests/idris2/error002/run | 3 - tests/idris2/error003/run | 3 - tests/idris2/error005/run | 3 - tests/idris2/error006/run | 3 - tests/idris2/error007/run | 3 - tests/idris2/error010/run | 3 - tests/idris2/error013/run | 3 - tests/idris2/error014/run | 3 - tests/idris2/error015/run | 3 - tests/idris2/error022/run | 3 - tests/idris2/error024/run | 3 - tests/idris2/error026/run | 3 - tests/idris2/error027/run | 3 - tests/idris2/eta001/run | 3 - .../{ => evaluator}/evaluator001/Issue650.idr | 0 .../{ => evaluator}/evaluator001/expected | 0 tests/idris2/evaluator/evaluator001/run | 3 + .../{ => evaluator}/evaluator002/Lib.idr | 0 .../{ => evaluator}/evaluator002/Main.idr | 0 .../{ => evaluator}/evaluator002/expected | 0 .../idris2/{ => evaluator}/evaluator002/input | 0 tests/idris2/{ => evaluator}/evaluator002/run | 2 +- .../{ => evaluator}/evaluator003/Issue705.idr | 0 .../{ => evaluator}/evaluator003/expected | 0 .../idris2/{ => evaluator}/evaluator003/input | 0 tests/idris2/{ => evaluator}/evaluator003/run | 2 +- .../evaluator004/Issue1282.idr | 0 .../{ => evaluator}/evaluator004/expected | 0 .../idris2/{ => evaluator}/evaluator004/input | 0 tests/idris2/{ => evaluator}/evaluator004/run | 2 +- .../{ => evaluator}/interpreter001/expected | 0 .../{ => evaluator}/interpreter001/input | 0 tests/idris2/evaluator/interpreter001/run | 3 + .../{ => evaluator}/interpreter002/expected | 0 .../{ => evaluator}/interpreter002/input | 0 .../idris2/{ => evaluator}/interpreter002/run | 2 +- .../{ => evaluator}/interpreter003/expected | 0 .../{ => evaluator}/interpreter003/input | 0 tests/idris2/evaluator/interpreter003/run | 3 + .../{ => evaluator}/interpreter004/expected | 0 .../{ => evaluator}/interpreter004/input | 0 tests/idris2/evaluator/interpreter004/run | 3 + .../interpreter005/Issue37.idr | 0 .../interpreter005/Issue37.lidr | 0 .../{ => evaluator}/interpreter005/expected | 0 .../{ => evaluator}/interpreter005/input | 0 .../idris2/{ => evaluator}/interpreter005/run | 2 +- .../{ => evaluator}/interpreter006/expected | 0 .../{ => evaluator}/interpreter006/input | 0 tests/idris2/evaluator/interpreter006/run | 3 + .../{ => evaluator}/interpreter007/expected | 0 .../{ => evaluator}/interpreter007/input | 0 tests/idris2/evaluator/interpreter007/run | 3 + .../interpreter008/Issue2041.idr | 0 .../{ => evaluator}/interpreter008/expected | 0 .../{ => evaluator}/interpreter008/input | 0 .../idris2/{ => evaluator}/interpreter008/run | 2 +- tests/idris2/{ => evaluator}/spec001/Desc.idr | 0 .../idris2/{ => evaluator}/spec001/Desc2.idr | 0 .../{ => evaluator}/spec001/Identity.idr | 0 .../idris2/{ => evaluator}/spec001/Mult3.idr | 0 tests/idris2/{ => evaluator}/spec001/expected | 0 tests/idris2/{ => evaluator}/spec001/run | 2 +- tests/idris2/evaluator001/run | 3 - .../idris2/{ => failing}/failing001/Fail.idr | 0 .../idris2/{ => failing}/failing001/expected | 0 tests/idris2/failing/failing001/run | 3 + .../{ => failing}/failing002/FailingBug.idr | 0 .../idris2/{ => failing}/failing002/expected | 0 tests/idris2/{ => failing}/failing002/run | 6 +- .../failing003/FailingTotality.idr | 0 .../idris2/{ => failing}/failing003/expected | 0 tests/idris2/{ => failing}/failing003/run | 2 +- .../{ => failing}/failing004/Issue2821.idr | 0 .../idris2/{ => failing}/failing004/expected | 0 tests/idris2/failing/failing004/run | 3 + tests/idris2/failing001/run | 3 - tests/idris2/failing004/run | 3 - tests/idris2/idiom001/run | 3 - tests/idris2/import005/run | 3 - tests/idris2/import007/run | 3 - .../interactive001/LocType.idr | 0 .../{ => interactive}/interactive001/expected | 0 .../{ => interactive}/interactive001/input | 0 .../{ => interactive}/interactive001/run | 2 +- .../interactive002/IEdit.idr | 0 .../{ => interactive}/interactive002/expected | 0 .../{ => interactive}/interactive002/input | 0 .../{ => interactive}/interactive002/run | 2 +- .../interactive003/IEdit.idr | 0 .../interactive003/IEdit2.idr | 0 .../{ => interactive}/interactive003/expected | 0 .../{ => interactive}/interactive003/input | 0 .../{ => interactive}/interactive003/input2 | 0 .../{ => interactive}/interactive003/run | 2 +- .../interactive004/IEdit.idr | 0 .../{ => interactive}/interactive004/expected | 0 .../{ => interactive}/interactive004/input | 0 .../interactive004}/run | 2 +- .../interactive005/IEdit.idr | 0 .../{ => interactive}/interactive005/expected | 0 .../{ => interactive}/interactive005/input | 0 .../interactive005}/run | 2 +- .../interactive006/IEdit.idr | 0 .../{ => interactive}/interactive006/expected | 0 .../{ => interactive}/interactive006/input | 0 .../interactive006}/run | 2 +- .../interactive007/IEdit.idr | 0 .../{ => interactive}/interactive007/expected | 0 .../{ => interactive}/interactive007/input | 0 tests/idris2/interactive/interactive007/run | 3 + .../interactive008/IEdit.idr | 0 .../{ => interactive}/interactive008/expected | 0 .../{ => interactive}/interactive008/input | 0 tests/idris2/interactive/interactive008/run | 3 + .../{ => interactive}/interactive009/Door.idr | 0 .../{ => interactive}/interactive009/expected | 0 .../{ => interactive}/interactive009/input | 0 .../{ => interactive}/interactive009/run | 2 +- .../interactive010/IEdit.idr | 0 .../{ => interactive}/interactive010/expected | 0 .../{ => interactive}/interactive010/input | 0 tests/idris2/interactive/interactive010/run | 3 + .../interactive011/IEdit.idr | 0 .../{ => interactive}/interactive011/expected | 0 .../{ => interactive}/interactive011/input | 0 tests/idris2/interactive/interactive011/run | 3 + .../interactive012/WithLift.idr | 0 .../{ => interactive}/interactive012/expected | 0 .../{ => interactive}/interactive012/input | 0 .../{ => interactive}/interactive012/run | 2 +- .../interactive013/Spacing.idr | 0 .../{ => interactive}/interactive013/expected | 0 .../{ => interactive}/interactive013/input | 0 .../{ => interactive}/interactive013/run | 2 +- .../{ => interactive}/interactive014/case.idr | 0 .../{ => interactive}/interactive014/expected | 0 .../{ => interactive}/interactive014/input | 0 .../{ => interactive}/interactive014/run | 2 +- .../interactive015/IEdit.idr | 0 .../{ => interactive}/interactive015/expected | 0 .../{ => interactive}/interactive015/input | 0 tests/idris2/interactive/interactive015/run | 3 + .../{ => interactive}/interactive016/Cont.idr | 0 .../{ => interactive}/interactive016/expected | 0 .../{ => interactive}/interactive016/input | 0 .../{ => interactive}/interactive016/run | 2 +- .../{ => interactive}/interactive017/RLE.idr | 0 .../{ => interactive}/interactive017/expected | 0 .../{ => interactive}/interactive017/input | 0 .../{ => interactive}/interactive017/run | 2 +- .../interactive018/PlusPrf.idr | 0 .../{ => interactive}/interactive018/expected | 0 .../{ => interactive}/interactive018/input | 0 .../{ => interactive}/interactive018/run | 2 +- .../interactive019/TypeSearch.idr | 0 .../{ => interactive}/interactive019/expected | 0 .../{ => interactive}/interactive019/input | 0 .../{ => interactive}/interactive019/run | 2 +- .../interactive020/Issue835.idr | 0 .../{ => interactive}/interactive020/expected | 0 .../{ => interactive}/interactive020/input | 0 .../{ => interactive}/interactive020/run | 2 +- .../interactive021/TypeAtDoNotation.idr | 0 .../{ => interactive}/interactive021/expected | 0 .../{ => interactive}/interactive021/input | 0 .../{ => interactive}/interactive021/run | 2 +- .../interactive022/TypeAtBangSyntax.idr | 0 .../{ => interactive}/interactive022/expected | 0 .../{ => interactive}/interactive022/input | 0 .../{ => interactive}/interactive022/run | 2 +- .../interactive023/TypeAtLambda.idr | 0 .../{ => interactive}/interactive023/expected | 0 .../{ => interactive}/interactive023/input | 0 .../{ => interactive}/interactive023/run | 2 +- .../interactive024/TypeAtAsPatterns.idr | 0 .../{ => interactive}/interactive024/expected | 0 .../{ => interactive}/interactive024/input | 0 .../{ => interactive}/interactive024/run | 2 +- .../interactive025/TypeAtInterfaces.idr | 0 .../{ => interactive}/interactive025/expected | 0 .../{ => interactive}/interactive025/input | 0 .../{ => interactive}/interactive025/run | 2 +- .../interactive026/TypeAtRecords.idr | 0 .../{ => interactive}/interactive026/expected | 0 .../{ => interactive}/interactive026/input | 0 .../{ => interactive}/interactive026/run | 2 +- .../interactive027/TypeAtLocalVars.idr | 0 .../{ => interactive}/interactive027/expected | 0 .../{ => interactive}/interactive027/input | 0 .../{ => interactive}/interactive027/run | 2 +- .../{ => interactive}/interactive028/expected | 0 .../{ => interactive}/interactive028/input | 0 tests/idris2/interactive/interactive028/run | 3 + .../interactive029/Issue834.idr | 0 .../{ => interactive}/interactive029/expected | 0 .../{ => interactive}/interactive029/input | 0 .../{ => interactive}/interactive029/run | 2 +- .../{ => interactive}/interactive030/expected | 0 .../{ => interactive}/interactive030/input | 0 tests/idris2/interactive/interactive030/run | 3 + .../interactive031/Signatures.idr | 0 .../{ => interactive}/interactive031/expected | 0 .../{ => interactive}/interactive031/input | 0 .../{ => interactive}/interactive031/run | 2 +- .../interactive032/Uninh.idr | 0 .../{ => interactive}/interactive032/expected | 0 .../{ => interactive}/interactive032/input | 0 .../{ => interactive}/interactive032/run | 2 +- .../interactive033/UninhIndent.idr | 0 .../{ => interactive}/interactive033/expected | 0 .../{ => interactive}/interactive033/input | 0 .../{ => interactive}/interactive033/run | 2 +- .../{ => interactive}/interactive034/expected | 0 .../{ => interactive}/interactive034/input | 0 .../{ => interactive}/interactive034/run | 2 +- .../interactive034/timeout.idr | 0 .../{ => interactive}/interactive035/expected | 0 .../{ => interactive}/interactive035/input | 0 .../{ => interactive}/interactive035/run | 2 +- .../interactive035/unify.idr | 0 .../interactive036/casefn.idr | 0 .../{ => interactive}/interactive036/expected | 0 .../{ => interactive}/interactive036/input | 0 .../{ => interactive}/interactive036/run | 2 +- .../interactive037/Holes.idr | 0 .../{ => interactive}/interactive037/expected | 0 .../{ => interactive}/interactive037/input | 0 .../{ => interactive}/interactive037/run | 2 +- .../interactive038/IEdit.idr | 0 .../{ => interactive}/interactive038/expected | 0 .../{ => interactive}/interactive038/input | 0 tests/idris2/interactive/interactive038/run | 3 + .../interactive039/CS_Syntax.idr | 0 .../{ => interactive}/interactive039/expected | 0 .../{ => interactive}/interactive039/input | 0 .../{ => interactive}/interactive039/run | 2 +- .../{ => interactive}/interactive040/expected | 0 .../{ => interactive}/interactive040/input | 0 tests/idris2/interactive/interactive040/run | 3 + .../interactive041/Issue1741.idr | 0 .../{ => interactive}/interactive041/expected | 0 .../{ => interactive}/interactive041/input | 0 .../{ => interactive}/interactive041/run | 2 +- .../interactive042/Issue35-2.idr | 0 .../interactive042/Issue35.idr | 0 .../{ => interactive}/interactive042/expected | 0 .../{ => interactive}/interactive042/input | 0 .../{ => interactive}/interactive042/run | 2 +- .../interactive043/ImplicitSplits.idr | 0 .../{ => interactive}/interactive043/expected | 0 .../{ => interactive}/interactive043/input | 0 .../{ => interactive}/interactive043/run | 2 +- .../interactive044/SplitShadow.idr | 0 .../{ => interactive}/interactive044/expected | 0 .../{ => interactive}/interactive044/input | 0 .../{ => interactive}/interactive044/run | 2 +- .../interactive045/Issue1742.idr | 0 .../{ => interactive}/interactive045/expected | 0 .../{ => interactive}/interactive045/input | 0 .../{ => interactive}/interactive045/run | 2 +- .../interactive046/Issue2712.idr | 0 .../{ => interactive}/interactive046/expected | 0 tests/idris2/interactive/interactive046/run | 3 + tests/idris2/interactive007/run | 3 - tests/idris2/interactive008/run | 3 - tests/idris2/interactive010/run | 3 - tests/idris2/interactive011/run | 3 - tests/idris2/interactive015/run | 3 - tests/idris2/interactive028/run | 3 - tests/idris2/interactive030/run | 3 - tests/idris2/interactive038/run | 3 - tests/idris2/interactive040/run | 3 - tests/idris2/interactive046/run | 3 - .../{ => interface}/interface001/IFace.idr | 0 .../{ => interface}/interface001/IFace1.idr | 0 .../{ => interface}/interface001/Stuff.idr | 0 .../{ => interface}/interface001/expected | 0 .../idris2/{ => interface}/interface001/input | 0 .../{ => interface}/interface001/input1 | 0 tests/idris2/{ => interface}/interface001/run | 2 +- .../{ => interface}/interface002/Functor.idr | 0 .../{ => interface}/interface002/Stuff.idr | 0 .../{ => interface}/interface002/expected | 0 .../idris2/{ => interface}/interface002/input | 0 tests/idris2/{ => interface}/interface002/run | 2 +- .../{ => interface}/interface003/Do.idr | 0 .../{ => interface}/interface003/expected | 0 .../idris2/{ => interface}/interface003/input | 0 tests/idris2/interface/interface003/run | 3 + .../{ => interface}/interface004/Do.idr | 0 .../{ => interface}/interface004/expected | 0 .../idris2/{ => interface}/interface004/input | 0 tests/idris2/interface/interface004/run | 3 + .../{ => interface}/interface005/Deps.idr | 0 .../{ => interface}/interface005/expected | 0 tests/idris2/interface/interface005/run | 3 + .../{ => interface}/interface006/Apply.idr | 0 .../interface006/Biapplicative.idr | 0 .../{ => interface}/interface006/Bimonad.idr | 0 .../{ => interface}/interface006/expected | 0 tests/idris2/interface/interface006/run | 3 + .../idris2/{ => interface}/interface007/A.idr | 0 .../{ => interface}/interface007/expected | 0 tests/idris2/{ => interface}/interface007/run | 2 +- .../{ => interface}/interface008/Deps.idr | 0 .../{ => interface}/interface008/expected | 0 tests/idris2/interface/interface008/run | 3 + .../{ => interface}/interface009/Odd.idr | 0 .../{ => interface}/interface009/expected | 0 .../idris2/{ => interface}/interface009/input | 0 tests/idris2/{ => interface}/interface009/run | 2 +- .../{ => interface}/interface010/Dep.idr | 0 .../{ => interface}/interface010/expected | 0 tests/idris2/interface/interface010/run | 3 + .../{ => interface}/interface011/FuncImpl.idr | 0 .../{ => interface}/interface011/expected | 0 tests/idris2/interface/interface011/run | 3 + .../{ => interface}/interface012/Defmeth.idr | 0 .../{ => interface}/interface012/expected | 0 tests/idris2/interface/interface012/run | 3 + .../{ => interface}/interface013/TypeInt.idr | 0 .../{ => interface}/interface013/expected | 0 tests/idris2/interface/interface013/run | 3 + .../{ => interface}/interface014/DepInt.idr | 0 .../{ => interface}/interface014/expected | 0 tests/idris2/interface/interface014/run | 3 + .../{ => interface}/interface015/expected | 0 .../{ => interface}/interface015/gnu.idr | 0 tests/idris2/interface/interface015/run | 3 + .../{ => interface}/interface016/TwoNum.idr | 0 .../{ => interface}/interface016/expected | 0 tests/idris2/interface/interface016/run | 3 + .../{ => interface}/interface017/Tricho.idr | 0 .../{ => interface}/interface017/expected | 0 tests/idris2/interface/interface017/run | 3 + .../{ => interface}/interface018/Sized.idr | 0 .../{ => interface}/interface018/Sized2.idr | 0 .../{ => interface}/interface018/Sized3.idr | 0 .../{ => interface}/interface018/expected | 0 .../idris2/{ => interface}/interface018/input | 0 tests/idris2/{ => interface}/interface018/run | 2 +- .../interface019/LocalHints.idr | 0 .../{ => interface}/interface019/expected | 0 tests/idris2/interface/interface019/run | 3 + .../interface020/LocalInterface.idr | 0 .../{ => interface}/interface020/expected | 0 tests/idris2/{ => interface}/interface020/run | 2 +- .../interface021/LocalHint.idr | 0 .../{ => interface}/interface021/expected | 0 tests/idris2/interface/interface021/run | 3 + .../interface022/DotMethod.idr | 0 .../{ => interface}/interface022/expected | 0 .../idris2/{ => interface}/interface022/input | 0 tests/idris2/{ => interface}/interface022/run | 2 +- .../{ => interface}/interface023/AppComp.idr | 0 .../{ => interface}/interface023/expected | 0 .../idris2/{ => interface}/interface023/input | 0 tests/idris2/{ => interface}/interface023/run | 2 +- .../{ => interface}/interface024/EH.idr | 0 .../{ => interface}/interface024/expected | 0 tests/idris2/{ => interface}/interface024/run | 2 +- .../interface025/AutoSearchHide1.idr | 0 .../interface025/AutoSearchHide2.idr | 0 .../{ => interface}/interface025/expected | 0 .../interface025}/input | 0 tests/idris2/{ => interface}/interface025/run | 2 +- .../interface026/UninhabitedRec.idr | 0 .../{ => interface}/interface026/expected | 0 .../interface026}/input | 0 tests/idris2/{ => interface}/interface026/run | 2 +- .../{ => interface}/interface027/expected | 0 .../idris2/{ => interface}/interface027/input | 0 .../{ => interface}/interface027/params.idr | 0 tests/idris2/{ => interface}/interface027/run | 2 +- .../interface028/InterfaceArgs.idr | 0 .../{ => interface}/interface028/expected | 0 tests/idris2/{ => interface}/interface028/run | 2 +- .../interface029/ForwardImpl1.idr | 0 .../interface029/ForwardImpl2.idr | 0 .../interface029/ForwardImpl3.idr | 0 .../{ => interface}/interface029/expected | 0 tests/idris2/{ => interface}/interface029/run | 2 +- tests/idris2/interface003/run | 3 - tests/idris2/interface004/run | 3 - tests/idris2/interface005/run | 3 - tests/idris2/interface006/run | 3 - tests/idris2/interface008/run | 3 - tests/idris2/interface010/run | 3 - tests/idris2/interface011/run | 3 - tests/idris2/interface012/run | 3 - tests/idris2/interface013/run | 3 - tests/idris2/interface014/run | 3 - tests/idris2/interface015/run | 3 - tests/idris2/interface016/run | 3 - tests/idris2/interface017/run | 3 - tests/idris2/interface019/run | 3 - tests/idris2/interface021/run | 3 - tests/idris2/interpolation003/run | 3 - tests/idris2/interpreter001/run | 3 - tests/idris2/interpreter003/run | 3 - tests/idris2/interpreter004/run | 3 - tests/idris2/interpreter006/run | 3 - tests/idris2/interpreter007/run | 3 - tests/idris2/{ => linear}/linear001/Door.idr | 0 tests/idris2/{ => linear}/linear001/Stuff.idr | 0 tests/idris2/{ => linear}/linear001/expected | 0 tests/idris2/{ => linear}/linear001/run | 2 +- tests/idris2/{ => linear}/linear002/Door.idr | 0 tests/idris2/{ => linear}/linear002/Stuff.idr | 0 tests/idris2/{ => linear}/linear002/expected | 0 tests/idris2/{ => linear}/linear002/input | 0 tests/idris2/{ => linear}/linear002/run | 2 +- .../idris2/{ => linear}/linear003/Linear.idr | 0 tests/idris2/{ => linear}/linear003/expected | 0 tests/idris2/{ => linear}/linear003/input | 0 tests/idris2/{ => linear}/linear003/run | 2 +- tests/idris2/{ => linear}/linear005/Door.idr | 0 .../idris2/{ => linear}/linear005/Linear.idr | 0 tests/idris2/{ => linear}/linear005/expected | 0 tests/idris2/{ => linear}/linear005/input | 0 tests/idris2/{ => linear}/linear005/run | 2 +- tests/idris2/{ => linear}/linear006/ZFun.idr | 0 tests/idris2/{ => linear}/linear006/expected | 0 tests/idris2/{ => linear}/linear006/input | 0 tests/idris2/{ => linear}/linear006/run | 2 +- tests/idris2/{ => linear}/linear007/LCase.idr | 0 tests/idris2/{ => linear}/linear007/expected | 0 tests/idris2/linear/linear007/run | 3 + tests/idris2/{ => linear}/linear008/Door.idr | 0 tests/idris2/{ => linear}/linear008/expected | 0 tests/idris2/linear/linear008/run | 3 + tests/idris2/{ => linear}/linear009/expected | 0 tests/idris2/{ => linear}/linear009/input | 0 tests/idris2/{ => linear}/linear009/qtt.idr | 0 tests/idris2/{ => linear}/linear009/run | 2 +- tests/idris2/{ => linear}/linear010/Door.idr | 0 tests/idris2/{ => linear}/linear010/expected | 0 tests/idris2/{ => linear}/linear010/run | 2 +- .../idris2/{ => linear}/linear011/Network.idr | 0 tests/idris2/{ => linear}/linear011/expected | 0 tests/idris2/{ => linear}/linear011/input | 0 tests/idris2/{ => linear}/linear011/run | 2 +- tests/idris2/{ => linear}/linear012/expected | 0 tests/idris2/{ => linear}/linear012/input | 0 .../{ => linear}/linear012/linholes.idr | 0 tests/idris2/{ => linear}/linear012/run | 2 +- .../{ => linear}/linear013/Issue758.idr | 0 tests/idris2/{ => linear}/linear013/expected | 0 tests/idris2/{ => linear}/linear013/input | 0 tests/idris2/{ => linear}/linear013/run | 2 +- .../idris2/{ => linear}/linear014/Issue55.idr | 0 tests/idris2/{ => linear}/linear014/expected | 0 tests/idris2/linear/linear014/run | 3 + .../{ => linear}/linear015/Issue1861.idr | 0 tests/idris2/{ => linear}/linear015/expected | 0 tests/idris2/linear/linear015/run | 3 + .../{ => linear}/linear016/Issue2895.idr | 0 .../{ => linear}/linear016/Issue2895_2.idr | 0 tests/idris2/{ => linear}/linear016/expected | 0 tests/idris2/{ => linear}/linear016/run | 2 +- tests/idris2/{ => linear}/linear017/As.idr | 0 tests/idris2/{ => linear}/linear017/expected | 0 tests/idris2/linear/linear017/run | 3 + tests/idris2/linear004/Erase.idr | 13 - tests/idris2/linear004/Stuff.idr | 68 --- tests/idris2/linear004/expected | 46 -- tests/idris2/linear004/input | 11 - tests/idris2/linear004/run | 3 - tests/idris2/linear007/run | 3 - tests/idris2/linear008/run | 3 - tests/idris2/linear014/run | 3 - tests/idris2/linear015/run | 3 - tests/idris2/linear017/run | 3 - .../{ => literate}/literate001/IEdit.lidr | 0 .../{ => literate}/literate001/expected | 0 tests/idris2/{ => literate}/literate001/input | 0 tests/idris2/{ => literate}/literate001/run | 2 +- .../{ => literate}/literate002/IEdit.lidr | 0 .../{ => literate}/literate002/IEdit2.lidr | 0 .../{ => literate}/literate002/expected | 0 tests/idris2/{ => literate}/literate002/input | 0 .../idris2/{ => literate}/literate002/input2 | 0 tests/idris2/{ => literate}/literate002/run | 2 +- .../{ => literate}/literate003/IEdit.lidr | 0 .../{ => literate}/literate003/expected | 0 tests/idris2/{ => literate}/literate003/input | 0 .../{literate005 => literate/literate003}/run | 2 +- .../{ => literate}/literate004/IEdit.lidr | 0 .../{ => literate}/literate004/expected | 0 tests/idris2/{ => literate}/literate004/input | 0 .../{literate003 => literate/literate004}/run | 2 +- .../{ => literate}/literate005/IEdit.lidr | 0 .../{ => literate}/literate005/expected | 0 tests/idris2/{ => literate}/literate005/input | 0 .../{literate004 => literate/literate005}/run | 2 +- .../{ => literate}/literate006/Door.lidr | 0 .../{ => literate}/literate006/expected | 0 tests/idris2/{ => literate}/literate006/input | 0 tests/idris2/{ => literate}/literate006/run | 2 +- .../{ => literate}/literate007/IEdit.lidr | 0 .../{ => literate}/literate007/IEdit.org | 0 .../{ => literate}/literate007/IEditOrg.org | 0 .../{ => literate}/literate007/expected | 0 tests/idris2/{ => literate}/literate007/input | 0 .../idris2/{ => literate}/literate007/input2 | 0 tests/idris2/{ => literate}/literate007/run | 2 +- .../{ => literate}/literate008/IEdit.lidr | 0 .../{ => literate}/literate008/expected | 0 tests/idris2/{ => literate}/literate008/input | 0 tests/idris2/literate/literate008/run | 3 + .../{ => literate}/literate009/WithLift.lidr | 0 .../{ => literate}/literate009/expected | 0 tests/idris2/{ => literate}/literate009/input | 0 tests/idris2/{ => literate}/literate009/run | 2 +- .../literate010/MyFirstIdrisProgram.org | 0 .../{ => literate}/literate010/expected | 0 tests/idris2/{ => literate}/literate010/run | 2 +- .../{ => literate}/literate011/IEdit.md | 0 .../{ => literate}/literate011/expected | 0 tests/idris2/{ => literate}/literate011/input | 0 tests/idris2/{ => literate}/literate011/run | 2 +- .../{ => literate}/literate012/IEdit.org | 0 .../{ => literate}/literate012/IEdit2.org | 0 .../{ => literate}/literate012/expected | 0 tests/idris2/{ => literate}/literate012/input | 0 tests/idris2/{ => literate}/literate012/run | 2 +- .../{ => literate}/literate013/Lit.lidr | 0 .../{ => literate}/literate013/LitTeX.tex | 0 .../{ => literate}/literate013/expected | 0 tests/idris2/{ => literate}/literate013/run | 2 +- .../{ => literate}/literate014/expected | 0 tests/idris2/{ => literate}/literate014/input | 0 tests/idris2/{ => literate}/literate014/run | 2 +- .../{ => literate}/literate014/with.lidr | 0 .../{ => literate}/literate015/case.lidr | 0 .../{ => literate}/literate015/expected | 0 tests/idris2/{ => literate}/literate015/input | 0 tests/idris2/{ => literate}/literate015/run | 2 +- .../{ => literate}/literate016/IEdit.org | 0 .../{ => literate}/literate016/IEdit2.org | 0 .../{ => literate}/literate016/expected | 0 tests/idris2/{ => literate}/literate016/input | 0 tests/idris2/{ => literate}/literate016/run | 2 +- .../{ => literate}/literate017/.gitignore | 0 .../{ => literate}/literate017/expected | 0 tests/idris2/{ => literate}/literate017/input | 0 .../literate017/project-expected.ipkg | 0 tests/idris2/{ => literate}/literate017/run | 2 +- .../{ => literate}/literate017/src/A/A.org | 0 .../{ => literate}/literate017/src/A/AB.lidr | 0 .../{ => literate}/literate017/src/B/B.md | 0 .../literate017/src/B/BA.markdown | 0 .../{ => literate}/literate017/src/Main.idr | 0 tests/idris2/literate008/run | 3 - tests/idris2/{ => misc}/docs001/expected | 0 tests/idris2/{ => misc}/docs001/input | 0 tests/idris2/misc/docs001/run | 3 + tests/idris2/{ => misc}/docs002/Doc.idr | 0 tests/idris2/{ => misc}/docs002/expected | 0 tests/idris2/{ => misc}/docs002/input | 0 tests/idris2/{ => misc}/docs002/run | 2 +- tests/idris2/{ => misc}/docs003/RecordDoc.idr | 0 tests/idris2/{ => misc}/docs003/expected | 0 tests/idris2/{ => misc}/docs003/input | 0 tests/idris2/{ => misc}/docs003/run | 2 +- tests/idris2/{ => misc}/docs004/DocImpl.idr | 0 tests/idris2/{ => misc}/docs004/List.idr | 0 tests/idris2/{ => misc}/docs004/expected | 0 tests/idris2/{ => misc}/docs004/input | 0 tests/idris2/{ => misc}/docs004/input2 | 0 tests/idris2/{ => misc}/docs004/run | 2 +- tests/idris2/{ => misc}/docs005/expected | 0 tests/idris2/{ => misc}/docs005/input | 0 tests/idris2/misc/docs005/run | 3 + tests/idris2/{ => misc}/eta001/Issue1370.idr | 0 tests/idris2/{ => misc}/eta001/expected | 0 tests/idris2/misc/eta001/run | 3 + .../{ => misc}/golden001/000-hello/expected | 0 .../idris2/{ => misc}/golden001/000-hello/run | 0 tests/idris2/{ => misc}/golden001/Main.idr | 0 tests/idris2/{ => misc}/golden001/Test.idr | 0 tests/idris2/{ => misc}/golden001/expected | 0 tests/idris2/{ => misc}/golden001/hello.ipkg | 0 tests/idris2/{ => misc}/golden001/run | 2 +- tests/idris2/{ => misc}/golden001/test.ipkg | 0 tests/idris2/{ => misc}/import001/Mult.idr | 0 tests/idris2/{ => misc}/import001/Nat.idr | 0 tests/idris2/{ => misc}/import001/Test.idr | 0 tests/idris2/{ => misc}/import001/expected | 0 tests/idris2/{ => misc}/import001/input | 0 tests/idris2/{ => misc}/import001/run | 2 +- tests/idris2/{ => misc}/import002/Mult.idr | 0 tests/idris2/{ => misc}/import002/Nat.idr | 0 tests/idris2/{ => misc}/import002/Test.idr | 0 tests/idris2/{ => misc}/import002/expected | 0 tests/idris2/{ => misc}/import002/run | 2 +- tests/idris2/{ => misc}/import003/A.idr | 0 tests/idris2/{ => misc}/import003/B.idr | 0 tests/idris2/{ => misc}/import003/C.idr | 0 tests/idris2/{ => misc}/import003/expected | 0 tests/idris2/{ => misc}/import003/input | 0 tests/idris2/{ => misc}/import003/run | 2 +- tests/idris2/{ => misc}/import004/Cycle1.idr | 0 tests/idris2/{ => misc}/import004/Cycle2.idr | 0 tests/idris2/{ => misc}/import004/Loop.idr | 0 tests/idris2/{ => misc}/import004/expected | 0 tests/idris2/{ => misc}/import004/run | 2 +- tests/idris2/{ => misc}/import005/As.idr | 0 tests/idris2/{ => misc}/import005/Test.idr | 0 tests/idris2/{ => misc}/import005/expected | 0 tests/idris2/{ => misc}/import005/input | 0 tests/idris2/misc/import005/run | 3 + tests/idris2/{ => misc}/import006/A/B.idr | 0 tests/idris2/{ => misc}/import006/A/C.idr | 0 tests/idris2/{ => misc}/import006/cyclic.ipkg | 0 tests/idris2/{ => misc}/import006/expected | 0 tests/idris2/{ => misc}/import006/run | 2 +- tests/idris2/{ => misc}/import007/Mod.idr | 0 tests/idris2/{ => misc}/import007/Mod1.idr | 0 tests/idris2/{ => misc}/import007/Mod2.idr | 0 tests/idris2/{ => misc}/import007/expected | 0 tests/idris2/misc/import007/run | 3 + tests/idris2/{ => misc}/import008/Exe/Mod.idr | 0 .../{ => misc}/import008/Exe/depends/lib-0 | 0 .../idris2/{ => misc}/import008/Exe/exe.ipkg | 0 .../idris2/{ => misc}/import008/Lib/Conf.idr | 0 .../idris2/{ => misc}/import008/Lib/lib.ipkg | 0 tests/idris2/{ => misc}/import008/expected | 0 tests/idris2/{ => misc}/import008/run | 2 +- tests/idris2/{ => misc}/import009/Import.idr | 0 tests/idris2/{ => misc}/import009/Infix.idr | 0 tests/idris2/{ => misc}/import009/Prefix.idr | 0 tests/idris2/{ => misc}/import009/Resugar.idr | 0 tests/idris2/{ => misc}/import009/Test.idr | 0 tests/idris2/{ => misc}/import009/expected | 0 tests/idris2/{ => misc}/import009/run | 2 +- .../{ => misc}/inlining001/Inlining.idr | 0 tests/idris2/{ => misc}/inlining001/expected | 0 tests/idris2/{ => misc}/inlining001/input | 0 tests/idris2/{ => misc}/inlining001/run | 2 +- tests/idris2/{ => misc}/lazy001/Lazy.idr | 0 tests/idris2/{ => misc}/lazy001/expected | 0 tests/idris2/{ => misc}/lazy001/input | 0 tests/idris2/{ => misc}/lazy001/run | 2 +- .../idris2/{ => misc}/lazy002/LazyFoldlM.idr | 0 tests/idris2/{ => misc}/lazy002/expected | 0 tests/idris2/{ => misc}/lazy002/input | 0 tests/idris2/{ => misc}/lazy002/run | 2 +- tests/idris2/{ => misc}/lazy003/DelayLam.idr | 0 tests/idris2/{ => misc}/lazy003/expected | 0 tests/idris2/{ => misc}/lazy003/input | 0 tests/idris2/{ => misc}/lazy003/run | 2 +- tests/idris2/{ => misc}/namespace001/Dup.idr | 0 .../idris2/{ => misc}/namespace001/Scope.idr | 0 tests/idris2/{ => misc}/namespace001/expected | 0 tests/idris2/{ => misc}/namespace001/run | 2 +- .../{ => misc}/namespace002/Issue1313.idr | 0 tests/idris2/{ => misc}/namespace002/expected | 0 tests/idris2/{ => misc}/namespace002/run | 2 +- tests/idris2/{ => misc}/namespace003/Test.idr | 0 tests/idris2/{ => misc}/namespace003/expected | 0 tests/idris2/{ => misc}/namespace003/run | 2 +- .../idris2/{ => misc}/namespace004/Export.idr | 0 tests/idris2/{ => misc}/namespace004/expected | 0 tests/idris2/misc/namespace004/run | 3 + tests/idris2/{ => misc}/namespace005/Lib1.idr | 0 tests/idris2/{ => misc}/namespace005/Lib2.idr | 0 .../{ => misc}/namespace005/LibPre1.idr | 0 .../{ => misc}/namespace005/LibPre2.idr | 0 .../idris2/{ => misc}/namespace005/Main0.idr | 0 .../idris2/{ => misc}/namespace005/Main1.idr | 0 .../idris2/{ => misc}/namespace005/Main3.idr | 0 .../{ => misc}/namespace005/MainConflict.idr | 0 .../{ => misc}/namespace005/MainFail.idr | 0 .../{ => misc}/namespace005/MainPre0.idr | 0 .../{ => misc}/namespace005/MainPre1.idr | 0 .../{ => misc}/namespace005/NonConflict1.idr | 0 .../{ => misc}/namespace005/NonConflict2.idr | 0 tests/idris2/{ => misc}/namespace005/expected | 0 tests/idris2/{ => misc}/namespace005/run | 2 +- tests/idris2/{ => misc}/params001/expected | 0 tests/idris2/{ => misc}/params001/param.idr | 0 .../idris2/{ => misc}/params001/parambad.idr | 0 tests/idris2/{ => misc}/params001/run | 2 +- .../{ => misc}/params002/ParamsPrint.idr | 0 tests/idris2/{ => misc}/params002/expected | 0 tests/idris2/{ => misc}/params002/input | 0 tests/idris2/{ => misc}/params002/run | 2 +- .../idris2/{ => misc}/params003/casesplit.idr | 0 tests/idris2/{ => misc}/params003/expected | 0 tests/idris2/{ => misc}/params003/input | 0 tests/idris2/{ => misc}/params003/run | 2 +- .../{ => misc}/pretty001/Issue1328A.idr | 0 tests/idris2/{ => misc}/pretty001/expected | 0 tests/idris2/{ => misc}/pretty001/input | 0 tests/idris2/{ => misc}/pretty001/run | 2 +- tests/idris2/{ => misc}/pretty002/expected | 0 tests/idris2/{ => misc}/pretty002/input | 0 tests/idris2/misc/pretty002/run | 3 + tests/idris2/{ => misc}/primloop/PrimLoop.idr | 0 tests/idris2/{ => misc}/primloop/expected | 0 tests/idris2/misc/primloop/run | 3 + .../quantifiers001/TestQuantifiers.idr | 0 .../idris2/{ => misc}/quantifiers001/expected | 0 tests/idris2/{ => misc}/quantifiers001/input | 0 tests/idris2/{ => misc}/quantifiers001/run | 2 +- tests/idris2/{ => misc}/real001/Channel.idr | 0 tests/idris2/{ => misc}/real001/Linear.idr | 0 tests/idris2/{ => misc}/real001/MakeChans.idr | 0 tests/idris2/{ => misc}/real001/TestProto.idr | 0 tests/idris2/{ => misc}/real001/expected | 0 tests/idris2/{ => misc}/real001/run | 2 +- tests/idris2/{ => misc}/real002/Store.idr | 0 tests/idris2/{ => misc}/real002/StoreL.idr | 0 tests/idris2/{ => misc}/real002/expected | 0 tests/idris2/{ => misc}/real002/run | 2 +- .../{ => misc}/unification001/Issue647.idr | 0 .../idris2/{ => misc}/unification001/expected | 0 tests/idris2/misc/unification001/run | 3 + tests/idris2/{ => misc}/with003/Main.idr | 0 tests/idris2/{ => misc}/with003/expected | 0 tests/idris2/{ => misc}/with003/input | 0 tests/idris2/{ => misc}/with003/run | 2 +- tests/idris2/namespace004/run | 3 - tests/idris2/{ => perf}/perf001/Big.idr | 0 tests/idris2/{ => perf}/perf001/expected | 0 tests/idris2/perf/perf001/run | 3 + tests/idris2/{ => perf}/perf002/Big.idr | 0 tests/idris2/{ => perf}/perf002/expected | 0 tests/idris2/perf/perf002/run | 3 + tests/idris2/{ => perf}/perf003/Auto.idr | 0 tests/idris2/{ => perf}/perf003/expected | 0 tests/idris2/perf/perf003/run | 3 + tests/idris2/{ => perf}/perf004/bigdpair.idr | 0 tests/idris2/{ => perf}/perf004/expected | 0 tests/idris2/perf/perf004/run | 3 + tests/idris2/{ => perf}/perf005/Bad1.idr | 0 tests/idris2/{ => perf}/perf005/Bad2.idr | 0 tests/idris2/{ => perf}/perf005/Bad3.idr | 0 tests/idris2/{ => perf}/perf005/Lambda.idr | 0 .../{ => perf}/perf005/NoRegression.idr | 0 tests/idris2/{ => perf}/perf005/expected | 0 tests/idris2/{ => perf}/perf005/run | 2 +- tests/idris2/{ => perf}/perf007/Slooow.idr | 0 tests/idris2/{ => perf}/perf007/expected | 0 tests/idris2/perf/perf007/run | 3 + tests/idris2/{ => perf}/perf008/FinPerf.idr | 0 tests/idris2/{ => perf}/perf008/expected | 0 tests/idris2/perf/perf008/run | 3 + tests/idris2/{ => perf}/perf009/A.idr | 0 tests/idris2/{ => perf}/perf009/B.idr | 0 tests/idris2/{ => perf}/perf009/C.idr | 0 tests/idris2/{ => perf}/perf009/expected | 0 tests/idris2/perf/perf009/run | 3 + tests/idris2/{ => perf}/perf010/Printf.idr | 0 tests/idris2/{ => perf}/perf010/expected | 0 tests/idris2/perf/perf010/run | 3 + tests/idris2/{ => perf}/perf011/A.idr | 0 tests/idris2/{ => perf}/perf011/B.idr | 0 tests/idris2/{ => perf}/perf011/C.idr | 0 tests/idris2/{ => perf}/perf011/expected | 0 tests/idris2/perf/perf011/run | 3 + tests/idris2/{ => perf}/perf012/Main.idr | 0 tests/idris2/{ => perf}/perf012/expected | 0 tests/idris2/{ => perf}/perf012/run | 2 +- tests/idris2/{ => perf}/perf2202/Church.idr | 0 tests/idris2/{ => perf}/perf2202/expected | 0 .../idris2/{ => perf}/perf2202/many10000.idr | 0 tests/idris2/{ => perf}/perf2202/run | 2 +- tests/idris2/perf001/run | 3 - tests/idris2/perf002/run | 3 - tests/idris2/perf003/run | 3 - tests/idris2/perf004/run | 3 - tests/idris2/perf007/run | 3 - tests/idris2/perf008/run | 3 - tests/idris2/perf009/run | 3 - tests/idris2/perf010/run | 3 - tests/idris2/perf011/run | 3 - tests/idris2/perror001/run | 3 - tests/idris2/perror002/run | 3 - tests/idris2/perror004/run | 3 - tests/idris2/perror005/run | 3 - tests/idris2/perror006/run | 3 - tests/idris2/perror019/run | 3 - tests/idris2/perror021/run | 3 - tests/idris2/perror022/run | 3 - tests/idris2/perror023/run | 3 - tests/idris2/perror024/run | 3 - tests/idris2/perror025/run | 3 - tests/idris2/perror026/run | 3 - tests/idris2/perror027/run | 3 - tests/idris2/perror028/run | 3 - tests/idris2/perror029/run | 3 - tests/idris2/{ => pkg}/pkg001/Dummy.idr | 0 tests/idris2/{ => pkg}/pkg001/dummy.ipkg | 0 tests/idris2/{ => pkg}/pkg001/expected | 0 tests/idris2/{ => pkg}/pkg001/run | 2 +- tests/idris2/{ => pkg}/pkg002/dummy.ipkg | 0 tests/idris2/{ => pkg}/pkg002/expected | 0 tests/idris2/{ => pkg}/pkg002/run | 2 +- .../idris2/{ => pkg}/pkg002/src/Top/Dummy.idr | 0 tests/idris2/{ => pkg}/pkg003/Main.idr | 0 tests/idris2/{ => pkg}/pkg003/expected | 0 tests/idris2/{ => pkg}/pkg003/run | 2 +- tests/idris2/{ => pkg}/pkg003/testpkg.ipkg | 0 tests/idris2/{ => pkg}/pkg004/Dummy.idr | 0 tests/idris2/{ => pkg}/pkg004/dummy.ipkg | 0 tests/idris2/{ => pkg}/pkg004/expected | 0 tests/idris2/{ => pkg}/pkg004/input | 0 tests/idris2/{ => pkg}/pkg004/run | 2 +- tests/idris2/{ => pkg}/pkg005/Foo.idr | 0 tests/idris2/{ => pkg}/pkg005/expected | 0 tests/idris2/{ => pkg}/pkg005/foo.ipkg | 0 tests/idris2/{ => pkg}/pkg005/input | 0 tests/idris2/{ => pkg}/pkg005/run | 2 +- .../pkg006/depends/bar-1.0.1/bar.ipkg | 0 .../pkg006/depends/bar-baz/bar-baz.ipkg | 0 .../{ => pkg}/pkg006/depends/foo-0.5/foo.ipkg | 0 .../pkg006/depends/foo-bar-1.3.1/foo-bar.ipkg | 0 .../{ => pkg}/pkg006/depends/quux/quux.ipkg | 0 tests/idris2/{ => pkg}/pkg006/expected | 0 tests/idris2/{ => pkg}/pkg006/run | 2 +- tests/idris2/{ => pkg}/pkg006/test1.ipkg | 0 tests/idris2/{ => pkg}/pkg006/test2.ipkg | 0 tests/idris2/{ => pkg}/pkg006/test3.ipkg | 0 tests/idris2/{ => pkg}/pkg006/test4.ipkg | 0 tests/idris2/{ => pkg}/pkg006/test5.ipkg | 0 .../pkg007/A/Path/Of/Dires/First.idr | 0 .../pkg007/A/Path/Of/Dires/Second.idr | 0 .../{ => pkg}/pkg007/Another/Fourth.idr | 0 .../{ => pkg}/pkg007/Another/One/Third.idr | 0 tests/idris2/{ => pkg}/pkg007/expected | 0 tests/idris2/{ => pkg}/pkg007/input | 0 tests/idris2/{ => pkg}/pkg007/input2 | 0 tests/idris2/{ => pkg}/pkg007/run | 2 +- .../{ => pkg}/pkg007/src/And/A/Proof.idr | 0 .../{ => pkg}/pkg007/src/Yet/Another/Path.idr | 0 tests/idris2/{ => pkg}/pkg008/Bar.idr | 0 tests/idris2/{ => pkg}/pkg008/Foo.idr | 0 tests/idris2/{ => pkg}/pkg008/bar.ipkg | 0 tests/idris2/{ => pkg}/pkg008/expected | 0 tests/idris2/{ => pkg}/pkg008/foo.ipkg | 0 tests/idris2/{ => pkg}/pkg008/run | 2 +- tests/idris2/{ => pkg}/pkg009/expected | 0 tests/idris2/{ => pkg}/pkg009/run | 2 +- .../idris2/{ => pkg}/pkg009/testpkg/Main.idr | 0 .../{ => pkg}/pkg009/testpkg/testpkg.ipkg | 0 tests/idris2/{ => pkg}/pkg010/Main.idr | 0 tests/idris2/{ => pkg}/pkg010/expected | 0 tests/idris2/{ => pkg}/pkg010/run | 2 +- tests/idris2/{ => pkg}/pkg010/testpkg.ipkg | 0 .../{ => pkg}/pkg011/dot-slash-dot/Main.idr | 0 .../pkg011/dot-slash-dot/testpkg.ipkg | 0 .../pkg011/dot-slash-name-slash/src/Main.idr | 0 .../pkg011/dot-slash-name-slash/testpkg.ipkg | 0 .../{ => pkg}/pkg011/dot-slash/Main.idr | 0 .../{ => pkg}/pkg011/dot-slash/testpkg.ipkg | 0 tests/idris2/{ => pkg}/pkg011/dot/Main.idr | 0 .../idris2/{ => pkg}/pkg011/dot/testpkg.ipkg | 0 tests/idris2/{ => pkg}/pkg011/expected | 0 tests/idris2/{ => pkg}/pkg011/run | 2 +- .../{ => pkg}/pkg011/sibling/pkg/testpkg.ipkg | 0 .../{ => pkg}/pkg011/sibling/src/Main.idr | 0 tests/idris2/{ => pkg}/pkg012/Main.idr | 0 tests/idris2/{ => pkg}/pkg012/expected | 0 tests/idris2/{ => pkg}/pkg012/run | 2 +- tests/idris2/{ => pkg}/pkg012/testpkg.ipkg | 0 .../pkg013/depends/bar-0.7.2/bar.ipkg | 0 .../pkg013/depends/foo-0.1.0/foo.ipkg | 0 tests/idris2/{ => pkg}/pkg013/expected | 0 tests/idris2/{ => pkg}/pkg013/run | 2 +- tests/idris2/{ => pkg}/pkg013/test.ipkg | 0 .../pkg014/depends/bar-0.1.0/bar.ipkg | 0 .../pkg014/depends/bar-0.1.1/bar.ipkg | 0 .../pkg014/depends/baz-0.1.0/baz.ipkg | 0 .../pkg014/depends/baz-0.2.0/baz.ipkg | 0 .../pkg014/depends/baz-0.3.0/baz.ipkg | 0 .../pkg014/depends/foo-0.1.0/foo.ipkg | 0 .../pkg014/depends/foo-0.2.0/foo.ipkg | 0 .../pkg014/depends/foo-0.3.0/foo.ipkg | 0 tests/idris2/{ => pkg}/pkg014/expected | 0 tests/idris2/{ => pkg}/pkg014/run | 2 +- tests/idris2/{ => pkg}/pkg014/test.ipkg | 0 .../pkg015/depends/bar-0.1.0/bar.ipkg | 0 .../pkg015/depends/bar-0.1.1/bar.ipkg | 0 .../pkg015/depends/baz-0.1.0/baz.ipkg | 0 .../pkg015/depends/baz-0.2.0/baz.ipkg | 0 .../pkg015/depends/baz-0.3.0/baz.ipkg | 0 .../pkg015/depends/foo-0.1.0/foo.ipkg | 0 .../pkg015/depends/foo-0.2.0/foo.ipkg | 0 .../pkg015/depends/foo-0.3.0/foo.ipkg | 0 .../pkg015/depends/prz-0.1.0/prz.ipkg | 0 tests/idris2/{ => pkg}/pkg015/expected | 0 tests/idris2/{ => pkg}/pkg015/run | 0 tests/idris2/{ => pkg}/pkg015/test.ipkg | 0 tests/idris2/{ => pkg}/pkg016/bar.ipkg | 0 tests/idris2/{ => pkg}/pkg016/bar/Bar.idr | 0 tests/idris2/{ => pkg}/pkg016/baz.ipkg | 0 tests/idris2/{ => pkg}/pkg016/baz/Baz.idr | 0 tests/idris2/{ => pkg}/pkg016/expected | 0 tests/idris2/{ => pkg}/pkg016/foo.ipkg | 0 tests/idris2/{ => pkg}/pkg016/foo/Foo.idr | 0 tests/idris2/{ => pkg}/pkg016/run | 0 tests/idris2/{ => pkg}/pkg016/src/Test.idr | 0 tests/idris2/{ => pkg}/pkg016/test.ipkg | 0 tests/idris2/{ => pkg}/pkg017/a1/A.idr | 0 tests/idris2/{ => pkg}/pkg017/a1/a1.ipkg | 0 tests/idris2/{ => pkg}/pkg017/a2/A.idr | 0 tests/idris2/{ => pkg}/pkg017/a2/a2.ipkg | 0 tests/idris2/{ => pkg}/pkg017/b1/b1.ipkg | 0 tests/idris2/{ => pkg}/pkg017/b1/src/B1.idr | 0 tests/idris2/{ => pkg}/pkg017/b2/b2.ipkg | 0 tests/idris2/{ => pkg}/pkg017/b2/src/B2.idr | 0 tests/idris2/{ => pkg}/pkg017/expected | 0 tests/idris2/{ => pkg}/pkg017/input1 | 0 tests/idris2/{ => pkg}/pkg017/input2 | 0 tests/idris2/{ => pkg}/pkg017/run | 2 +- tests/idris2/positivity001/run | 3 - tests/idris2/positivity002/run | 3 - tests/idris2/positivity003/run | 3 - tests/idris2/pretty002/run | 3 - tests/idris2/primloop/run | 3 - tests/idris2/record003/run | 3 - tests/idris2/record010/run | 3 - tests/idris2/record011/run | 3 - tests/idris2/record012/run | 3 - tests/idris2/record014/run | 3 - tests/idris2/record015/run | 3 - tests/idris2/record018/run | 3 - .../{ => reflection}/reflection001/expected | 0 .../{ => reflection}/reflection001/input | 0 .../{ => reflection}/reflection001/quote.idr | 0 .../idris2/{ => reflection}/reflection001/run | 2 +- .../{ => reflection}/reflection002/expected | 0 .../{ => reflection}/reflection002/input | 0 .../{ => reflection}/reflection002/power.idr | 0 .../idris2/{ => reflection}/reflection002/run | 2 +- .../{ => reflection}/reflection003/expected | 0 .../reflection003/refprims.idr | 0 tests/idris2/reflection/reflection003/run | 3 + .../{ => reflection}/reflection004/expected | 0 .../{ => reflection}/reflection004/input | 0 .../reflection004/refdecl.idr | 0 .../idris2/{ => reflection}/reflection004/run | 2 +- .../{ => reflection}/reflection005/expected | 0 .../{ => reflection}/reflection005/input | 0 .../reflection005/refdecl.idr | 0 .../idris2/{ => reflection}/reflection005/run | 2 +- .../{ => reflection}/reflection006/expected | 0 .../{ => reflection}/reflection006/input | 0 .../{ => reflection}/reflection006/refleq.idr | 0 tests/idris2/reflection/reflection006/run | 3 + .../reflection007/NatExpr.idr | 0 .../{ => reflection}/reflection007/expected | 0 .../{ => reflection}/reflection007/input | 0 .../idris2/{ => reflection}/reflection007/run | 2 +- .../{ => reflection}/reflection008/Interp.idr | 0 .../{ => reflection}/reflection008/expected | 0 .../{ => reflection}/reflection008/input | 0 .../idris2/{ => reflection}/reflection008/run | 2 +- .../{ => reflection}/reflection009/expected | 0 .../{ => reflection}/reflection009/perf.idr | 0 tests/idris2/reflection/reflection009/run | 3 + .../{ => reflection}/reflection010/Name.idr | 0 .../{ => reflection}/reflection010/expected | 0 .../idris2/{ => reflection}/reflection010/run | 2 +- .../{ => reflection}/reflection011/expected | 0 .../{ => reflection}/reflection011/input | 0 .../idris2/{ => reflection}/reflection011/run | 2 +- .../{ => reflection}/reflection011/tryref.idr | 0 .../{ => reflection}/reflection012/expected | 0 .../{ => reflection}/reflection012/input | 0 .../reflection012/nameinfo.idr | 0 .../idris2/{ => reflection}/reflection012/run | 2 +- .../reflection013/WithUnambig.idr | 0 .../{ => reflection}/reflection013/expected | 0 tests/idris2/reflection/reflection013/run | 3 + .../{ => reflection}/reflection014/expected | 0 .../reflection014/refdecl.idr | 0 tests/idris2/reflection/reflection014/run | 3 + .../reflection015/MacroRetFunc.idr | 0 .../{ => reflection}/reflection015/expected | 0 .../idris2/{ => reflection}/reflection015/run | 2 +- .../reflection016/BindElabScBug.idr | 0 .../{ => reflection}/reflection016/Eta.idr | 0 .../{ => reflection}/reflection016/expected | 0 .../{ => reflection}/reflection016/input | 0 .../idris2/{ => reflection}/reflection016/run | 2 +- .../reflection017/CanElabType.idr | 0 .../reflection017/StillCantEscape.idr | 0 .../{ => reflection}/reflection017/expected | 0 .../idris2/{ => reflection}/reflection017/run | 2 +- .../reflection018/AtTypeLevel.idr | 0 .../{ => reflection}/reflection018/expected | 0 tests/idris2/reflection/reflection018/run | 3 + .../reflection019/ElabScriptWarning.idr | 0 .../{ => reflection}/reflection019/expected | 0 .../{ => reflection}/reflection019/input | 0 .../idris2/{ => reflection}/reflection019/run | 2 +- .../{ => reflection}/reflection019/test.ipkg | 0 .../reflection020/FromDecls.idr | 0 .../reflection020/FromName.idr | 0 .../reflection020/FromTTImp.idr | 0 .../{ => reflection}/reflection020/expected | 0 .../idris2/{ => reflection}/reflection020/run | 2 +- .../reflection021/QuoteSearch.idr | 0 .../{ => reflection}/reflection021/expected | 0 tests/idris2/reflection/reflection021/run | 3 + tests/idris2/reflection003/run | 3 - tests/idris2/reflection006/run | 3 - tests/idris2/reflection009/run | 3 - tests/idris2/reflection013/run | 3 - tests/idris2/reflection014/run | 3 - tests/idris2/reflection018/run | 3 - tests/idris2/reflection021/run | 3 - tests/idris2/{ => reg}/reg001/D.idr | 0 tests/idris2/{ => reg}/reg001/expected | 0 tests/idris2/reg/reg001/run | 3 + tests/idris2/{ => reg}/reg002/expected | 0 tests/idris2/{ => reg}/reg002/linm.idr | 0 tests/idris2/reg/reg002/run | 3 + tests/idris2/{ => reg}/reg003/Holes.idr | 0 tests/idris2/{ => reg}/reg003/expected | 0 tests/idris2/{ => reg}/reg003/input | 0 tests/idris2/{ => reg}/reg003/run | 2 +- tests/idris2/{ => reg}/reg004/ambig.idr | 0 tests/idris2/{ => reg}/reg004/expected | 0 tests/idris2/reg/reg004/run | 3 + tests/idris2/{ => reg}/reg005/expected | 0 tests/idris2/{ => reg}/reg005/iftype.idr | 0 tests/idris2/reg/reg005/run | 3 + tests/idris2/{ => reg}/reg006/Cmd.idr | 0 tests/idris2/{ => reg}/reg006/expected | 0 tests/idris2/reg/reg006/run | 3 + tests/idris2/{ => reg}/reg007/Main.idr | 0 tests/idris2/{ => reg}/reg007/expected | 0 tests/idris2/reg/reg007/run | 3 + tests/idris2/{ => reg}/reg008/Vending.idr | 0 tests/idris2/{ => reg}/reg008/expected | 0 tests/idris2/{ => reg}/reg008/input | 0 tests/idris2/{ => reg}/reg008/run | 2 +- tests/idris2/{ => reg}/reg009/Case.idr | 0 tests/idris2/{ => reg}/reg009/expected | 0 tests/idris2/reg/reg009/run | 3 + tests/idris2/{ => reg}/reg010/Recordname.idr | 0 tests/idris2/{ => reg}/reg010/expected | 0 tests/idris2/reg/reg010/run | 3 + tests/idris2/{ => reg}/reg011/expected | 0 tests/idris2/{ => reg}/reg011/mut.idr | 0 tests/idris2/reg/reg011/run | 3 + tests/idris2/{ => reg}/reg012/Foo.idr | 0 tests/idris2/{ => reg}/reg012/expected | 0 tests/idris2/reg/reg012/run | 3 + .../{ => reg}/reg013/UnboundImplicits.idr | 0 tests/idris2/{ => reg}/reg013/expected | 0 tests/idris2/{ => reg}/reg013/run | 2 +- tests/idris2/{ => reg}/reg014/casecase.idr | 0 tests/idris2/{ => reg}/reg014/expected | 0 tests/idris2/{ => reg}/reg014/run | 2 +- tests/idris2/{ => reg}/reg015/anyfail.idr | 0 tests/idris2/{ => reg}/reg015/expected | 0 tests/idris2/reg/reg015/run | 3 + tests/idris2/{ => reg}/reg016/Using.idr | 0 tests/idris2/{ => reg}/reg016/expected | 0 tests/idris2/reg/reg016/run | 3 + tests/idris2/{ => reg}/reg017/expected | 0 tests/idris2/{ => reg}/reg017/lammult.idr | 0 tests/idris2/reg/reg017/run | 3 + tests/idris2/{ => reg}/reg018/cycle.idr | 0 tests/idris2/{ => reg}/reg018/expected | 0 tests/idris2/reg/reg018/run | 3 + tests/idris2/{ => reg}/reg019/expected | 0 tests/idris2/{ => reg}/reg019/lazybug.idr | 0 tests/idris2/reg/reg019/run | 3 + tests/idris2/{ => reg}/reg020/expected | 0 tests/idris2/{ => reg}/reg020/input | 0 tests/idris2/{ => reg}/reg020/matchlits.idr | 0 tests/idris2/{ => reg}/reg020/run | 2 +- tests/idris2/{ => reg}/reg021/case.idr | 0 tests/idris2/{ => reg}/reg021/expected | 0 tests/idris2/reg/reg021/run | 3 + tests/idris2/{ => reg}/reg022/case.idr | 0 tests/idris2/{ => reg}/reg022/expected | 0 tests/idris2/{ => reg}/reg022/input | 0 tests/idris2/reg/reg022/run | 3 + tests/idris2/{ => reg}/reg023/boom.idr | 0 tests/idris2/{ => reg}/reg023/expected | 0 tests/idris2/reg/reg023/run | 3 + tests/idris2/{ => reg}/reg024/expected | 0 tests/idris2/{ => reg}/reg024/input | 0 tests/idris2/{ => reg}/reg024/run | 2 +- tests/idris2/{ => reg}/reg024/split.idr | 0 tests/idris2/{ => reg}/reg025/expected | 0 tests/idris2/{ => reg}/reg025/input | 0 tests/idris2/{ => reg}/reg025/lift.idr | 0 tests/idris2/{ => reg}/reg025/run | 2 +- tests/idris2/{ => reg}/reg026/Meh.idr | 0 tests/idris2/{ => reg}/reg026/expected | 0 tests/idris2/reg/reg026/run | 3 + tests/idris2/{ => reg}/reg027/expected | 0 tests/idris2/{ => reg}/reg027/pwhere.idr | 0 tests/idris2/reg/reg027/run | 3 + tests/idris2/{ => reg}/reg028/Test.idr | 0 tests/idris2/{ => reg}/reg028/expected | 0 tests/idris2/reg/reg028/run | 3 + tests/idris2/{ => reg}/reg029/expected | 0 tests/idris2/{ => reg}/reg029/lqueue.idr | 0 tests/idris2/reg/reg029/run | 3 + tests/idris2/{ => reg}/reg030/A.idr | 0 tests/idris2/{ => reg}/reg030/B.idr | 0 tests/idris2/{ => reg}/reg030/C.idr | 0 tests/idris2/{ => reg}/reg030/expected | 0 tests/idris2/reg/reg030/run | 3 + tests/idris2/{ => reg}/reg031/dpair.idr | 0 tests/idris2/{ => reg}/reg031/expected | 0 tests/idris2/reg/reg031/run | 3 + tests/idris2/{ => reg}/reg032/expected | 0 tests/idris2/{ => reg}/reg032/recupdate.idr | 0 tests/idris2/reg/reg032/run | 3 + tests/idris2/{ => reg}/reg033/DerivingEq.idr | 0 tests/idris2/{ => reg}/reg033/expected | 0 tests/idris2/reg/reg033/run | 3 + tests/idris2/{ => reg}/reg033/test.idr | 0 tests/idris2/{ => reg}/reg034/expected | 0 tests/idris2/reg/reg034/run | 3 + tests/idris2/{ => reg}/reg034/void.idr | 0 tests/idris2/{ => reg}/reg035/Implicit.idr | 0 tests/idris2/{ => reg}/reg035/expected | 0 tests/idris2/reg/reg035/run | 3 + tests/idris2/{ => reg}/reg036/Test.idr | 0 tests/idris2/{ => reg}/reg036/expected | 0 tests/idris2/reg/reg036/run | 3 + tests/idris2/{ => reg}/reg037/Test.idr | 0 tests/idris2/{ => reg}/reg037/expected | 0 tests/idris2/reg/reg037/run | 3 + tests/idris2/{ => reg}/reg038/Test1.idr | 0 tests/idris2/{ => reg}/reg038/Test2.idr | 0 tests/idris2/{ => reg}/reg038/expected | 0 tests/idris2/{ => reg}/reg038/run | 2 +- tests/idris2/{ => reg}/reg039/dupdup.idr | 0 tests/idris2/{ => reg}/reg039/expected | 0 tests/idris2/reg/reg039/run | 3 + tests/idris2/{ => reg}/reg040/CoverBug.idr | 0 tests/idris2/{ => reg}/reg040/expected | 0 tests/idris2/reg/reg040/run | 3 + tests/idris2/{ => reg}/reg041/expected | 0 tests/idris2/reg/reg041/run | 3 + tests/idris2/{ => reg}/reg041/tuple.idr | 0 tests/idris2/{ => reg}/reg042/NatOpts.idr | 0 tests/idris2/{ => reg}/reg042/expected | 0 tests/idris2/{ => reg}/reg042/input | 0 tests/idris2/{ => reg}/reg042/run | 2 +- tests/idris2/{ => reg}/reg043/NotFake.idr | 0 tests/idris2/{ => reg}/reg043/TestFake.idr | 0 tests/idris2/{ => reg}/reg043/expected | 0 tests/idris2/{ => reg}/reg043/run | 2 +- tests/idris2/{ => reg}/reg044/Methods.idr | 0 tests/idris2/{ => reg}/reg044/expected | 0 tests/idris2/reg/reg044/run | 3 + tests/idris2/{ => reg}/reg045/expected | 0 tests/idris2/reg/reg045/run | 3 + tests/idris2/{ => reg}/reg045/withparams.idr | 0 tests/idris2/{ => reg}/reg046/Postpone.idr | 0 tests/idris2/{ => reg}/reg046/expected | 0 tests/idris2/{ => reg}/reg046/input | 0 tests/idris2/{ => reg}/reg046/run | 2 +- .../{ => reg}/reg047/QualifiedDoBang.idr | 0 tests/idris2/{ => reg}/reg047/expected | 0 tests/idris2/{ => reg}/reg047/input | 0 tests/idris2/{ => reg}/reg047/run | 2 +- tests/idris2/{ => reg}/reg048/expected | 0 tests/idris2/{ => reg}/reg048/inferror.idr | 0 tests/idris2/{ => reg}/reg048/run | 2 +- tests/idris2/{ => reg}/reg049/expected | 0 tests/idris2/{ => reg}/reg049/lettype.idr | 0 tests/idris2/{ => reg}/reg049/run | 2 +- tests/idris2/{ => reg}/reg050/expected | 0 tests/idris2/{ => reg}/reg050/loopy.idr | 0 tests/idris2/{ => reg}/reg050/run | 2 +- tests/idris2/{ => reg}/reg051/BigFins.idr | 0 tests/idris2/{ => reg}/reg051/expected | 0 tests/idris2/reg/reg051/run | 2 + tests/idris2/{ => reg}/reg051/test.ipkg | 0 tests/idris2/{ => reg}/reg052/DPairQuote.idr | 0 tests/idris2/{ => reg}/reg052/expected | 0 tests/idris2/reg/reg052/run | 2 + tests/idris2/{ => reg}/reg052/test.ipkg | 0 tests/idris2/reg001/run | 3 - tests/idris2/reg002/run | 3 - tests/idris2/reg004/run | 3 - tests/idris2/reg005/run | 3 - tests/idris2/reg006/run | 3 - tests/idris2/reg007/run | 3 - tests/idris2/reg009/run | 3 - tests/idris2/reg010/run | 3 - tests/idris2/reg011/run | 3 - tests/idris2/reg012/run | 3 - tests/idris2/reg015/run | 3 - tests/idris2/reg016/run | 3 - tests/idris2/reg017/run | 3 - tests/idris2/reg018/run | 3 - tests/idris2/reg019/run | 3 - tests/idris2/reg021/run | 3 - tests/idris2/reg022/run | 3 - tests/idris2/reg023/run | 3 - tests/idris2/reg026/run | 3 - tests/idris2/reg027/run | 3 - tests/idris2/reg028/run | 3 - tests/idris2/reg029/run | 3 - tests/idris2/reg030/run | 3 - tests/idris2/reg031/run | 3 - tests/idris2/reg032/run | 3 - tests/idris2/reg033/run | 3 - tests/idris2/reg034/run | 3 - tests/idris2/reg035/run | 3 - tests/idris2/reg036/run | 3 - tests/idris2/reg037/run | 3 - tests/idris2/reg039/run | 3 - tests/idris2/reg040/run | 3 - tests/idris2/reg041/run | 3 - tests/idris2/reg044/run | 3 - tests/idris2/reg045/run | 3 - tests/idris2/reg051/run | 2 - tests/idris2/reg052/run | 2 - tests/idris2/{ => repl}/repl001/expected | 0 tests/idris2/{ => repl}/repl001/input | 0 tests/idris2/repl/repl001/run | 3 + tests/idris2/{ => repl}/repl002/expected | 0 tests/idris2/{ => repl}/repl002/input | 0 tests/idris2/repl/repl002/run | 3 + tests/idris2/{ => repl}/repl003/expected | 0 tests/idris2/{ => repl}/repl003/input | 0 tests/idris2/repl/repl003/run | 3 + tests/idris2/{ => repl}/repl004/expected | 0 tests/idris2/{ => repl}/repl004/input | 0 tests/idris2/repl/repl004/run | 3 + tests/idris2/{ => repl}/repl005/expected | 0 tests/idris2/{ => repl}/repl005/input | 0 tests/idris2/repl/repl005/run | 3 + tests/idris2/{ => repl}/repl006/expected | 0 tests/idris2/{ => repl}/repl006/input | 0 tests/idris2/repl/repl006/run | 3 + tests/idris2/repl001/run | 3 - tests/idris2/repl002/run | 3 - tests/idris2/repl003/run | 3 - tests/idris2/repl004/run | 3 - tests/idris2/repl005/run | 3 - tests/idris2/repl006/run | 3 - tests/idris2/rewrite001/run | 3 - .../{ => schemeeval}/schemeeval001/expected | 0 .../{ => schemeeval}/schemeeval001/input | 0 tests/idris2/schemeeval/schemeeval001/run | 3 + .../{ => schemeeval}/schemeeval002/expected | 0 .../{ => schemeeval}/schemeeval002/input | 0 tests/idris2/schemeeval/schemeeval002/run | 3 + .../{ => schemeeval}/schemeeval003/expected | 0 .../{ => schemeeval}/schemeeval003/input | 0 tests/idris2/schemeeval/schemeeval003/run | 3 + .../{ => schemeeval}/schemeeval004/expected | 0 .../{ => schemeeval}/schemeeval004/input | 0 .../{ => schemeeval}/schemeeval004/list.idr | 0 .../idris2/{ => schemeeval}/schemeeval004/run | 2 +- .../{ => schemeeval}/schemeeval005/Printf.idr | 0 .../{ => schemeeval}/schemeeval005/expected | 0 .../{ => schemeeval}/schemeeval005/input | 0 .../idris2/{ => schemeeval}/schemeeval005/run | 2 +- .../{ => schemeeval}/schemeeval006/expected | 0 .../{ => schemeeval}/schemeeval006/input | 0 tests/idris2/schemeeval/schemeeval006/run | 3 + tests/idris2/schemeeval001/run | 3 - tests/idris2/schemeeval002/run | 3 - tests/idris2/schemeeval003/run | 3 - tests/idris2/schemeeval006/run | 3 - .../termination001/AgdaIssue6059.idr | 0 .../{ => termination}/termination001/expected | 0 .../{ => termination}/termination001/run | 2 +- .../{ => total}/positivity001/Issue660.idr | 0 .../idris2/{ => total}/positivity001/expected | 0 tests/idris2/total/positivity001/run | 3 + .../{ => total}/positivity002/Issue660.idr | 0 .../idris2/{ => total}/positivity002/expected | 0 tests/idris2/total/positivity002/run | 3 + .../{ => total}/positivity003/Issue660.idr | 0 .../idris2/{ => total}/positivity003/expected | 0 tests/idris2/total/positivity003/run | 3 + .../{ => total}/positivity004/Issue1771-1.idr | 0 .../{ => total}/positivity004/Issue1771-2.idr | 0 .../{ => total}/positivity004/Issue1771-3.idr | 0 .../idris2/{ => total}/positivity004/expected | 0 tests/idris2/{ => total}/positivity004/run | 2 +- tests/idris2/{ => total}/total001/Total.idr | 0 tests/idris2/{ => total}/total001/expected | 0 tests/idris2/{ => total}/total001/input | 0 tests/idris2/{ => total}/total001/run | 2 +- tests/idris2/{ => total}/total002/Total.idr | 0 tests/idris2/{ => total}/total002/expected | 0 tests/idris2/{ => total}/total002/input | 0 tests/idris2/{ => total}/total002/run | 2 +- tests/idris2/{ => total}/total003/Total.idr | 0 tests/idris2/{ => total}/total003/expected | 0 tests/idris2/{ => total}/total003/input | 0 tests/idris2/{ => total}/total003/run | 2 +- tests/idris2/{ => total}/total004/Total.idr | 0 tests/idris2/{ => total}/total004/expected | 0 tests/idris2/{ => total}/total004/input | 0 tests/idris2/{ => total}/total004/run | 2 +- tests/idris2/{ => total}/total005/Total.idr | 0 tests/idris2/{ => total}/total005/expected | 0 tests/idris2/{ => total}/total005/input | 0 tests/idris2/total/total005/run | 3 + tests/idris2/{ => total}/total006/Total.idr | 0 tests/idris2/{ => total}/total006/expected | 0 tests/idris2/{ => total}/total006/input | 0 tests/idris2/total/total006/run | 3 + tests/idris2/{ => total}/total007/expected | 0 tests/idris2/{ => total}/total007/partial.idr | 0 tests/idris2/total/total007/run | 3 + tests/idris2/{ => total}/total008/expected | 0 tests/idris2/{ => total}/total008/partial.idr | 0 tests/idris2/total/total008/run | 3 + tests/idris2/{ => total}/total009/expected | 0 tests/idris2/total/total009/run | 3 + tests/idris2/{ => total}/total009/tree.idr | 0 .../{ => total}/total010/PartialWith.idr | 0 tests/idris2/{ => total}/total010/expected | 0 tests/idris2/total/total010/run | 3 + .../idris2/{ => total}/total011/Issue1460.idr | 0 .../idris2/{ => total}/total011/Issue1782.idr | 0 .../idris2/{ => total}/total011/Issue1828.idr | 0 .../{ => total}/total011/Issue1859-2.idr | 0 .../idris2/{ => total}/total011/Issue1859.idr | 0 tests/idris2/{ => total}/total011/expected | 0 tests/idris2/{ => total}/total011/run | 2 +- .../idris2/{ => total}/total012/Issue1828.idr | 0 .../{ => total}/total012/TotallyTotal.idr | 0 tests/idris2/{ => total}/total012/expected | 0 tests/idris2/{ => total}/total012/run | 2 +- .../idris2/{ => total}/total013/Issue1404.idr | 0 tests/idris2/{ => total}/total013/expected | 0 tests/idris2/total/total013/run | 3 + .../idris2/{ => total}/total014/FunCompTC.idr | 0 tests/idris2/{ => total}/total014/expected | 0 tests/idris2/total/total014/run | 3 + .../{ => total}/total015/CoveringData.idr | 0 tests/idris2/{ => total}/total015/expected | 0 tests/idris2/{ => total}/total015/run | 2 +- .../{ => total}/total016/AssertPositivity.idr | 0 .../total016/LazyPositivityCheck.idr | 0 tests/idris2/{ => total}/total016/expected | 0 tests/idris2/{ => total}/total016/run | 2 +- tests/idris2/{ => total}/total017/Paper.idr | 0 tests/idris2/{ => total}/total017/expected | 0 tests/idris2/total/total017/run | 3 + .../idris2/{ => total}/total018/Issue2448.idr | 0 tests/idris2/{ => total}/total018/expected | 0 tests/idris2/total/total018/run | 3 + tests/idris2/{ => total}/total019/Check.idr | 0 tests/idris2/{ => total}/total019/expected | 0 tests/idris2/total/total019/run | 3 + tests/idris2/{ => total}/total020/Check.idr | 0 tests/idris2/{ => total}/total020/expected | 0 tests/idris2/total/total020/run | 3 + tests/idris2/{ => total}/total020/test.ipkg | 0 tests/idris2/total005/run | 3 - tests/idris2/total006/run | 3 - tests/idris2/total007/run | 3 - tests/idris2/total008/run | 3 - tests/idris2/total009/run | 3 - tests/idris2/total010/run | 3 - tests/idris2/total013/run | 3 - tests/idris2/total014/run | 3 - tests/idris2/total017/run | 3 - tests/idris2/total018/run | 3 - tests/idris2/total019/run | 3 - tests/idris2/total020/run | 3 - tests/idris2/unification001/run | 3 - .../idris2/{ => warning}/warning001/Holes.idr | 0 .../{ => warning}/warning001/Issue1401.idr | 0 .../{ => warning}/warning001/Issue539.idr | 0 .../{ => warning}/warning001/Issue621.idr | 0 .../{ => warning}/warning001/PR1407.idr | 0 .../idris2/{ => warning}/warning001/expected | 0 tests/idris2/{ => warning}/warning001/run | 2 +- tests/idris2/{ => warning}/warning002/Foo.idr | 0 .../idris2/{ => warning}/warning002/Main.idr | 0 .../{ => warning}/warning002/deprecated.ipkg | 0 .../idris2/{ => warning}/warning002/expected | 0 tests/idris2/{ => warning}/warning002/run | 2 +- .../idris2/{ => warning}/warning003/Main.idr | 0 .../{ => warning}/warning003/deprecated.ipkg | 0 .../idris2/{ => warning}/warning003/expected | 0 tests/idris2/warning/warning003/run | 3 + .../idris2/{ => warning}/warning004/Lib1.idr | 0 .../idris2/{ => warning}/warning004/Lib2.idr | 0 .../idris2/{ => warning}/warning004/Main1.idr | 0 .../idris2/{ => warning}/warning004/expected | 0 tests/idris2/{ => warning}/warning004/run | 2 +- tests/idris2/warning003/run | 3 - tests/idris2/{ => with}/with001/Temp.idr | 0 tests/idris2/{ => with}/with001/expected | 0 tests/idris2/with/with001/run | 3 + tests/idris2/{ => with}/with002/Temp.idr | 0 tests/idris2/{ => with}/with002/expected | 0 tests/idris2/with/with002/run | 3 + .../idris2/{ => with}/with004/Issue637-2.idr | 0 .../idris2/{ => with}/with004/Issue637-3.idr | 0 tests/idris2/{ => with}/with004/Issue637.idr | 0 tests/idris2/{ => with}/with004/expected | 0 tests/idris2/{ => with}/with004/input | 0 tests/idris2/{ => with}/with004/run | 2 +- tests/idris2/{ => with}/with005/Issue893.idr | 0 tests/idris2/{ => with}/with005/WithProof.idr | 0 tests/idris2/{ => with}/with005/expected | 0 tests/idris2/{ => with}/with005/run | 2 +- .../idris2/{ => with}/with006/SparseWith.idr | 0 tests/idris2/{ => with}/with006/expected | 0 tests/idris2/with/with006/run | 3 + tests/idris2/{ => with}/with007/With0.idr | 0 tests/idris2/{ => with}/with007/expected | 0 tests/idris2/with/with007/run | 3 + .../idris2/{ => with}/with008/WithClause.idr | 0 tests/idris2/{ => with}/with008/expected | 0 tests/idris2/with/with008/run | 3 + .../idris2/{ => with}/with009/WithClause.idr | 0 tests/idris2/{ => with}/with009/expected | 0 tests/idris2/with/with009/run | 3 + .../idris2/{ => with}/with010/NestedWith.idr | 0 tests/idris2/{ => with}/with010/expected | 0 tests/idris2/with/with010/run | 3 + .../{ => with}/with011/WithImplicits.idr | 0 tests/idris2/{ => with}/with011/expected | 0 tests/idris2/{ => with}/with011/run | 2 +- tests/idris2/with001/run | 3 - tests/idris2/with002/run | 3 - tests/idris2/with006/run | 3 - tests/idris2/with007/run | 3 - tests/idris2/with008/run | 3 - tests/idris2/with009/run | 3 - tests/idris2/with010/run | 3 - tests/node/node014/Echo.idr | 57 --- tests/node/node014/expected | 2 - tests/node/node014/run | 3 - tests/racket/barrier001/Main.idr | 2 +- 2218 files changed, 1049 insertions(+), 1461 deletions(-) rename tests/idris2/{api001 => api}/README (100%) rename tests/idris2/{ => api}/api001/Hello.idr (100%) rename tests/idris2/{ => api}/api001/LazyCodegen.idr (100%) rename tests/idris2/{ => api}/api001/expected (100%) rename tests/idris2/{ => api}/api001/run (88%) rename tests/idris2/{ => basic}/basic001/Vect.idr (100%) rename tests/idris2/{ => basic}/basic001/expected (100%) rename tests/idris2/{ => basic}/basic001/input (100%) rename tests/idris2/{ => basic}/basic001/run (61%) rename tests/idris2/{ => basic}/basic002/Do.idr (100%) rename tests/idris2/{ => basic}/basic002/expected (100%) rename tests/idris2/{ => basic}/basic002/input (100%) rename tests/idris2/{ => basic}/basic002/run (60%) rename tests/idris2/{ => basic}/basic003/Ambig1.idr (100%) rename tests/idris2/{ => basic}/basic003/Ambig2.idr (100%) rename tests/idris2/{ => basic}/basic003/expected (100%) rename tests/idris2/{ => basic}/basic003/run (78%) rename tests/idris2/{ => basic}/basic004/Stuff.idr (100%) rename tests/idris2/{ => basic}/basic004/Wheres.idr (100%) rename tests/idris2/{ => basic}/basic004/expected (100%) rename tests/idris2/{ => basic}/basic004/input (100%) rename tests/idris2/{ => basic}/basic004/run (62%) rename tests/idris2/{ => basic}/basic005/NoInfer.idr (100%) rename tests/idris2/{ => basic}/basic005/expected (100%) rename tests/idris2/{ => basic}/basic005/run (65%) rename tests/idris2/{ => basic}/basic006/PMLet.idr (100%) rename tests/idris2/{ => basic}/basic006/Stuff.idr (100%) rename tests/idris2/{ => basic}/basic006/expected (100%) rename tests/idris2/{ => basic}/basic006/input (100%) rename tests/idris2/{ => basic}/basic006/run (61%) rename tests/idris2/{ => basic}/basic007/DoLocal.idr (100%) rename tests/idris2/{ => basic}/basic007/Stuff.idr (100%) rename tests/idris2/{ => basic}/basic007/expected (100%) rename tests/idris2/{ => basic}/basic007/input (100%) rename tests/idris2/{ => basic}/basic007/run (63%) rename tests/idris2/{ => basic}/basic008/If.idr (100%) rename tests/idris2/{ => basic}/basic008/Stuff.idr (100%) rename tests/idris2/{ => basic}/basic008/expected (100%) rename tests/idris2/{ => basic}/basic008/input (100%) rename tests/idris2/{ => basic}/basic008/run (60%) rename tests/idris2/{ => basic}/basic009/LetCase.idr (100%) rename tests/idris2/{ => basic}/basic009/Stuff.idr (100%) rename tests/idris2/{ => basic}/basic009/expected (100%) rename tests/idris2/{ => basic}/basic009/input (100%) rename tests/idris2/{ => basic}/basic009/run (63%) rename tests/idris2/{ => basic}/basic010/Comp.idr (100%) rename tests/idris2/{ => basic}/basic010/expected (100%) rename tests/idris2/{ => basic}/basic010/input (100%) rename tests/idris2/{ => basic}/basic010/run (61%) rename tests/idris2/{ => basic}/basic011/Dots1.idr (100%) rename tests/idris2/{ => basic}/basic011/Dots2.idr (100%) rename tests/idris2/{ => basic}/basic011/Dots3.idr (100%) rename tests/idris2/{ => basic}/basic011/expected (100%) rename tests/idris2/{ => basic}/basic011/run (67%) rename tests/idris2/{ => basic}/basic012/VIndex.idr (100%) rename tests/idris2/{ => basic}/basic012/expected (100%) create mode 100755 tests/idris2/basic/basic012/run rename tests/idris2/{ => basic}/basic013/Implicits.idr (100%) rename tests/idris2/{ => basic}/basic013/expected (100%) create mode 100755 tests/idris2/basic/basic013/run rename tests/idris2/{ => basic}/basic014/Rewrite.idr (100%) rename tests/idris2/{ => basic}/basic014/expected (100%) create mode 100755 tests/idris2/basic/basic014/run rename tests/idris2/{ => basic}/basic015/George.idr (100%) rename tests/idris2/{ => basic}/basic015/expected (100%) create mode 100755 tests/idris2/basic/basic015/run rename tests/idris2/{ => basic}/basic016/Eta.idr (100%) rename tests/idris2/{ => basic}/basic016/Eta2.idr (100%) rename tests/idris2/{ => basic}/basic016/expected (100%) rename tests/idris2/{ => basic}/basic016/run (55%) rename tests/idris2/{ => basic}/basic017/CaseInf.idr (100%) rename tests/idris2/{ => basic}/basic017/expected (100%) rename tests/idris2/{ => basic}/basic017/input (100%) rename tests/idris2/{ => basic}/basic017/run (53%) rename tests/idris2/{ => basic}/basic018/Fin.idr (100%) rename tests/idris2/{ => basic}/basic018/expected (100%) create mode 100755 tests/idris2/basic/basic018/run rename tests/idris2/{ => basic}/basic019/CaseBlock.idr (100%) rename tests/idris2/{ => basic}/basic019/expected (100%) rename tests/idris2/{ => basic}/basic019/input (100%) rename tests/idris2/{ => basic}/basic019/run (55%) rename tests/idris2/{ => basic}/basic020/Mut.idr (100%) rename tests/idris2/{ => basic}/basic020/expected (100%) rename tests/idris2/{ => basic}/basic020/input (100%) rename tests/idris2/{ => basic}/basic020/run (50%) rename tests/idris2/{ => basic}/basic021/CaseDep.idr (100%) rename tests/idris2/{ => basic}/basic021/expected (100%) rename tests/idris2/{ => basic}/basic021/input (100%) rename tests/idris2/{ => basic}/basic021/run (53%) rename tests/idris2/{ => basic}/basic022/Erase.idr (100%) rename tests/idris2/{ => basic}/basic022/expected (100%) rename tests/idris2/{ => basic}/basic022/input (100%) rename tests/idris2/{ => basic}/basic022/run (52%) rename tests/idris2/{ => basic}/basic023/Params.idr (100%) rename tests/idris2/{ => basic}/basic023/expected (100%) rename tests/idris2/{ => basic}/basic023/input (100%) rename tests/idris2/{ => basic}/basic023/run (52%) rename tests/idris2/{ => basic}/basic024/PatLam.idr (100%) rename tests/idris2/{ => basic}/basic024/expected (100%) rename tests/idris2/{ => basic}/basic024/input (100%) rename tests/idris2/{ => basic}/basic024/run (52%) rename tests/idris2/{ => basic}/basic025/expected (100%) rename tests/idris2/{ => basic}/basic025/input (100%) create mode 100755 tests/idris2/basic/basic025/run rename tests/idris2/{ => basic}/basic026/Erl.idr (100%) rename tests/idris2/{ => basic}/basic026/expected (100%) create mode 100755 tests/idris2/basic/basic026/run rename tests/idris2/{ => basic}/basic027/Temp.idr (100%) rename tests/idris2/{ => basic}/basic027/expected (100%) create mode 100755 tests/idris2/basic/basic027/run rename tests/idris2/{ => basic}/basic028/Do.idr (100%) rename tests/idris2/{ => basic}/basic028/expected (100%) rename tests/idris2/{ => basic}/basic028/input (100%) rename tests/idris2/{ => basic}/basic028/run (66%) rename tests/idris2/{ => basic}/basic029/Params.idr (100%) rename tests/idris2/{ => basic}/basic029/expected (100%) rename tests/idris2/{ => basic}/basic029/input (100%) rename tests/idris2/{ => basic}/basic029/run (52%) rename tests/idris2/{ => basic}/basic030/arity.idr (100%) rename tests/idris2/{ => basic}/basic030/expected (100%) create mode 100644 tests/idris2/basic/basic030/run rename tests/idris2/{ => basic}/basic031/erased.idr (100%) rename tests/idris2/{ => basic}/basic031/expected (100%) create mode 100644 tests/idris2/basic/basic031/run rename tests/idris2/{ => basic}/basic032/Idiom.idr (100%) rename tests/idris2/{ => basic}/basic032/Idiom2.idr (100%) rename tests/idris2/{ => basic}/basic032/expected (100%) rename tests/idris2/{ => basic}/basic032/input (100%) rename tests/idris2/{ => basic}/basic032/run (64%) rename tests/idris2/{ => basic}/basic033/expected (100%) create mode 100644 tests/idris2/basic/basic033/run rename tests/idris2/{ => basic}/basic033/unboundimps.idr (100%) rename tests/idris2/{ => basic}/basic034/expected (100%) rename tests/idris2/{ => basic}/basic034/lets.idr (100%) create mode 100644 tests/idris2/basic/basic034/run rename tests/idris2/{ => basic}/basic035/expected (100%) rename tests/idris2/{ => basic}/basic035/input (100%) rename tests/idris2/{ => basic}/basic035/run (52%) rename tests/idris2/{ => basic}/basic035/using.idr (100%) rename tests/idris2/{ => basic}/basic036/defimp.idr (100%) rename tests/idris2/{ => basic}/basic036/expected (100%) rename tests/idris2/{ => basic}/basic036/input (100%) rename tests/idris2/{ => basic}/basic036/run (52%) rename tests/idris2/{ => basic}/basic037/Comments.idr (100%) rename tests/idris2/{ => basic}/basic037/Issue279.idr (100%) rename tests/idris2/{ => basic}/basic037/expected (100%) rename tests/idris2/{ => basic}/basic037/run (76%) rename tests/idris2/{ => basic}/basic038/Resugar.idr (100%) rename tests/idris2/{ => basic}/basic038/expected (100%) rename tests/idris2/{ => basic}/basic038/input (100%) rename tests/idris2/{ => basic}/basic038/run (53%) rename tests/idris2/{ => basic}/basic039/Main.idr (100%) rename tests/idris2/{ => basic}/basic039/expected (100%) rename tests/idris2/{ => basic}/basic039/input (100%) rename tests/idris2/{ => basic}/basic039/run (51%) rename tests/idris2/{ => basic}/basic040/Default.idr (100%) rename tests/idris2/{ => basic}/basic040/expected (100%) rename tests/idris2/{ => basic}/basic040/run (57%) rename tests/idris2/{ => basic}/basic041/QDo.idr (100%) rename tests/idris2/{ => basic}/basic041/expected (100%) create mode 100755 tests/idris2/basic/basic041/run rename tests/idris2/{ => basic}/basic042/LiteralsInteger.idr (100%) rename tests/idris2/{ => basic}/basic042/LiteralsString.idr (100%) rename tests/idris2/{ => basic}/basic042/expected (100%) rename tests/idris2/{ => basic}/basic042/input (100%) rename tests/idris2/{ => basic}/basic042/input2 (100%) rename tests/idris2/{ => basic}/basic042/run (74%) rename tests/idris2/{ => basic}/basic043/BitCasts.idr (100%) rename tests/idris2/{ => basic}/basic043/expected (100%) rename tests/idris2/{ => basic}/basic043/input (100%) rename tests/idris2/{ => basic}/basic043/run (54%) rename tests/idris2/{ => basic}/basic044/Term.idr (100%) rename tests/idris2/{ => basic}/basic044/Vec.idr (100%) rename tests/idris2/{ => basic}/basic044/expected (100%) rename tests/idris2/{ => basic}/basic044/run (94%) rename tests/idris2/{ => basic}/basic045/Main.idr (100%) rename tests/idris2/{ => basic}/basic045/expected (100%) create mode 100644 tests/idris2/basic/basic045/run rename tests/idris2/{ => basic}/basic046/TupleSections.idr (100%) rename tests/idris2/{ => basic}/basic046/expected (100%) rename tests/idris2/{ => basic}/basic046/input (100%) rename tests/idris2/{ => basic}/basic046/run (58%) rename tests/idris2/{ => basic}/basic047/InterleavingLets.idr (100%) rename tests/idris2/{ => basic}/basic047/expected (100%) rename tests/idris2/{ => basic}/basic047/input (100%) rename tests/idris2/{ => basic}/basic047/run (60%) rename tests/idris2/{ => basic}/basic049/Fld.idr (100%) rename tests/idris2/{ => basic}/basic049/expected (100%) rename tests/idris2/{ => basic}/basic049/input (100%) rename tests/idris2/{ => basic}/basic049/run (50%) rename tests/idris2/{ => basic}/basic050/Ilc.idr (100%) rename tests/idris2/{ => basic}/basic050/expected (100%) rename tests/idris2/{ => basic}/basic050/input (100%) rename tests/idris2/{ => basic}/basic050/run (50%) rename tests/idris2/{ => basic}/basic051/Issue833.idr (100%) rename tests/idris2/{ => basic}/basic051/expected (100%) create mode 100755 tests/idris2/basic/basic051/run rename tests/idris2/{ => basic}/basic052/DoubleClBrace.idr (100%) rename tests/idris2/{ => basic}/basic052/expected (100%) rename tests/idris2/{ => basic}/basic052/input (100%) rename tests/idris2/{ => basic}/basic052/run (58%) rename tests/idris2/{ => basic}/basic053/UnderscoredIntegerLiterals.idr (100%) rename tests/idris2/{ => basic}/basic053/expected (100%) rename tests/idris2/{ => basic}/basic053/input (100%) rename tests/idris2/{ => basic}/basic053/run (66%) rename tests/idris2/{ => basic}/basic054/Issue1023.idr (100%) rename tests/idris2/{ => basic}/basic054/expected (100%) create mode 100755 tests/idris2/basic/basic054/run rename tests/idris2/{ => basic}/basic055/BitOps.idr (100%) rename tests/idris2/{ => basic}/basic055/expected (100%) create mode 100644 tests/idris2/basic/basic055/run rename tests/idris2/{ => basic}/basic056/DoubleLit.idr (100%) rename tests/idris2/{ => basic}/basic056/expected (100%) create mode 100644 tests/idris2/basic/basic056/run rename tests/idris2/{ => basic}/basic057/LetIn.idr (100%) rename tests/idris2/{ => basic}/basic057/expected (100%) create mode 100644 tests/idris2/basic/basic057/run rename tests/idris2/{ => basic}/basic058/DataTypeOp.idr (100%) rename tests/idris2/{ => basic}/basic058/DataTypeProj.idr (100%) rename tests/idris2/{ => basic}/basic058/expected (100%) rename tests/idris2/{ => basic}/basic058/run (65%) rename tests/idris2/{ => basic}/basic059/MultiClaim.idr (100%) rename tests/idris2/{ => basic}/basic059/expected (100%) create mode 100644 tests/idris2/basic/basic059/run rename tests/idris2/{ => basic}/basic060/Snoc.idr (100%) rename tests/idris2/{ => basic}/basic060/expected (100%) rename tests/idris2/{ => basic}/basic060/input (100%) rename tests/idris2/{ => basic}/basic060/run (51%) rename tests/idris2/{ => basic}/basic061/IgnoreDo.idr (100%) rename tests/idris2/{ => basic}/basic061/expected (100%) rename tests/idris2/{ => basic}/basic061/input (100%) rename tests/idris2/{ => basic}/basic061/run (54%) rename tests/idris2/{ => basic}/basic062/Issue1943.idr (100%) rename tests/idris2/{ => basic}/basic062/expected (100%) create mode 100755 tests/idris2/basic/basic062/run rename tests/idris2/{ => basic}/basic063/NoDeclaration.idr (100%) rename tests/idris2/{ => basic}/basic063/expected (100%) rename tests/idris2/{ => basic}/basic063/input (100%) rename tests/idris2/{ => basic}/basic063/run (58%) rename tests/idris2/{ => basic}/basic064/Issue2072.idr (100%) rename tests/idris2/{ => basic}/basic064/expected (100%) rename tests/idris2/{ => basic}/basic064/input (100%) rename tests/idris2/{ => basic}/basic064/run (55%) rename tests/idris2/{ => basic}/basic065/Issue215.idr (100%) rename tests/idris2/{ => basic}/basic065/expected (100%) create mode 100755 tests/idris2/basic/basic065/run rename tests/idris2/{ => basic}/basic066/comment.idr (100%) rename tests/idris2/{ => basic}/basic066/expected (100%) create mode 100755 tests/idris2/basic/basic066/run rename tests/idris2/{ => basic}/basic067/expected (100%) rename tests/idris2/{ => basic}/basic067/input (100%) rename tests/idris2/{ => basic}/basic067/run (86%) rename tests/idris2/{ => basic}/basic067/unclosed1.idr (100%) rename tests/idris2/{ => basic}/basic067/unclosed2.idr (100%) rename tests/idris2/{ => basic}/basic067/unclosed3.idr (100%) rename tests/idris2/{ => basic}/basic068/Issue2138.idr (100%) rename tests/idris2/{ => basic}/basic068/expected (100%) create mode 100755 tests/idris2/basic/basic068/run rename tests/idris2/{ => basic}/basic069/DebugInfo.idr (100%) rename tests/idris2/{ => basic}/basic069/expected (100%) create mode 100755 tests/idris2/basic/basic069/run rename tests/idris2/{ => basic}/basic070/Issue2592.idr (100%) rename tests/idris2/{ => basic}/basic070/Issue2593.idr (100%) rename tests/idris2/{ => basic}/basic070/Issue2782.idr (100%) rename tests/idris2/{ => basic}/basic070/Issue3016.idr (100%) rename tests/idris2/{ => basic}/basic070/expected (100%) rename tests/idris2/{ => basic}/basic070/run (77%) rename tests/idris2/{ => basic}/basic071/A.idr (100%) rename tests/idris2/{ => basic}/basic071/B.idr (100%) rename tests/idris2/{ => basic}/basic071/expected (100%) rename tests/idris2/{ => basic}/basic071/run (100%) rename tests/idris2/{ => basic}/case001/InlineCase.idr (100%) rename tests/idris2/{ => basic}/case001/expected (100%) rename tests/idris2/{ => basic}/case001/run (100%) rename tests/idris2/{ => basic}/dotted001/Issue2726.idr (100%) rename tests/idris2/{ => basic}/dotted001/expected (100%) rename tests/idris2/{ => basic}/dotted001/run (51%) rename tests/idris2/{ => basic}/idiom001/Main.idr (100%) rename tests/idris2/{ => basic}/idiom001/expected (100%) create mode 100644 tests/idris2/basic/idiom001/run rename tests/idris2/{ => basic}/interpolation001/IfThenElse.idr (100%) rename tests/idris2/{ => basic}/interpolation001/expected (100%) rename tests/idris2/{ => basic}/interpolation001/run (64%) rename tests/idris2/{ => basic}/interpolation002/StringLiteral.idr (100%) rename tests/idris2/{ => basic}/interpolation002/expected (100%) rename tests/idris2/{ => basic}/interpolation002/run (61%) rename tests/idris2/{ => basic}/interpolation003/Test.idr (100%) rename tests/idris2/{ => basic}/interpolation003/expected (100%) rename tests/idris2/{ => basic}/interpolation003/input (100%) rename tests/idris2/{builtin001 => basic/interpolation003}/run (51%) rename tests/idris2/{ => basic}/interpolation004/StringLiteral.idr (100%) rename tests/idris2/{ => basic}/interpolation004/expected (100%) rename tests/idris2/{ => basic}/interpolation004/run (51%) rename tests/idris2/{ => basic}/literals001/Test.idr (100%) rename tests/idris2/{ => basic}/literals001/expected (100%) rename tests/idris2/{ => basic}/literals001/run (100%) rename tests/idris2/{ => basic}/rewrite001/Issue2573.idr (100%) rename tests/idris2/{ => basic}/rewrite001/expected (100%) create mode 100755 tests/idris2/basic/rewrite001/run delete mode 100755 tests/idris2/basic012/run delete mode 100755 tests/idris2/basic013/run delete mode 100755 tests/idris2/basic014/run delete mode 100755 tests/idris2/basic015/run delete mode 100755 tests/idris2/basic018/run delete mode 100755 tests/idris2/basic025/run delete mode 100755 tests/idris2/basic026/run delete mode 100755 tests/idris2/basic027/run delete mode 100644 tests/idris2/basic030/run delete mode 100644 tests/idris2/basic031/run delete mode 100644 tests/idris2/basic033/run delete mode 100644 tests/idris2/basic034/run delete mode 100755 tests/idris2/basic041/run delete mode 100644 tests/idris2/basic045/run delete mode 100755 tests/idris2/basic051/run delete mode 100755 tests/idris2/basic054/run delete mode 100644 tests/idris2/basic055/run delete mode 100644 tests/idris2/basic056/run delete mode 100644 tests/idris2/basic057/run delete mode 100644 tests/idris2/basic059/run delete mode 100755 tests/idris2/basic062/run delete mode 100755 tests/idris2/basic065/run delete mode 100755 tests/idris2/basic066/run delete mode 100755 tests/idris2/basic068/run delete mode 100755 tests/idris2/basic069/run rename tests/idris2/{ => builtin}/builtin001/Test.idr (100%) rename tests/idris2/{ => builtin}/builtin001/expected (100%) rename tests/idris2/{ => builtin}/builtin001/input (100%) rename tests/idris2/{builtin002 => builtin/builtin001}/run (51%) rename tests/idris2/{ => builtin}/builtin002/Test.idr (100%) rename tests/idris2/{ => builtin}/builtin002/expected (100%) rename tests/idris2/{ => builtin}/builtin002/input (100%) rename tests/idris2/{builtin003 => builtin/builtin002}/run (51%) rename tests/idris2/{ => builtin}/builtin003/Test.idr (100%) rename tests/idris2/{ => builtin}/builtin003/expected (100%) rename tests/idris2/{ => builtin}/builtin003/input (100%) rename tests/idris2/{builtin004 => builtin/builtin003}/run (51%) rename tests/idris2/{ => builtin}/builtin004/Test.idr (100%) rename tests/idris2/{ => builtin}/builtin004/expected (100%) rename tests/idris2/{ => builtin}/builtin004/input (100%) create mode 100755 tests/idris2/builtin/builtin004/run rename tests/idris2/{ => builtin}/builtin005/Test.idr (100%) rename tests/idris2/{ => builtin}/builtin005/expected (100%) rename tests/idris2/{ => builtin}/builtin005/input (100%) create mode 100755 tests/idris2/builtin/builtin005/run rename tests/idris2/{ => builtin}/builtin006/Test.idr (100%) rename tests/idris2/{ => builtin}/builtin006/expected (100%) rename tests/idris2/{ => builtin}/builtin006/input (100%) create mode 100755 tests/idris2/builtin/builtin006/run rename tests/idris2/{ => builtin}/builtin007/Test.idr (100%) rename tests/idris2/{ => builtin}/builtin007/expected (100%) rename tests/idris2/{ => builtin}/builtin007/input (100%) create mode 100755 tests/idris2/builtin/builtin007/run rename tests/idris2/{ => builtin}/builtin008/Test.idr (100%) rename tests/idris2/{ => builtin}/builtin008/expected (100%) rename tests/idris2/{ => builtin}/builtin008/input (100%) create mode 100755 tests/idris2/builtin/builtin008/run rename tests/idris2/{ => builtin}/builtin009/Test.idr (100%) rename tests/idris2/{ => builtin}/builtin009/expected (100%) create mode 100755 tests/idris2/builtin/builtin009/run rename tests/idris2/{ => builtin}/builtin010/Test.idr (100%) rename tests/idris2/{ => builtin}/builtin010/expected (100%) rename tests/idris2/{ => builtin}/builtin010/input (100%) create mode 100755 tests/idris2/builtin/builtin010/run rename tests/idris2/{ => builtin}/builtin011/Test.idr (100%) rename tests/idris2/{ => builtin}/builtin011/expected (100%) create mode 100755 tests/idris2/builtin/builtin011/run rename tests/idris2/{ => builtin}/builtin012/Issue1799.idr (100%) rename tests/idris2/{ => builtin}/builtin012/expected (100%) create mode 100755 tests/idris2/builtin/builtin012/run delete mode 100755 tests/idris2/builtin005/run delete mode 100755 tests/idris2/builtin006/run delete mode 100755 tests/idris2/builtin007/run delete mode 100755 tests/idris2/builtin008/run delete mode 100755 tests/idris2/builtin009/run delete mode 100755 tests/idris2/builtin010/run delete mode 100755 tests/idris2/builtin011/run delete mode 100755 tests/idris2/builtin012/run rename tests/idris2/{ => casetree}/casetree001/IsS.idr (100%) rename tests/idris2/{ => casetree}/casetree001/Issue762.idr (100%) rename tests/idris2/{ => casetree}/casetree001/expected (100%) rename tests/idris2/{ => casetree}/casetree001/run (58%) rename tests/idris2/{ => casetree}/casetree002/DefaultCases.idr (100%) rename tests/idris2/{ => casetree}/casetree002/Issue1079.idr (100%) rename tests/idris2/{ => casetree}/casetree002/expected (100%) rename tests/idris2/{ => casetree}/casetree002/run (64%) rename tests/idris2/{ => casetree}/casetree003/ForcedPats.idr (100%) rename tests/idris2/{ => casetree}/casetree003/expected (100%) rename tests/idris2/{ => casetree}/casetree003/input (100%) rename tests/idris2/{ => casetree}/casetree003/run (56%) rename tests/idris2/{ => casetree}/casetree004/LocalArgs.idr (100%) rename tests/idris2/{ => casetree}/casetree004/PiMatch.idr (100%) rename tests/idris2/{ => casetree}/casetree004/expected (100%) rename tests/idris2/{ => casetree}/casetree004/run (61%) rename tests/idris2/{ => coverage}/coverage001/Vect.idr (100%) rename tests/idris2/{ => coverage}/coverage001/Vect2.idr (100%) rename tests/idris2/{ => coverage}/coverage001/Vect3.idr (100%) rename tests/idris2/{ => coverage}/coverage001/expected (100%) rename tests/idris2/{ => coverage}/coverage001/input (100%) rename tests/idris2/{ => coverage}/coverage001/run (73%) rename tests/idris2/{ => coverage}/coverage002/Vect.idr (100%) rename tests/idris2/{ => coverage}/coverage002/expected (100%) rename tests/idris2/{ => coverage}/coverage002/input (100%) rename tests/idris2/{ => coverage}/coverage002/run (51%) rename tests/idris2/{ => coverage}/coverage003/Cover.idr (100%) rename tests/idris2/{ => coverage}/coverage003/expected (100%) rename tests/idris2/{ => coverage}/coverage003/input (100%) rename tests/idris2/{ => coverage}/coverage003/run (52%) rename tests/idris2/{ => coverage}/coverage004/Cover.idr (100%) rename tests/idris2/{ => coverage}/coverage004/expected (100%) rename tests/idris2/{ => coverage}/coverage004/input (100%) rename tests/idris2/{ => coverage}/coverage004/run (52%) rename tests/idris2/{ => coverage}/coverage005/Cover.idr (100%) rename tests/idris2/{ => coverage}/coverage005/expected (100%) rename tests/idris2/{ => coverage}/coverage005/input (100%) rename tests/idris2/{ => coverage}/coverage005/run (52%) rename tests/idris2/{ => coverage}/coverage006/expected (100%) rename tests/idris2/{ => coverage}/coverage006/foobar.idr (100%) rename tests/idris2/{ => coverage}/coverage006/input (100%) rename tests/idris2/{ => coverage}/coverage006/run (52%) rename tests/idris2/{ => coverage}/coverage007/eq.idr (100%) rename tests/idris2/{ => coverage}/coverage007/expected (100%) create mode 100755 tests/idris2/coverage/coverage007/run rename tests/idris2/{ => coverage}/coverage008/expected (100%) rename tests/idris2/{ => coverage}/coverage008/input (100%) rename tests/idris2/{ => coverage}/coverage008/run (51%) rename tests/idris2/{ => coverage}/coverage008/wcov.idr (100%) rename tests/idris2/{ => coverage}/coverage009/expected (100%) create mode 100755 tests/idris2/coverage/coverage009/run rename tests/idris2/{ => coverage}/coverage009/unreachable.idr (100%) rename tests/idris2/{ => coverage}/coverage010/casetot.idr (100%) rename tests/idris2/{ => coverage}/coverage010/expected (100%) create mode 100755 tests/idris2/coverage/coverage010/run rename tests/idris2/{ => coverage}/coverage011/Sing.idr (100%) rename tests/idris2/{ => coverage}/coverage011/expected (100%) create mode 100755 tests/idris2/coverage/coverage011/run rename tests/idris2/{ => coverage}/coverage012/Issue484.idr (100%) rename tests/idris2/{ => coverage}/coverage012/Issue899.idr (100%) rename tests/idris2/{ => coverage}/coverage012/expected (100%) rename tests/idris2/{ => coverage}/coverage012/run (61%) rename tests/idris2/{ => coverage}/coverage013/Issue1022-Refl.idr (100%) rename tests/idris2/{ => coverage}/coverage013/Issue1022.idr (100%) rename tests/idris2/{ => coverage}/coverage013/expected (100%) rename tests/idris2/{ => coverage}/coverage013/run (65%) rename tests/idris2/{ => coverage}/coverage014/Issue794.idr (100%) rename tests/idris2/{ => coverage}/coverage014/expected (100%) create mode 100755 tests/idris2/coverage/coverage014/run rename tests/idris2/{ => coverage}/coverage015/Issue1169.idr (100%) rename tests/idris2/{ => coverage}/coverage015/Issue1366.idr (100%) rename tests/idris2/{ => coverage}/coverage015/expected (100%) rename tests/idris2/{ => coverage}/coverage015/run (63%) rename tests/idris2/{ => coverage}/coverage016/Issue633-2.idr (100%) rename tests/idris2/{ => coverage}/coverage016/Issue633.idr (100%) rename tests/idris2/{ => coverage}/coverage016/expected (100%) rename tests/idris2/{ => coverage}/coverage016/run (63%) rename tests/idris2/{ => coverage}/coverage017/Issue1421.idr (100%) rename tests/idris2/{ => coverage}/coverage017/expected (100%) create mode 100755 tests/idris2/coverage/coverage017/run rename tests/idris2/{ => coverage}/coverage018/Issue1831_1.idr (100%) rename tests/idris2/{ => coverage}/coverage018/Issue1831_2.idr (100%) rename tests/idris2/{ => coverage}/coverage018/expected (100%) rename tests/idris2/{ => coverage}/coverage018/input (100%) rename tests/idris2/{ => coverage}/coverage018/run (72%) rename tests/idris2/{ => coverage}/coverage019/Issue1632.idr (100%) rename tests/idris2/{ => coverage}/coverage019/expected (100%) create mode 100755 tests/idris2/coverage/coverage019/run delete mode 100755 tests/idris2/coverage007/run delete mode 100755 tests/idris2/coverage009/run delete mode 100755 tests/idris2/coverage010/run delete mode 100755 tests/idris2/coverage011/run delete mode 100755 tests/idris2/coverage014/run delete mode 100755 tests/idris2/coverage017/run delete mode 100755 tests/idris2/coverage019/run rename tests/idris2/{ => data}/data001/Test.idr (100%) rename tests/idris2/{ => data}/data001/TestImpl.idr (100%) rename tests/idris2/{ => data}/data001/expected (100%) create mode 100755 tests/idris2/data/data001/run rename tests/idris2/{ => data}/data002/Test.idr (100%) rename tests/idris2/{ => data}/data002/expected (100%) rename tests/idris2/{ => data}/data002/input (100%) create mode 100755 tests/idris2/data/data002/run rename tests/idris2/{ => data}/record001/Record.idr (100%) rename tests/idris2/{ => data}/record001/expected (100%) rename tests/idris2/{ => data}/record001/input (100%) rename tests/idris2/{ => data}/record001/run (52%) rename tests/idris2/{ => data}/record002/Record.idr (100%) rename tests/idris2/{ => data}/record002/expected (100%) rename tests/idris2/{ => data}/record002/input (100%) rename tests/idris2/{ => data}/record002/run (52%) rename tests/idris2/{ => data}/record003/Record.idr (100%) rename tests/idris2/{ => data}/record003/expected (100%) create mode 100755 tests/idris2/data/record003/run rename tests/idris2/{ => data}/record004/Main.idr (100%) rename tests/idris2/{ => data}/record004/expected (100%) rename tests/idris2/{ => data}/record004/input (100%) rename tests/idris2/{ => data}/record004/run (51%) rename tests/idris2/{ => data}/record005/Fld.idr (100%) rename tests/idris2/{ => data}/record005/expected (100%) rename tests/idris2/{ => data}/record005/input (100%) rename tests/idris2/{ => data}/record005/run (50%) rename tests/idris2/{ => data}/record006/Fld.idr (100%) rename tests/idris2/{ => data}/record006/expected (100%) rename tests/idris2/{ => data}/record006/input (100%) rename tests/idris2/{ => data}/record006/run (50%) rename tests/idris2/{ => data}/record007/Bond.idr (100%) rename tests/idris2/{ => data}/record007/expected (100%) rename tests/idris2/{error008 => data/record007}/input (100%) rename tests/idris2/{ => data}/record007/run (51%) rename tests/idris2/{ => data}/record008/Postfix.idr (100%) rename tests/idris2/{ => data}/record008/expected (100%) rename tests/idris2/{ => data}/record008/input (100%) rename tests/idris2/{ => data}/record008/run (53%) rename tests/idris2/{ => data}/record009/expected (100%) rename tests/idris2/{ => data}/record009/input (100%) rename tests/idris2/{ => data}/record009/record.idr (100%) rename tests/idris2/{ => data}/record009/run (52%) rename tests/idris2/{ => data}/record010/expected (100%) rename tests/idris2/{ => data}/record010/record.idr (100%) create mode 100755 tests/idris2/data/record010/run rename tests/idris2/{ => data}/record011/Issue2095.idr (100%) rename tests/idris2/{ => data}/record011/expected (100%) create mode 100755 tests/idris2/data/record011/run rename tests/idris2/{ => data}/record012/Issue2065.idr (100%) rename tests/idris2/{ => data}/record012/expected (100%) create mode 100755 tests/idris2/data/record012/run rename tests/idris2/{ => data}/record013/Issue1945.idr (100%) rename tests/idris2/{ => data}/record013/expected (100%) rename tests/idris2/{ => data}/record013/input (100%) rename tests/idris2/{ => data}/record013/run (55%) rename tests/idris2/{ => data}/record014/Issue1404.idr (100%) rename tests/idris2/{ => data}/record014/expected (100%) create mode 100755 tests/idris2/data/record014/run rename tests/idris2/{ => data}/record015/Issue2176.idr (100%) rename tests/idris2/{ => data}/record015/expected (100%) create mode 100755 tests/idris2/data/record015/run rename tests/idris2/{ => data}/record016/HoleRecord.idr (100%) rename tests/idris2/{ => data}/record016/expected (100%) rename tests/idris2/{ => data}/record016/input (100%) rename tests/idris2/{ => data}/record016/run (56%) rename tests/idris2/{ => data}/record017/RecordOptions.idr (100%) rename tests/idris2/{ => data}/record017/expected (100%) rename tests/idris2/{ => data}/record017/run (54%) rename tests/idris2/{ => data}/record018/expected (100%) rename tests/idris2/{ => data}/record018/mut.idr (100%) create mode 100755 tests/idris2/data/record018/run rename tests/idris2/{ => data}/record019/BindParams.idr (100%) rename tests/idris2/{ => data}/record019/expected (100%) rename tests/idris2/{ => data}/record019/run (52%) delete mode 100755 tests/idris2/data001/run delete mode 100755 tests/idris2/data002/run rename tests/idris2/{ => debug}/debug001/TypePat.idr (100%) rename tests/idris2/{ => debug}/debug001/expected (100%) rename tests/idris2/{ => debug}/debug001/input (100%) rename tests/idris2/{ => debug}/debug001/run (53%) delete mode 100755 tests/idris2/docs001/run delete mode 100755 tests/idris2/docs005/run rename tests/idris2/{ => error}/error001/Error.idr (100%) rename tests/idris2/{ => error}/error001/expected (100%) create mode 100755 tests/idris2/error/error001/run rename tests/idris2/{ => error}/error002/Error.idr (100%) rename tests/idris2/{ => error}/error002/expected (100%) create mode 100755 tests/idris2/error/error002/run rename tests/idris2/{ => error}/error003/Error.idr (100%) rename tests/idris2/{ => error}/error003/expected (100%) create mode 100755 tests/idris2/error/error003/run rename tests/idris2/{ => error}/error004/Error1.idr (100%) rename tests/idris2/{ => error}/error004/Error2.idr (100%) rename tests/idris2/{ => error}/error004/expected (100%) rename tests/idris2/{ => error}/error004/run (59%) rename tests/idris2/{ => error}/error005/IfErr.idr (100%) rename tests/idris2/{ => error}/error005/expected (100%) create mode 100755 tests/idris2/error/error005/run rename tests/idris2/{ => error}/error006/IfErr.idr (100%) rename tests/idris2/{ => error}/error006/expected (100%) create mode 100755 tests/idris2/error/error006/run rename tests/idris2/{ => error}/error007/CongErr.idr (100%) rename tests/idris2/{ => error}/error007/expected (100%) create mode 100755 tests/idris2/error/error007/run rename tests/idris2/{ => error}/error008/expected (100%) rename tests/idris2/{error009 => error/error008}/input (100%) rename tests/idris2/{ => error}/error008/run (57%) rename tests/idris2/{ => error}/error009/Exists.idr (100%) rename tests/idris2/{ => error}/error009/expected (100%) rename tests/idris2/{interface025 => error/error009}/input (100%) rename tests/idris2/{ => error}/error009/run (52%) rename tests/idris2/{ => error}/error010/Loop.idr (100%) rename tests/idris2/{ => error}/error010/expected (100%) create mode 100755 tests/idris2/error/error010/run rename tests/idris2/{ => error}/error011/ConstructorDuplicate.idr (100%) rename tests/idris2/{ => error}/error011/expected (100%) rename tests/idris2/{ => error}/error011/run (57%) rename tests/idris2/{ => error}/error012/expected (100%) rename tests/idris2/{ => error}/error012/run (100%) rename tests/idris2/{ => error}/error013/Issue361.idr (100%) rename tests/idris2/{ => error}/error013/expected (100%) create mode 100755 tests/idris2/error/error013/run rename tests/idris2/{ => error}/error014/Issue735.idr (100%) rename tests/idris2/{ => error}/error014/expected (100%) create mode 100755 tests/idris2/error/error014/run rename tests/idris2/{ => error}/error015/Issue110.idr (100%) rename tests/idris2/{ => error}/error015/expected (100%) create mode 100755 tests/idris2/error/error015/run rename tests/idris2/{ => error}/error016/Issue1230.idr (100%) rename tests/idris2/{ => error}/error016/expected (100%) rename tests/idris2/{ => error}/error016/input (100%) rename tests/idris2/{ => error}/error016/run (60%) rename tests/idris2/{ => error}/error017/Issue962-case.idr (100%) rename tests/idris2/{ => error}/error017/Issue962.idr (100%) rename tests/idris2/{ => error}/error017/expected (100%) rename tests/idris2/{ => error}/error017/run (64%) rename tests/idris2/{ => error}/error018/Issue1031-2.idr (100%) rename tests/idris2/{ => error}/error018/Issue1031-3.idr (100%) rename tests/idris2/{ => error}/error018/Issue1031-4.idr (100%) rename tests/idris2/{ => error}/error018/Issue1031.idr (100%) rename tests/idris2/{ => error}/error018/expected (100%) rename tests/idris2/{ => error}/error018/run (78%) rename tests/idris2/{ => error}/error019/Error.idr (100%) rename tests/idris2/{ => error}/error019/expected (100%) rename tests/idris2/{ => error}/error019/run (51%) rename tests/idris2/{ => error}/error020/Error.idr (100%) rename tests/idris2/{ => error}/error020/expected (100%) rename tests/idris2/{ => error}/error020/run (51%) rename tests/idris2/{ => error}/error021/DeepAmbig.idr (100%) rename tests/idris2/{ => error}/error021/expected (100%) rename tests/idris2/{ => error}/error021/run (54%) rename tests/idris2/{ => error}/error022/UpdateLoc.idr (100%) rename tests/idris2/{ => error}/error022/expected (100%) create mode 100644 tests/idris2/error/error022/run rename tests/idris2/{ => error}/error023/Error1.idr (100%) rename tests/idris2/{ => error}/error023/Error2.idr (100%) rename tests/idris2/{ => error}/error023/expected (100%) rename tests/idris2/{ => error}/error023/run (59%) rename tests/idris2/{ => error}/error024/Error1.idr (100%) rename tests/idris2/{ => error}/error024/expected (100%) create mode 100755 tests/idris2/error/error024/run rename tests/idris2/{ => error}/error025/IAlternativePrints.idr (100%) rename tests/idris2/{ => error}/error025/expected (100%) rename tests/idris2/{ => error}/error025/run (55%) rename tests/idris2/{ => error}/error026/DoBlockFC.idr (100%) rename tests/idris2/{ => error}/error026/expected (100%) create mode 100755 tests/idris2/error/error026/run rename tests/idris2/{ => error}/error027/Issue2950.idr (100%) rename tests/idris2/{ => error}/error027/expected (100%) create mode 100755 tests/idris2/error/error027/run rename tests/idris2/{ => error}/perror001/PError.idr (100%) rename tests/idris2/{ => error}/perror001/expected (100%) create mode 100755 tests/idris2/error/perror001/run rename tests/idris2/{ => error}/perror002/PError.idr (100%) rename tests/idris2/{ => error}/perror002/expected (100%) create mode 100755 tests/idris2/error/perror002/run rename tests/idris2/{ => error}/perror003/PError.idr (100%) rename tests/idris2/{ => error}/perror003/PError2.idr (100%) rename tests/idris2/{ => error}/perror003/expected (100%) rename tests/idris2/{ => error}/perror003/run (60%) rename tests/idris2/{ => error}/perror004/PError.idr (100%) rename tests/idris2/{ => error}/perror004/expected (100%) create mode 100755 tests/idris2/error/perror004/run rename tests/idris2/{ => error}/perror005/PError.idr (100%) rename tests/idris2/{ => error}/perror005/expected (100%) create mode 100755 tests/idris2/error/perror005/run rename tests/idris2/{ => error}/perror006/PError.idr (100%) rename tests/idris2/{ => error}/perror006/expected (100%) create mode 100755 tests/idris2/error/perror006/run rename tests/idris2/{ => error}/perror007/StrError1.idr (100%) rename tests/idris2/{ => error}/perror007/StrError10.idr (100%) rename tests/idris2/{ => error}/perror007/StrError11.idr (100%) rename tests/idris2/{ => error}/perror007/StrError12.idr (100%) rename tests/idris2/{ => error}/perror007/StrError2.idr (100%) rename tests/idris2/{ => error}/perror007/StrError3.idr (100%) rename tests/idris2/{ => error}/perror007/StrError4.idr (100%) rename tests/idris2/{ => error}/perror007/StrError5.idr (100%) rename tests/idris2/{ => error}/perror007/StrError6.idr (100%) rename tests/idris2/{ => error}/perror007/StrError7.idr (100%) rename tests/idris2/{ => error}/perror007/StrError8.idr (100%) rename tests/idris2/{ => error}/perror007/StrError9.idr (100%) rename tests/idris2/{ => error}/perror007/expected (100%) rename tests/idris2/{ => error}/perror007/run (93%) rename tests/idris2/{ => error}/perror008/Issue1224a.idr (100%) rename tests/idris2/{ => error}/perror008/Issue1224b.idr (100%) rename tests/idris2/{ => error}/perror008/Issue710a.idr (100%) rename tests/idris2/{ => error}/perror008/Issue710b.idr (100%) rename tests/idris2/{ => error}/perror008/Issue710c.idr (100%) rename tests/idris2/{ => error}/perror008/Issue710d.idr (100%) rename tests/idris2/{ => error}/perror008/Issue710e.idr (100%) rename tests/idris2/{ => error}/perror008/Issue710f.idr (100%) rename tests/idris2/{ => error}/perror008/expected (100%) rename tests/idris2/{ => error}/perror008/run (90%) rename tests/idris2/{ => error}/perror009/Error1.idr (100%) rename tests/idris2/{ => error}/perror009/expected (100%) rename tests/idris2/{ => error}/perror009/run (52%) rename tests/idris2/{ => error}/perror010/NamedReturn1.idr (100%) rename tests/idris2/{ => error}/perror010/NamedReturn2.idr (100%) rename tests/idris2/{ => error}/perror010/NamedReturn3.idr (100%) rename tests/idris2/{ => error}/perror010/NamedReturn4.idr (100%) rename tests/idris2/{ => error}/perror010/TrailingLam.idr (100%) rename tests/idris2/{ => error}/perror010/expected (100%) rename tests/idris2/{ => error}/perror010/input (100%) rename tests/idris2/{ => error}/perror010/run (86%) rename tests/idris2/{ => error}/perror011/Issue1345.idr (100%) rename tests/idris2/{ => error}/perror011/Issue1496-1.idr (100%) rename tests/idris2/{ => error}/perror011/Issue1496-2.idr (100%) rename tests/idris2/{ => error}/perror011/Main.idr (100%) rename tests/idris2/{ => error}/perror011/Pretty.idr (100%) rename tests/idris2/{ => error}/perror011/expected (100%) rename tests/idris2/{ => error}/perror011/foo.ipkg (100%) rename tests/idris2/{ => error}/perror011/run (82%) rename tests/idris2/{ => error}/perror012/CaseParseError.idr (100%) rename tests/idris2/{ => error}/perror012/LamParseError.idr (100%) rename tests/idris2/{ => error}/perror012/expected (100%) rename tests/idris2/{ => error}/perror012/run (67%) rename tests/idris2/{ => error}/perror013/EmptyFailing.idr (100%) rename tests/idris2/{ => error}/perror013/EmptyMutual.idr (100%) rename tests/idris2/{ => error}/perror013/EmptyParameters.idr (100%) rename tests/idris2/{ => error}/perror013/EmptyUsing.idr (100%) rename tests/idris2/{ => error}/perror013/expected (100%) rename tests/idris2/{ => error}/perror013/run (83%) rename tests/idris2/{ => error}/perror013/x.ipkg (100%) rename tests/idris2/{ => error}/perror014/ParseList.idr (100%) rename tests/idris2/{ => error}/perror014/expected (100%) rename tests/idris2/{ => error}/perror014/run (54%) rename tests/idris2/{ => error}/perror015/ParseWith.idr (100%) rename tests/idris2/{ => error}/perror015/expected (100%) rename tests/idris2/{ => error}/perror015/run (54%) rename tests/idris2/{ => error}/perror016/ParseIf.idr (100%) rename tests/idris2/{ => error}/perror016/ParseIf2.idr (100%) rename tests/idris2/{ => error}/perror016/ParseIf3.idr (100%) rename tests/idris2/{ => error}/perror016/ParseIf4.idr (100%) rename tests/idris2/{ => error}/perror016/expected (100%) rename tests/idris2/{ => error}/perror016/run (81%) rename tests/idris2/{ => error}/perror017/ParseImpl.idr (100%) rename tests/idris2/{ => error}/perror017/expected (100%) rename tests/idris2/{ => error}/perror017/run (54%) rename tests/idris2/{ => error}/perror018/ParseRecord.idr (100%) rename tests/idris2/{ => error}/perror018/ParseRecord2.idr (100%) rename tests/idris2/{ => error}/perror018/ParseRecord3.idr (100%) rename tests/idris2/{ => error}/perror018/expected (100%) rename tests/idris2/{ => error}/perror018/run (79%) rename tests/idris2/{ => error}/perror019/ImplError.idr (100%) rename tests/idris2/{ => error}/perror019/expected (100%) create mode 100644 tests/idris2/error/perror019/run rename tests/idris2/{ => error}/perror020/Issue2769.idr (100%) rename tests/idris2/{ => error}/perror020/Issue2769b.idr (100%) rename tests/idris2/{ => error}/perror020/expected (100%) rename tests/idris2/{ => error}/perror020/run (63%) rename tests/idris2/{ => error}/perror021/Implicit.idr (100%) rename tests/idris2/{ => error}/perror021/expected (100%) create mode 100644 tests/idris2/error/perror021/run rename tests/idris2/{ => error}/perror022/Indent.idr (100%) rename tests/idris2/{ => error}/perror022/expected (100%) create mode 100644 tests/idris2/error/perror022/run rename tests/idris2/{ => error}/perror023/ParseError.idr (100%) rename tests/idris2/{ => error}/perror023/expected (100%) create mode 100755 tests/idris2/error/perror023/run rename tests/idris2/{ => error}/perror024/ParseError.idr (100%) rename tests/idris2/{ => error}/perror024/expected (100%) create mode 100755 tests/idris2/error/perror024/run rename tests/idris2/{ => error}/perror025/DataWhere.idr (100%) rename tests/idris2/{ => error}/perror025/expected (100%) create mode 100755 tests/idris2/error/perror025/run rename tests/idris2/{ => error}/perror026/Micro.idr (100%) rename tests/idris2/{ => error}/perror026/expected (100%) create mode 100755 tests/idris2/error/perror026/run rename tests/idris2/{ => error}/perror027/Outdent.idr (100%) rename tests/idris2/{ => error}/perror027/expected (100%) create mode 100755 tests/idris2/error/perror027/run rename tests/idris2/{ => error}/perror028/LetInDo.idr (100%) rename tests/idris2/{ => error}/perror028/expected (100%) create mode 100644 tests/idris2/error/perror028/run rename tests/idris2/{ => error}/perror029/DelayParse.idr (100%) rename tests/idris2/{ => error}/perror029/expected (100%) create mode 100644 tests/idris2/error/perror029/run delete mode 100755 tests/idris2/error001/run delete mode 100755 tests/idris2/error002/run delete mode 100755 tests/idris2/error003/run delete mode 100755 tests/idris2/error005/run delete mode 100755 tests/idris2/error006/run delete mode 100755 tests/idris2/error007/run delete mode 100755 tests/idris2/error010/run delete mode 100755 tests/idris2/error013/run delete mode 100755 tests/idris2/error014/run delete mode 100755 tests/idris2/error015/run delete mode 100644 tests/idris2/error022/run delete mode 100755 tests/idris2/error024/run delete mode 100755 tests/idris2/error026/run delete mode 100755 tests/idris2/error027/run delete mode 100755 tests/idris2/eta001/run rename tests/idris2/{ => evaluator}/evaluator001/Issue650.idr (100%) rename tests/idris2/{ => evaluator}/evaluator001/expected (100%) create mode 100755 tests/idris2/evaluator/evaluator001/run rename tests/idris2/{ => evaluator}/evaluator002/Lib.idr (100%) rename tests/idris2/{ => evaluator}/evaluator002/Main.idr (100%) rename tests/idris2/{ => evaluator}/evaluator002/expected (100%) rename tests/idris2/{ => evaluator}/evaluator002/input (100%) rename tests/idris2/{ => evaluator}/evaluator002/run (64%) rename tests/idris2/{ => evaluator}/evaluator003/Issue705.idr (100%) rename tests/idris2/{ => evaluator}/evaluator003/expected (100%) rename tests/idris2/{ => evaluator}/evaluator003/input (100%) rename tests/idris2/{ => evaluator}/evaluator003/run (54%) rename tests/idris2/{ => evaluator}/evaluator004/Issue1282.idr (100%) rename tests/idris2/{ => evaluator}/evaluator004/expected (100%) rename tests/idris2/{ => evaluator}/evaluator004/input (100%) rename tests/idris2/{ => evaluator}/evaluator004/run (55%) rename tests/idris2/{ => evaluator}/interpreter001/expected (100%) rename tests/idris2/{ => evaluator}/interpreter001/input (100%) create mode 100755 tests/idris2/evaluator/interpreter001/run rename tests/idris2/{ => evaluator}/interpreter002/expected (100%) rename tests/idris2/{ => evaluator}/interpreter002/input (100%) rename tests/idris2/{ => evaluator}/interpreter002/run (54%) rename tests/idris2/{ => evaluator}/interpreter003/expected (100%) rename tests/idris2/{ => evaluator}/interpreter003/input (100%) create mode 100755 tests/idris2/evaluator/interpreter003/run rename tests/idris2/{ => evaluator}/interpreter004/expected (100%) rename tests/idris2/{ => evaluator}/interpreter004/input (100%) create mode 100755 tests/idris2/evaluator/interpreter004/run rename tests/idris2/{ => evaluator}/interpreter005/Issue37.idr (100%) rename tests/idris2/{ => evaluator}/interpreter005/Issue37.lidr (100%) rename tests/idris2/{ => evaluator}/interpreter005/expected (100%) rename tests/idris2/{ => evaluator}/interpreter005/input (100%) rename tests/idris2/{ => evaluator}/interpreter005/run (54%) rename tests/idris2/{ => evaluator}/interpreter006/expected (100%) rename tests/idris2/{ => evaluator}/interpreter006/input (100%) create mode 100755 tests/idris2/evaluator/interpreter006/run rename tests/idris2/{ => evaluator}/interpreter007/expected (100%) rename tests/idris2/{ => evaluator}/interpreter007/input (100%) create mode 100755 tests/idris2/evaluator/interpreter007/run rename tests/idris2/{ => evaluator}/interpreter008/Issue2041.idr (100%) rename tests/idris2/{ => evaluator}/interpreter008/expected (100%) rename tests/idris2/{ => evaluator}/interpreter008/input (100%) rename tests/idris2/{ => evaluator}/interpreter008/run (55%) rename tests/idris2/{ => evaluator}/spec001/Desc.idr (100%) rename tests/idris2/{ => evaluator}/spec001/Desc2.idr (100%) rename tests/idris2/{ => evaluator}/spec001/Identity.idr (100%) rename tests/idris2/{ => evaluator}/spec001/Mult3.idr (100%) rename tests/idris2/{ => evaluator}/spec001/expected (100%) rename tests/idris2/{ => evaluator}/spec001/run (91%) delete mode 100755 tests/idris2/evaluator001/run rename tests/idris2/{ => failing}/failing001/Fail.idr (100%) rename tests/idris2/{ => failing}/failing001/expected (100%) create mode 100755 tests/idris2/failing/failing001/run rename tests/idris2/{ => failing}/failing002/FailingBug.idr (100%) rename tests/idris2/{ => failing}/failing002/expected (100%) rename tests/idris2/{ => failing}/failing002/run (58%) rename tests/idris2/{ => failing}/failing003/FailingTotality.idr (100%) rename tests/idris2/{ => failing}/failing003/expected (100%) rename tests/idris2/{ => failing}/failing003/run (52%) rename tests/idris2/{ => failing}/failing004/Issue2821.idr (100%) rename tests/idris2/{ => failing}/failing004/expected (100%) create mode 100755 tests/idris2/failing/failing004/run delete mode 100755 tests/idris2/failing001/run delete mode 100755 tests/idris2/failing004/run delete mode 100644 tests/idris2/idiom001/run delete mode 100644 tests/idris2/import005/run delete mode 100644 tests/idris2/import007/run rename tests/idris2/{ => interactive}/interactive001/LocType.idr (100%) rename tests/idris2/{ => interactive}/interactive001/expected (100%) rename tests/idris2/{ => interactive}/interactive001/input (100%) rename tests/idris2/{ => interactive}/interactive001/run (53%) rename tests/idris2/{ => interactive}/interactive002/IEdit.idr (100%) rename tests/idris2/{ => interactive}/interactive002/expected (100%) rename tests/idris2/{ => interactive}/interactive002/input (100%) rename tests/idris2/{ => interactive}/interactive002/run (52%) rename tests/idris2/{ => interactive}/interactive003/IEdit.idr (100%) rename tests/idris2/{ => interactive}/interactive003/IEdit2.idr (100%) rename tests/idris2/{ => interactive}/interactive003/expected (100%) rename tests/idris2/{ => interactive}/interactive003/input (100%) rename tests/idris2/{ => interactive}/interactive003/input2 (100%) rename tests/idris2/{ => interactive}/interactive003/run (68%) rename tests/idris2/{ => interactive}/interactive004/IEdit.idr (100%) rename tests/idris2/{ => interactive}/interactive004/expected (100%) rename tests/idris2/{ => interactive}/interactive004/input (100%) rename tests/idris2/{interactive006 => interactive/interactive004}/run (52%) rename tests/idris2/{ => interactive}/interactive005/IEdit.idr (100%) rename tests/idris2/{ => interactive}/interactive005/expected (100%) rename tests/idris2/{ => interactive}/interactive005/input (100%) rename tests/idris2/{interactive004 => interactive/interactive005}/run (52%) rename tests/idris2/{ => interactive}/interactive006/IEdit.idr (100%) rename tests/idris2/{ => interactive}/interactive006/expected (100%) rename tests/idris2/{ => interactive}/interactive006/input (100%) rename tests/idris2/{interactive005 => interactive/interactive006}/run (52%) rename tests/idris2/{ => interactive}/interactive007/IEdit.idr (100%) rename tests/idris2/{ => interactive}/interactive007/expected (100%) rename tests/idris2/{ => interactive}/interactive007/input (100%) create mode 100755 tests/idris2/interactive/interactive007/run rename tests/idris2/{ => interactive}/interactive008/IEdit.idr (100%) rename tests/idris2/{ => interactive}/interactive008/expected (100%) rename tests/idris2/{ => interactive}/interactive008/input (100%) create mode 100755 tests/idris2/interactive/interactive008/run rename tests/idris2/{ => interactive}/interactive009/Door.idr (100%) rename tests/idris2/{ => interactive}/interactive009/expected (100%) rename tests/idris2/{ => interactive}/interactive009/input (100%) rename tests/idris2/{ => interactive}/interactive009/run (51%) rename tests/idris2/{ => interactive}/interactive010/IEdit.idr (100%) rename tests/idris2/{ => interactive}/interactive010/expected (100%) rename tests/idris2/{ => interactive}/interactive010/input (100%) create mode 100755 tests/idris2/interactive/interactive010/run rename tests/idris2/{ => interactive}/interactive011/IEdit.idr (100%) rename tests/idris2/{ => interactive}/interactive011/expected (100%) rename tests/idris2/{ => interactive}/interactive011/input (100%) create mode 100755 tests/idris2/interactive/interactive011/run rename tests/idris2/{ => interactive}/interactive012/WithLift.idr (100%) rename tests/idris2/{ => interactive}/interactive012/expected (100%) rename tests/idris2/{ => interactive}/interactive012/input (100%) rename tests/idris2/{ => interactive}/interactive012/run (54%) rename tests/idris2/{ => interactive}/interactive013/Spacing.idr (100%) rename tests/idris2/{ => interactive}/interactive013/expected (100%) rename tests/idris2/{ => interactive}/interactive013/input (100%) rename tests/idris2/{ => interactive}/interactive013/run (53%) rename tests/idris2/{ => interactive}/interactive014/case.idr (100%) rename tests/idris2/{ => interactive}/interactive014/expected (100%) rename tests/idris2/{ => interactive}/interactive014/input (100%) rename tests/idris2/{ => interactive}/interactive014/run (51%) rename tests/idris2/{ => interactive}/interactive015/IEdit.idr (100%) rename tests/idris2/{ => interactive}/interactive015/expected (100%) rename tests/idris2/{ => interactive}/interactive015/input (100%) create mode 100755 tests/idris2/interactive/interactive015/run rename tests/idris2/{ => interactive}/interactive016/Cont.idr (100%) rename tests/idris2/{ => interactive}/interactive016/expected (100%) rename tests/idris2/{ => interactive}/interactive016/input (100%) rename tests/idris2/{ => interactive}/interactive016/run (51%) rename tests/idris2/{ => interactive}/interactive017/RLE.idr (100%) rename tests/idris2/{ => interactive}/interactive017/expected (100%) rename tests/idris2/{ => interactive}/interactive017/input (100%) rename tests/idris2/{ => interactive}/interactive017/run (50%) rename tests/idris2/{ => interactive}/interactive018/PlusPrf.idr (100%) rename tests/idris2/{ => interactive}/interactive018/expected (100%) rename tests/idris2/{ => interactive}/interactive018/input (100%) rename tests/idris2/{ => interactive}/interactive018/run (53%) rename tests/idris2/{ => interactive}/interactive019/TypeSearch.idr (100%) rename tests/idris2/{ => interactive}/interactive019/expected (100%) rename tests/idris2/{ => interactive}/interactive019/input (100%) rename tests/idris2/{ => interactive}/interactive019/run (56%) rename tests/idris2/{ => interactive}/interactive020/Issue835.idr (100%) rename tests/idris2/{ => interactive}/interactive020/expected (100%) rename tests/idris2/{ => interactive}/interactive020/input (100%) rename tests/idris2/{ => interactive}/interactive020/run (54%) rename tests/idris2/{ => interactive}/interactive021/TypeAtDoNotation.idr (100%) rename tests/idris2/{ => interactive}/interactive021/expected (100%) rename tests/idris2/{ => interactive}/interactive021/input (100%) rename tests/idris2/{ => interactive}/interactive021/run (60%) rename tests/idris2/{ => interactive}/interactive022/TypeAtBangSyntax.idr (100%) rename tests/idris2/{ => interactive}/interactive022/expected (100%) rename tests/idris2/{ => interactive}/interactive022/input (100%) rename tests/idris2/{ => interactive}/interactive022/run (60%) rename tests/idris2/{ => interactive}/interactive023/TypeAtLambda.idr (100%) rename tests/idris2/{ => interactive}/interactive023/expected (100%) rename tests/idris2/{ => interactive}/interactive023/input (100%) rename tests/idris2/{ => interactive}/interactive023/run (57%) rename tests/idris2/{ => interactive}/interactive024/TypeAtAsPatterns.idr (100%) rename tests/idris2/{ => interactive}/interactive024/expected (100%) rename tests/idris2/{ => interactive}/interactive024/input (100%) rename tests/idris2/{ => interactive}/interactive024/run (60%) rename tests/idris2/{ => interactive}/interactive025/TypeAtInterfaces.idr (100%) rename tests/idris2/{ => interactive}/interactive025/expected (100%) rename tests/idris2/{ => interactive}/interactive025/input (100%) rename tests/idris2/{ => interactive}/interactive025/run (60%) rename tests/idris2/{ => interactive}/interactive026/TypeAtRecords.idr (100%) rename tests/idris2/{ => interactive}/interactive026/expected (100%) rename tests/idris2/{ => interactive}/interactive026/input (100%) rename tests/idris2/{ => interactive}/interactive026/run (58%) rename tests/idris2/{ => interactive}/interactive027/TypeAtLocalVars.idr (100%) rename tests/idris2/{ => interactive}/interactive027/expected (100%) rename tests/idris2/{ => interactive}/interactive027/input (100%) rename tests/idris2/{ => interactive}/interactive027/run (60%) rename tests/idris2/{ => interactive}/interactive028/expected (100%) rename tests/idris2/{ => interactive}/interactive028/input (100%) create mode 100644 tests/idris2/interactive/interactive028/run rename tests/idris2/{ => interactive}/interactive029/Issue834.idr (100%) rename tests/idris2/{ => interactive}/interactive029/expected (100%) rename tests/idris2/{ => interactive}/interactive029/input (100%) rename tests/idris2/{ => interactive}/interactive029/run (54%) rename tests/idris2/{ => interactive}/interactive030/expected (100%) rename tests/idris2/{ => interactive}/interactive030/input (100%) create mode 100644 tests/idris2/interactive/interactive030/run rename tests/idris2/{ => interactive}/interactive031/Signatures.idr (100%) rename tests/idris2/{ => interactive}/interactive031/expected (100%) rename tests/idris2/{ => interactive}/interactive031/input (100%) rename tests/idris2/{ => interactive}/interactive031/run (56%) rename tests/idris2/{ => interactive}/interactive032/Uninh.idr (100%) rename tests/idris2/{ => interactive}/interactive032/expected (100%) rename tests/idris2/{ => interactive}/interactive032/input (100%) rename tests/idris2/{ => interactive}/interactive032/run (52%) rename tests/idris2/{ => interactive}/interactive033/UninhIndent.idr (100%) rename tests/idris2/{ => interactive}/interactive033/expected (100%) rename tests/idris2/{ => interactive}/interactive033/input (100%) rename tests/idris2/{ => interactive}/interactive033/run (57%) rename tests/idris2/{ => interactive}/interactive034/expected (100%) rename tests/idris2/{ => interactive}/interactive034/input (100%) rename tests/idris2/{ => interactive}/interactive034/run (53%) rename tests/idris2/{ => interactive}/interactive034/timeout.idr (100%) rename tests/idris2/{ => interactive}/interactive035/expected (100%) rename tests/idris2/{ => interactive}/interactive035/input (100%) rename tests/idris2/{ => interactive}/interactive035/run (52%) rename tests/idris2/{ => interactive}/interactive035/unify.idr (100%) rename tests/idris2/{ => interactive}/interactive036/casefn.idr (100%) rename tests/idris2/{ => interactive}/interactive036/expected (100%) rename tests/idris2/{ => interactive}/interactive036/input (100%) rename tests/idris2/{ => interactive}/interactive036/run (52%) rename tests/idris2/{ => interactive}/interactive037/Holes.idr (100%) rename tests/idris2/{ => interactive}/interactive037/expected (100%) rename tests/idris2/{ => interactive}/interactive037/input (100%) rename tests/idris2/{ => interactive}/interactive037/run (52%) rename tests/idris2/{ => interactive}/interactive038/IEdit.idr (100%) rename tests/idris2/{ => interactive}/interactive038/expected (100%) rename tests/idris2/{ => interactive}/interactive038/input (100%) create mode 100755 tests/idris2/interactive/interactive038/run rename tests/idris2/{ => interactive}/interactive039/CS_Syntax.idr (100%) rename tests/idris2/{ => interactive}/interactive039/expected (100%) rename tests/idris2/{ => interactive}/interactive039/input (100%) rename tests/idris2/{ => interactive}/interactive039/run (55%) rename tests/idris2/{ => interactive}/interactive040/expected (100%) rename tests/idris2/{ => interactive}/interactive040/input (100%) create mode 100755 tests/idris2/interactive/interactive040/run rename tests/idris2/{ => interactive}/interactive041/Issue1741.idr (100%) rename tests/idris2/{ => interactive}/interactive041/expected (100%) rename tests/idris2/{ => interactive}/interactive041/input (100%) rename tests/idris2/{ => interactive}/interactive041/run (55%) rename tests/idris2/{ => interactive}/interactive042/Issue35-2.idr (100%) rename tests/idris2/{ => interactive}/interactive042/Issue35.idr (100%) rename tests/idris2/{ => interactive}/interactive042/expected (100%) rename tests/idris2/{ => interactive}/interactive042/input (100%) rename tests/idris2/{ => interactive}/interactive042/run (87%) rename tests/idris2/{ => interactive}/interactive043/ImplicitSplits.idr (100%) rename tests/idris2/{ => interactive}/interactive043/expected (100%) rename tests/idris2/{ => interactive}/interactive043/input (100%) rename tests/idris2/{ => interactive}/interactive043/run (90%) rename tests/idris2/{ => interactive}/interactive044/SplitShadow.idr (100%) rename tests/idris2/{ => interactive}/interactive044/expected (100%) rename tests/idris2/{ => interactive}/interactive044/input (100%) rename tests/idris2/{ => interactive}/interactive044/run (84%) rename tests/idris2/{ => interactive}/interactive045/Issue1742.idr (100%) rename tests/idris2/{ => interactive}/interactive045/expected (100%) rename tests/idris2/{ => interactive}/interactive045/input (100%) rename tests/idris2/{ => interactive}/interactive045/run (83%) rename tests/idris2/{ => interactive}/interactive046/Issue2712.idr (100%) rename tests/idris2/{ => interactive}/interactive046/expected (100%) create mode 100755 tests/idris2/interactive/interactive046/run delete mode 100755 tests/idris2/interactive007/run delete mode 100755 tests/idris2/interactive008/run delete mode 100755 tests/idris2/interactive010/run delete mode 100755 tests/idris2/interactive011/run delete mode 100755 tests/idris2/interactive015/run delete mode 100644 tests/idris2/interactive028/run delete mode 100644 tests/idris2/interactive030/run delete mode 100755 tests/idris2/interactive038/run delete mode 100755 tests/idris2/interactive040/run delete mode 100755 tests/idris2/interactive046/run rename tests/idris2/{ => interface}/interface001/IFace.idr (100%) rename tests/idris2/{ => interface}/interface001/IFace1.idr (100%) rename tests/idris2/{ => interface}/interface001/Stuff.idr (100%) rename tests/idris2/{ => interface}/interface001/expected (100%) rename tests/idris2/{ => interface}/interface001/input (100%) rename tests/idris2/{ => interface}/interface001/input1 (100%) rename tests/idris2/{ => interface}/interface001/run (76%) rename tests/idris2/{ => interface}/interface002/Functor.idr (100%) rename tests/idris2/{ => interface}/interface002/Stuff.idr (100%) rename tests/idris2/{ => interface}/interface002/expected (100%) rename tests/idris2/{ => interface}/interface002/input (100%) rename tests/idris2/{ => interface}/interface002/run (63%) rename tests/idris2/{ => interface}/interface003/Do.idr (100%) rename tests/idris2/{ => interface}/interface003/expected (100%) rename tests/idris2/{ => interface}/interface003/input (100%) create mode 100755 tests/idris2/interface/interface003/run rename tests/idris2/{ => interface}/interface004/Do.idr (100%) rename tests/idris2/{ => interface}/interface004/expected (100%) rename tests/idris2/{ => interface}/interface004/input (100%) create mode 100755 tests/idris2/interface/interface004/run rename tests/idris2/{ => interface}/interface005/Deps.idr (100%) rename tests/idris2/{ => interface}/interface005/expected (100%) create mode 100755 tests/idris2/interface/interface005/run rename tests/idris2/{ => interface}/interface006/Apply.idr (100%) rename tests/idris2/{ => interface}/interface006/Biapplicative.idr (100%) rename tests/idris2/{ => interface}/interface006/Bimonad.idr (100%) rename tests/idris2/{ => interface}/interface006/expected (100%) create mode 100755 tests/idris2/interface/interface006/run rename tests/idris2/{ => interface}/interface007/A.idr (100%) rename tests/idris2/{ => interface}/interface007/expected (100%) rename tests/idris2/{ => interface}/interface007/run (57%) rename tests/idris2/{ => interface}/interface008/Deps.idr (100%) rename tests/idris2/{ => interface}/interface008/expected (100%) create mode 100755 tests/idris2/interface/interface008/run rename tests/idris2/{ => interface}/interface009/Odd.idr (100%) rename tests/idris2/{ => interface}/interface009/expected (100%) rename tests/idris2/{ => interface}/interface009/input (100%) rename tests/idris2/{ => interface}/interface009/run (50%) rename tests/idris2/{ => interface}/interface010/Dep.idr (100%) rename tests/idris2/{ => interface}/interface010/expected (100%) create mode 100755 tests/idris2/interface/interface010/run rename tests/idris2/{ => interface}/interface011/FuncImpl.idr (100%) rename tests/idris2/{ => interface}/interface011/expected (100%) create mode 100755 tests/idris2/interface/interface011/run rename tests/idris2/{ => interface}/interface012/Defmeth.idr (100%) rename tests/idris2/{ => interface}/interface012/expected (100%) create mode 100755 tests/idris2/interface/interface012/run rename tests/idris2/{ => interface}/interface013/TypeInt.idr (100%) rename tests/idris2/{ => interface}/interface013/expected (100%) create mode 100755 tests/idris2/interface/interface013/run rename tests/idris2/{ => interface}/interface014/DepInt.idr (100%) rename tests/idris2/{ => interface}/interface014/expected (100%) create mode 100755 tests/idris2/interface/interface014/run rename tests/idris2/{ => interface}/interface015/expected (100%) rename tests/idris2/{ => interface}/interface015/gnu.idr (100%) create mode 100755 tests/idris2/interface/interface015/run rename tests/idris2/{ => interface}/interface016/TwoNum.idr (100%) rename tests/idris2/{ => interface}/interface016/expected (100%) create mode 100755 tests/idris2/interface/interface016/run rename tests/idris2/{ => interface}/interface017/Tricho.idr (100%) rename tests/idris2/{ => interface}/interface017/expected (100%) create mode 100755 tests/idris2/interface/interface017/run rename tests/idris2/{ => interface}/interface018/Sized.idr (100%) rename tests/idris2/{ => interface}/interface018/Sized2.idr (100%) rename tests/idris2/{ => interface}/interface018/Sized3.idr (100%) rename tests/idris2/{ => interface}/interface018/expected (100%) rename tests/idris2/{ => interface}/interface018/input (100%) rename tests/idris2/{ => interface}/interface018/run (71%) rename tests/idris2/{ => interface}/interface019/LocalHints.idr (100%) rename tests/idris2/{ => interface}/interface019/expected (100%) create mode 100755 tests/idris2/interface/interface019/run rename tests/idris2/{ => interface}/interface020/LocalInterface.idr (100%) rename tests/idris2/{ => interface}/interface020/expected (100%) rename tests/idris2/{ => interface}/interface020/run (52%) rename tests/idris2/{ => interface}/interface021/LocalHint.idr (100%) rename tests/idris2/{ => interface}/interface021/expected (100%) create mode 100755 tests/idris2/interface/interface021/run rename tests/idris2/{ => interface}/interface022/DotMethod.idr (100%) rename tests/idris2/{ => interface}/interface022/expected (100%) rename tests/idris2/{ => interface}/interface022/input (100%) rename tests/idris2/{ => interface}/interface022/run (55%) rename tests/idris2/{ => interface}/interface023/AppComp.idr (100%) rename tests/idris2/{ => interface}/interface023/expected (100%) rename tests/idris2/{ => interface}/interface023/input (100%) rename tests/idris2/{ => interface}/interface023/run (53%) rename tests/idris2/{ => interface}/interface024/EH.idr (100%) rename tests/idris2/{ => interface}/interface024/expected (100%) rename tests/idris2/{ => interface}/interface024/run (51%) rename tests/idris2/{ => interface}/interface025/AutoSearchHide1.idr (100%) rename tests/idris2/{ => interface}/interface025/AutoSearchHide2.idr (100%) rename tests/idris2/{ => interface}/interface025/expected (100%) rename tests/idris2/{interface026 => interface/interface025}/input (100%) rename tests/idris2/{ => interface}/interface025/run (60%) rename tests/idris2/{ => interface}/interface026/UninhabitedRec.idr (100%) rename tests/idris2/{ => interface}/interface026/expected (100%) rename tests/idris2/{record007 => interface/interface026}/input (100%) rename tests/idris2/{ => interface}/interface026/run (59%) rename tests/idris2/{ => interface}/interface027/expected (100%) rename tests/idris2/{ => interface}/interface027/input (100%) rename tests/idris2/{ => interface}/interface027/params.idr (100%) rename tests/idris2/{ => interface}/interface027/run (52%) rename tests/idris2/{ => interface}/interface028/InterfaceArgs.idr (100%) rename tests/idris2/{ => interface}/interface028/expected (100%) rename tests/idris2/{ => interface}/interface028/run (51%) rename tests/idris2/{ => interface}/interface029/ForwardImpl1.idr (100%) rename tests/idris2/{ => interface}/interface029/ForwardImpl2.idr (100%) rename tests/idris2/{ => interface}/interface029/ForwardImpl3.idr (100%) rename tests/idris2/{ => interface}/interface029/expected (100%) rename tests/idris2/{ => interface}/interface029/run (73%) delete mode 100755 tests/idris2/interface003/run delete mode 100755 tests/idris2/interface004/run delete mode 100755 tests/idris2/interface005/run delete mode 100755 tests/idris2/interface006/run delete mode 100755 tests/idris2/interface008/run delete mode 100755 tests/idris2/interface010/run delete mode 100755 tests/idris2/interface011/run delete mode 100755 tests/idris2/interface012/run delete mode 100755 tests/idris2/interface013/run delete mode 100755 tests/idris2/interface014/run delete mode 100755 tests/idris2/interface015/run delete mode 100755 tests/idris2/interface016/run delete mode 100755 tests/idris2/interface017/run delete mode 100755 tests/idris2/interface019/run delete mode 100755 tests/idris2/interface021/run delete mode 100755 tests/idris2/interpolation003/run delete mode 100755 tests/idris2/interpreter001/run delete mode 100755 tests/idris2/interpreter003/run delete mode 100755 tests/idris2/interpreter004/run delete mode 100755 tests/idris2/interpreter006/run delete mode 100755 tests/idris2/interpreter007/run rename tests/idris2/{ => linear}/linear001/Door.idr (100%) rename tests/idris2/{ => linear}/linear001/Stuff.idr (100%) rename tests/idris2/{ => linear}/linear001/expected (100%) rename tests/idris2/{ => linear}/linear001/run (63%) rename tests/idris2/{ => linear}/linear002/Door.idr (100%) rename tests/idris2/{ => linear}/linear002/Stuff.idr (100%) rename tests/idris2/{ => linear}/linear002/expected (100%) rename tests/idris2/{ => linear}/linear002/input (100%) rename tests/idris2/{ => linear}/linear002/run (61%) rename tests/idris2/{ => linear}/linear003/Linear.idr (100%) rename tests/idris2/{ => linear}/linear003/expected (100%) rename tests/idris2/{ => linear}/linear003/input (100%) rename tests/idris2/{ => linear}/linear003/run (62%) rename tests/idris2/{ => linear}/linear005/Door.idr (100%) rename tests/idris2/{ => linear}/linear005/Linear.idr (100%) rename tests/idris2/{ => linear}/linear005/expected (100%) rename tests/idris2/{ => linear}/linear005/input (100%) rename tests/idris2/{ => linear}/linear005/run (51%) rename tests/idris2/{ => linear}/linear006/ZFun.idr (100%) rename tests/idris2/{ => linear}/linear006/expected (100%) rename tests/idris2/{ => linear}/linear006/input (100%) rename tests/idris2/{ => linear}/linear006/run (51%) rename tests/idris2/{ => linear}/linear007/LCase.idr (100%) rename tests/idris2/{ => linear}/linear007/expected (100%) create mode 100755 tests/idris2/linear/linear007/run rename tests/idris2/{ => linear}/linear008/Door.idr (100%) rename tests/idris2/{ => linear}/linear008/expected (100%) create mode 100644 tests/idris2/linear/linear008/run rename tests/idris2/{ => linear}/linear009/expected (100%) rename tests/idris2/{ => linear}/linear009/input (100%) rename tests/idris2/{ => linear}/linear009/qtt.idr (100%) rename tests/idris2/{ => linear}/linear009/run (50%) rename tests/idris2/{ => linear}/linear010/Door.idr (100%) rename tests/idris2/{ => linear}/linear010/expected (100%) rename tests/idris2/{ => linear}/linear010/run (52%) rename tests/idris2/{ => linear}/linear011/Network.idr (100%) rename tests/idris2/{ => linear}/linear011/expected (100%) rename tests/idris2/{ => linear}/linear011/input (100%) rename tests/idris2/{ => linear}/linear011/run (67%) rename tests/idris2/{ => linear}/linear012/expected (100%) rename tests/idris2/{ => linear}/linear012/input (100%) rename tests/idris2/{ => linear}/linear012/linholes.idr (100%) rename tests/idris2/{ => linear}/linear012/run (54%) rename tests/idris2/{ => linear}/linear013/Issue758.idr (100%) rename tests/idris2/{ => linear}/linear013/expected (100%) rename tests/idris2/{ => linear}/linear013/input (100%) rename tests/idris2/{ => linear}/linear013/run (54%) rename tests/idris2/{ => linear}/linear014/Issue55.idr (100%) rename tests/idris2/{ => linear}/linear014/expected (100%) create mode 100644 tests/idris2/linear/linear014/run rename tests/idris2/{ => linear}/linear015/Issue1861.idr (100%) rename tests/idris2/{ => linear}/linear015/expected (100%) create mode 100644 tests/idris2/linear/linear015/run rename tests/idris2/{ => linear}/linear016/Issue2895.idr (100%) rename tests/idris2/{ => linear}/linear016/Issue2895_2.idr (100%) rename tests/idris2/{ => linear}/linear016/expected (100%) rename tests/idris2/{ => linear}/linear016/run (64%) rename tests/idris2/{ => linear}/linear017/As.idr (100%) rename tests/idris2/{ => linear}/linear017/expected (100%) create mode 100644 tests/idris2/linear/linear017/run delete mode 100644 tests/idris2/linear004/Erase.idr delete mode 100644 tests/idris2/linear004/Stuff.idr delete mode 100644 tests/idris2/linear004/expected delete mode 100644 tests/idris2/linear004/input delete mode 100755 tests/idris2/linear004/run delete mode 100755 tests/idris2/linear007/run delete mode 100644 tests/idris2/linear008/run delete mode 100644 tests/idris2/linear014/run delete mode 100644 tests/idris2/linear015/run delete mode 100644 tests/idris2/linear017/run rename tests/idris2/{ => literate}/literate001/IEdit.lidr (100%) rename tests/idris2/{ => literate}/literate001/expected (100%) rename tests/idris2/{ => literate}/literate001/input (100%) rename tests/idris2/{ => literate}/literate001/run (52%) rename tests/idris2/{ => literate}/literate002/IEdit.lidr (100%) rename tests/idris2/{ => literate}/literate002/IEdit2.lidr (100%) rename tests/idris2/{ => literate}/literate002/expected (100%) rename tests/idris2/{ => literate}/literate002/input (100%) rename tests/idris2/{ => literate}/literate002/input2 (100%) rename tests/idris2/{ => literate}/literate002/run (69%) rename tests/idris2/{ => literate}/literate003/IEdit.lidr (100%) rename tests/idris2/{ => literate}/literate003/expected (100%) rename tests/idris2/{ => literate}/literate003/input (100%) rename tests/idris2/{literate005 => literate/literate003}/run (52%) rename tests/idris2/{ => literate}/literate004/IEdit.lidr (100%) rename tests/idris2/{ => literate}/literate004/expected (100%) rename tests/idris2/{ => literate}/literate004/input (100%) rename tests/idris2/{literate003 => literate/literate004}/run (52%) rename tests/idris2/{ => literate}/literate005/IEdit.lidr (100%) rename tests/idris2/{ => literate}/literate005/expected (100%) rename tests/idris2/{ => literate}/literate005/input (100%) rename tests/idris2/{literate004 => literate/literate005}/run (52%) rename tests/idris2/{ => literate}/literate006/Door.lidr (100%) rename tests/idris2/{ => literate}/literate006/expected (100%) rename tests/idris2/{ => literate}/literate006/input (100%) rename tests/idris2/{ => literate}/literate006/run (52%) rename tests/idris2/{ => literate}/literate007/IEdit.lidr (100%) rename tests/idris2/{ => literate}/literate007/IEdit.org (100%) rename tests/idris2/{ => literate}/literate007/IEditOrg.org (100%) rename tests/idris2/{ => literate}/literate007/expected (100%) rename tests/idris2/{ => literate}/literate007/input (100%) rename tests/idris2/{ => literate}/literate007/input2 (100%) rename tests/idris2/{ => literate}/literate007/run (70%) rename tests/idris2/{ => literate}/literate008/IEdit.lidr (100%) rename tests/idris2/{ => literate}/literate008/expected (100%) rename tests/idris2/{ => literate}/literate008/input (100%) create mode 100755 tests/idris2/literate/literate008/run rename tests/idris2/{ => literate}/literate009/WithLift.lidr (100%) rename tests/idris2/{ => literate}/literate009/expected (100%) rename tests/idris2/{ => literate}/literate009/input (100%) rename tests/idris2/{ => literate}/literate009/run (55%) rename tests/idris2/{ => literate}/literate010/MyFirstIdrisProgram.org (100%) rename tests/idris2/{ => literate}/literate010/expected (100%) rename tests/idris2/{ => literate}/literate010/run (56%) rename tests/idris2/{ => literate}/literate011/IEdit.md (100%) rename tests/idris2/{ => literate}/literate011/expected (100%) rename tests/idris2/{ => literate}/literate011/input (100%) rename tests/idris2/{ => literate}/literate011/run (51%) rename tests/idris2/{ => literate}/literate012/IEdit.org (100%) rename tests/idris2/{ => literate}/literate012/IEdit2.org (100%) rename tests/idris2/{ => literate}/literate012/expected (100%) rename tests/idris2/{ => literate}/literate012/input (100%) rename tests/idris2/{ => literate}/literate012/run (52%) rename tests/idris2/{ => literate}/literate013/Lit.lidr (100%) rename tests/idris2/{ => literate}/literate013/LitTeX.tex (100%) rename tests/idris2/{ => literate}/literate013/expected (100%) rename tests/idris2/{ => literate}/literate013/run (57%) rename tests/idris2/{ => literate}/literate014/expected (100%) rename tests/idris2/{ => literate}/literate014/input (100%) rename tests/idris2/{ => literate}/literate014/run (52%) rename tests/idris2/{ => literate}/literate014/with.lidr (100%) rename tests/idris2/{ => literate}/literate015/case.lidr (100%) rename tests/idris2/{ => literate}/literate015/expected (100%) rename tests/idris2/{ => literate}/literate015/input (100%) rename tests/idris2/{ => literate}/literate015/run (52%) rename tests/idris2/{ => literate}/literate016/IEdit.org (100%) rename tests/idris2/{ => literate}/literate016/IEdit2.org (100%) rename tests/idris2/{ => literate}/literate016/expected (100%) rename tests/idris2/{ => literate}/literate016/input (100%) rename tests/idris2/{ => literate}/literate016/run (52%) rename tests/idris2/{ => literate}/literate017/.gitignore (100%) rename tests/idris2/{ => literate}/literate017/expected (100%) rename tests/idris2/{ => literate}/literate017/input (100%) rename tests/idris2/{ => literate}/literate017/project-expected.ipkg (100%) rename tests/idris2/{ => literate}/literate017/run (82%) rename tests/idris2/{ => literate}/literate017/src/A/A.org (100%) rename tests/idris2/{ => literate}/literate017/src/A/AB.lidr (100%) rename tests/idris2/{ => literate}/literate017/src/B/B.md (100%) rename tests/idris2/{ => literate}/literate017/src/B/BA.markdown (100%) rename tests/idris2/{ => literate}/literate017/src/Main.idr (100%) delete mode 100755 tests/idris2/literate008/run rename tests/idris2/{ => misc}/docs001/expected (100%) rename tests/idris2/{ => misc}/docs001/input (100%) create mode 100755 tests/idris2/misc/docs001/run rename tests/idris2/{ => misc}/docs002/Doc.idr (100%) rename tests/idris2/{ => misc}/docs002/expected (100%) rename tests/idris2/{ => misc}/docs002/input (100%) rename tests/idris2/{ => misc}/docs002/run (50%) rename tests/idris2/{ => misc}/docs003/RecordDoc.idr (100%) rename tests/idris2/{ => misc}/docs003/expected (100%) rename tests/idris2/{ => misc}/docs003/input (100%) rename tests/idris2/{ => misc}/docs003/run (55%) rename tests/idris2/{ => misc}/docs004/DocImpl.idr (100%) rename tests/idris2/{ => misc}/docs004/List.idr (100%) rename tests/idris2/{ => misc}/docs004/expected (100%) rename tests/idris2/{ => misc}/docs004/input (100%) rename tests/idris2/{ => misc}/docs004/input2 (100%) rename tests/idris2/{ => misc}/docs004/run (73%) rename tests/idris2/{ => misc}/docs005/expected (100%) rename tests/idris2/{ => misc}/docs005/input (100%) create mode 100755 tests/idris2/misc/docs005/run rename tests/idris2/{ => misc}/eta001/Issue1370.idr (100%) rename tests/idris2/{ => misc}/eta001/expected (100%) create mode 100755 tests/idris2/misc/eta001/run rename tests/idris2/{ => misc}/golden001/000-hello/expected (100%) rename tests/idris2/{ => misc}/golden001/000-hello/run (100%) mode change 100644 => 100755 rename tests/idris2/{ => misc}/golden001/Main.idr (100%) rename tests/idris2/{ => misc}/golden001/Test.idr (100%) rename tests/idris2/{ => misc}/golden001/expected (100%) rename tests/idris2/{ => misc}/golden001/hello.ipkg (100%) rename tests/idris2/{ => misc}/golden001/run (81%) rename tests/idris2/{ => misc}/golden001/test.ipkg (100%) rename tests/idris2/{ => misc}/import001/Mult.idr (100%) rename tests/idris2/{ => misc}/import001/Nat.idr (100%) rename tests/idris2/{ => misc}/import001/Test.idr (100%) rename tests/idris2/{ => misc}/import001/expected (100%) rename tests/idris2/{ => misc}/import001/input (100%) rename tests/idris2/{ => misc}/import001/run (90%) rename tests/idris2/{ => misc}/import002/Mult.idr (100%) rename tests/idris2/{ => misc}/import002/Nat.idr (100%) rename tests/idris2/{ => misc}/import002/Test.idr (100%) rename tests/idris2/{ => misc}/import002/expected (100%) rename tests/idris2/{ => misc}/import002/run (54%) rename tests/idris2/{ => misc}/import003/A.idr (100%) rename tests/idris2/{ => misc}/import003/B.idr (100%) rename tests/idris2/{ => misc}/import003/C.idr (100%) rename tests/idris2/{ => misc}/import003/expected (100%) rename tests/idris2/{ => misc}/import003/input (100%) rename tests/idris2/{ => misc}/import003/run (64%) rename tests/idris2/{ => misc}/import004/Cycle1.idr (100%) rename tests/idris2/{ => misc}/import004/Cycle2.idr (100%) rename tests/idris2/{ => misc}/import004/Loop.idr (100%) rename tests/idris2/{ => misc}/import004/expected (100%) rename tests/idris2/{ => misc}/import004/run (59%) rename tests/idris2/{ => misc}/import005/As.idr (100%) rename tests/idris2/{ => misc}/import005/Test.idr (100%) rename tests/idris2/{ => misc}/import005/expected (100%) rename tests/idris2/{ => misc}/import005/input (100%) create mode 100644 tests/idris2/misc/import005/run rename tests/idris2/{ => misc}/import006/A/B.idr (100%) rename tests/idris2/{ => misc}/import006/A/C.idr (100%) rename tests/idris2/{ => misc}/import006/cyclic.ipkg (100%) rename tests/idris2/{ => misc}/import006/expected (100%) rename tests/idris2/{ => misc}/import006/run (53%) rename tests/idris2/{ => misc}/import007/Mod.idr (100%) rename tests/idris2/{ => misc}/import007/Mod1.idr (100%) rename tests/idris2/{ => misc}/import007/Mod2.idr (100%) rename tests/idris2/{ => misc}/import007/expected (100%) create mode 100644 tests/idris2/misc/import007/run rename tests/idris2/{ => misc}/import008/Exe/Mod.idr (100%) rename tests/idris2/{ => misc}/import008/Exe/depends/lib-0 (100%) rename tests/idris2/{ => misc}/import008/Exe/exe.ipkg (100%) rename tests/idris2/{ => misc}/import008/Lib/Conf.idr (100%) rename tests/idris2/{ => misc}/import008/Lib/lib.ipkg (100%) rename tests/idris2/{ => misc}/import008/expected (100%) rename tests/idris2/{ => misc}/import008/run (83%) rename tests/idris2/{ => misc}/import009/Import.idr (100%) rename tests/idris2/{ => misc}/import009/Infix.idr (100%) rename tests/idris2/{ => misc}/import009/Prefix.idr (100%) rename tests/idris2/{ => misc}/import009/Resugar.idr (100%) rename tests/idris2/{ => misc}/import009/Test.idr (100%) rename tests/idris2/{ => misc}/import009/expected (100%) rename tests/idris2/{ => misc}/import009/run (68%) rename tests/idris2/{ => misc}/inlining001/Inlining.idr (100%) rename tests/idris2/{ => misc}/inlining001/expected (100%) rename tests/idris2/{ => misc}/inlining001/input (100%) rename tests/idris2/{ => misc}/inlining001/run (54%) rename tests/idris2/{ => misc}/lazy001/Lazy.idr (100%) rename tests/idris2/{ => misc}/lazy001/expected (100%) rename tests/idris2/{ => misc}/lazy001/input (100%) rename tests/idris2/{ => misc}/lazy001/run (61%) rename tests/idris2/{ => misc}/lazy002/LazyFoldlM.idr (100%) rename tests/idris2/{ => misc}/lazy002/expected (100%) rename tests/idris2/{ => misc}/lazy002/input (100%) rename tests/idris2/{ => misc}/lazy002/run (63%) rename tests/idris2/{ => misc}/lazy003/DelayLam.idr (100%) rename tests/idris2/{ => misc}/lazy003/expected (100%) rename tests/idris2/{ => misc}/lazy003/input (100%) rename tests/idris2/{ => misc}/lazy003/run (54%) rename tests/idris2/{ => misc}/namespace001/Dup.idr (100%) rename tests/idris2/{ => misc}/namespace001/Scope.idr (100%) rename tests/idris2/{ => misc}/namespace001/expected (100%) rename tests/idris2/{ => misc}/namespace001/run (72%) rename tests/idris2/{ => misc}/namespace002/Issue1313.idr (100%) rename tests/idris2/{ => misc}/namespace002/expected (100%) rename tests/idris2/{ => misc}/namespace002/run (54%) rename tests/idris2/{ => misc}/namespace003/Test.idr (100%) rename tests/idris2/{ => misc}/namespace003/expected (100%) rename tests/idris2/{ => misc}/namespace003/run (50%) rename tests/idris2/{ => misc}/namespace004/Export.idr (100%) rename tests/idris2/{ => misc}/namespace004/expected (100%) create mode 100644 tests/idris2/misc/namespace004/run rename tests/idris2/{ => misc}/namespace005/Lib1.idr (100%) rename tests/idris2/{ => misc}/namespace005/Lib2.idr (100%) rename tests/idris2/{ => misc}/namespace005/LibPre1.idr (100%) rename tests/idris2/{ => misc}/namespace005/LibPre2.idr (100%) rename tests/idris2/{ => misc}/namespace005/Main0.idr (100%) rename tests/idris2/{ => misc}/namespace005/Main1.idr (100%) rename tests/idris2/{ => misc}/namespace005/Main3.idr (100%) rename tests/idris2/{ => misc}/namespace005/MainConflict.idr (100%) rename tests/idris2/{ => misc}/namespace005/MainFail.idr (100%) rename tests/idris2/{ => misc}/namespace005/MainPre0.idr (100%) rename tests/idris2/{ => misc}/namespace005/MainPre1.idr (100%) rename tests/idris2/{ => misc}/namespace005/NonConflict1.idr (100%) rename tests/idris2/{ => misc}/namespace005/NonConflict2.idr (100%) rename tests/idris2/{ => misc}/namespace005/expected (100%) rename tests/idris2/{ => misc}/namespace005/run (90%) rename tests/idris2/{ => misc}/params001/expected (100%) rename tests/idris2/{ => misc}/params001/param.idr (100%) rename tests/idris2/{ => misc}/params001/parambad.idr (100%) rename tests/idris2/{ => misc}/params001/run (60%) rename tests/idris2/{ => misc}/params002/ParamsPrint.idr (100%) rename tests/idris2/{ => misc}/params002/expected (100%) rename tests/idris2/{ => misc}/params002/input (100%) rename tests/idris2/{ => misc}/params002/run (57%) rename tests/idris2/{ => misc}/params003/casesplit.idr (100%) rename tests/idris2/{ => misc}/params003/expected (100%) rename tests/idris2/{ => misc}/params003/input (100%) rename tests/idris2/{ => misc}/params003/run (55%) rename tests/idris2/{ => misc}/pretty001/Issue1328A.idr (100%) rename tests/idris2/{ => misc}/pretty001/expected (100%) rename tests/idris2/{ => misc}/pretty001/input (100%) rename tests/idris2/{ => misc}/pretty001/run (56%) rename tests/idris2/{ => misc}/pretty002/expected (100%) rename tests/idris2/{ => misc}/pretty002/input (100%) create mode 100644 tests/idris2/misc/pretty002/run rename tests/idris2/{ => misc}/primloop/PrimLoop.idr (100%) rename tests/idris2/{ => misc}/primloop/expected (100%) create mode 100644 tests/idris2/misc/primloop/run rename tests/idris2/{ => misc}/quantifiers001/TestQuantifiers.idr (100%) rename tests/idris2/{ => misc}/quantifiers001/expected (100%) rename tests/idris2/{ => misc}/quantifiers001/input (100%) rename tests/idris2/{ => misc}/quantifiers001/run (60%) rename tests/idris2/{ => misc}/real001/Channel.idr (100%) rename tests/idris2/{ => misc}/real001/Linear.idr (100%) rename tests/idris2/{ => misc}/real001/MakeChans.idr (100%) rename tests/idris2/{ => misc}/real001/TestProto.idr (100%) rename tests/idris2/{ => misc}/real001/expected (100%) rename tests/idris2/{ => misc}/real001/run (63%) rename tests/idris2/{ => misc}/real002/Store.idr (100%) rename tests/idris2/{ => misc}/real002/StoreL.idr (100%) rename tests/idris2/{ => misc}/real002/expected (100%) rename tests/idris2/{ => misc}/real002/run (58%) rename tests/idris2/{ => misc}/unification001/Issue647.idr (100%) rename tests/idris2/{ => misc}/unification001/expected (100%) create mode 100755 tests/idris2/misc/unification001/run rename tests/idris2/{ => misc}/with003/Main.idr (100%) rename tests/idris2/{ => misc}/with003/expected (100%) rename tests/idris2/{ => misc}/with003/input (100%) rename tests/idris2/{ => misc}/with003/run (51%) delete mode 100644 tests/idris2/namespace004/run rename tests/idris2/{ => perf}/perf001/Big.idr (100%) rename tests/idris2/{ => perf}/perf001/expected (100%) create mode 100755 tests/idris2/perf/perf001/run rename tests/idris2/{ => perf}/perf002/Big.idr (100%) rename tests/idris2/{ => perf}/perf002/expected (100%) create mode 100755 tests/idris2/perf/perf002/run rename tests/idris2/{ => perf}/perf003/Auto.idr (100%) rename tests/idris2/{ => perf}/perf003/expected (100%) create mode 100755 tests/idris2/perf/perf003/run rename tests/idris2/{ => perf}/perf004/bigdpair.idr (100%) rename tests/idris2/{ => perf}/perf004/expected (100%) create mode 100755 tests/idris2/perf/perf004/run rename tests/idris2/{ => perf}/perf005/Bad1.idr (100%) rename tests/idris2/{ => perf}/perf005/Bad2.idr (100%) rename tests/idris2/{ => perf}/perf005/Bad3.idr (100%) rename tests/idris2/{ => perf}/perf005/Lambda.idr (100%) rename tests/idris2/{ => perf}/perf005/NoRegression.idr (100%) rename tests/idris2/{ => perf}/perf005/expected (100%) rename tests/idris2/{ => perf}/perf005/run (73%) rename tests/idris2/{ => perf}/perf007/Slooow.idr (100%) rename tests/idris2/{ => perf}/perf007/expected (100%) create mode 100644 tests/idris2/perf/perf007/run rename tests/idris2/{ => perf}/perf008/FinPerf.idr (100%) rename tests/idris2/{ => perf}/perf008/expected (100%) create mode 100644 tests/idris2/perf/perf008/run rename tests/idris2/{ => perf}/perf009/A.idr (100%) rename tests/idris2/{ => perf}/perf009/B.idr (100%) rename tests/idris2/{ => perf}/perf009/C.idr (100%) rename tests/idris2/{ => perf}/perf009/expected (100%) create mode 100644 tests/idris2/perf/perf009/run rename tests/idris2/{ => perf}/perf010/Printf.idr (100%) rename tests/idris2/{ => perf}/perf010/expected (100%) create mode 100644 tests/idris2/perf/perf010/run rename tests/idris2/{ => perf}/perf011/A.idr (100%) rename tests/idris2/{ => perf}/perf011/B.idr (100%) rename tests/idris2/{ => perf}/perf011/C.idr (100%) rename tests/idris2/{ => perf}/perf011/expected (100%) create mode 100644 tests/idris2/perf/perf011/run rename tests/idris2/{ => perf}/perf012/Main.idr (100%) rename tests/idris2/{ => perf}/perf012/expected (100%) rename tests/idris2/{ => perf}/perf012/run (51%) rename tests/idris2/{ => perf}/perf2202/Church.idr (100%) rename tests/idris2/{ => perf}/perf2202/expected (100%) rename tests/idris2/{ => perf}/perf2202/many10000.idr (100%) rename tests/idris2/{ => perf}/perf2202/run (61%) delete mode 100755 tests/idris2/perf001/run delete mode 100755 tests/idris2/perf002/run delete mode 100755 tests/idris2/perf003/run delete mode 100755 tests/idris2/perf004/run delete mode 100644 tests/idris2/perf007/run delete mode 100644 tests/idris2/perf008/run delete mode 100644 tests/idris2/perf009/run delete mode 100644 tests/idris2/perf010/run delete mode 100644 tests/idris2/perf011/run delete mode 100755 tests/idris2/perror001/run delete mode 100755 tests/idris2/perror002/run delete mode 100755 tests/idris2/perror004/run delete mode 100755 tests/idris2/perror005/run delete mode 100755 tests/idris2/perror006/run delete mode 100644 tests/idris2/perror019/run delete mode 100644 tests/idris2/perror021/run delete mode 100644 tests/idris2/perror022/run delete mode 100755 tests/idris2/perror023/run delete mode 100755 tests/idris2/perror024/run delete mode 100755 tests/idris2/perror025/run delete mode 100755 tests/idris2/perror026/run delete mode 100755 tests/idris2/perror027/run delete mode 100644 tests/idris2/perror028/run delete mode 100644 tests/idris2/perror029/run rename tests/idris2/{ => pkg}/pkg001/Dummy.idr (100%) rename tests/idris2/{ => pkg}/pkg001/dummy.ipkg (100%) rename tests/idris2/{ => pkg}/pkg001/expected (100%) rename tests/idris2/{ => pkg}/pkg001/run (52%) rename tests/idris2/{ => pkg}/pkg002/dummy.ipkg (100%) rename tests/idris2/{ => pkg}/pkg002/expected (100%) rename tests/idris2/{ => pkg}/pkg002/run (67%) rename tests/idris2/{ => pkg}/pkg002/src/Top/Dummy.idr (100%) rename tests/idris2/{ => pkg}/pkg003/Main.idr (100%) rename tests/idris2/{ => pkg}/pkg003/expected (100%) rename tests/idris2/{ => pkg}/pkg003/run (92%) rename tests/idris2/{ => pkg}/pkg003/testpkg.ipkg (100%) rename tests/idris2/{ => pkg}/pkg004/Dummy.idr (100%) rename tests/idris2/{ => pkg}/pkg004/dummy.ipkg (100%) rename tests/idris2/{ => pkg}/pkg004/expected (100%) rename tests/idris2/{ => pkg}/pkg004/input (100%) rename tests/idris2/{ => pkg}/pkg004/run (58%) rename tests/idris2/{ => pkg}/pkg005/Foo.idr (100%) rename tests/idris2/{ => pkg}/pkg005/expected (100%) rename tests/idris2/{ => pkg}/pkg005/foo.ipkg (100%) rename tests/idris2/{ => pkg}/pkg005/input (100%) rename tests/idris2/{ => pkg}/pkg005/run (78%) rename tests/idris2/{ => pkg}/pkg006/depends/bar-1.0.1/bar.ipkg (100%) rename tests/idris2/{ => pkg}/pkg006/depends/bar-baz/bar-baz.ipkg (100%) rename tests/idris2/{ => pkg}/pkg006/depends/foo-0.5/foo.ipkg (100%) rename tests/idris2/{ => pkg}/pkg006/depends/foo-bar-1.3.1/foo-bar.ipkg (100%) rename tests/idris2/{ => pkg}/pkg006/depends/quux/quux.ipkg (100%) rename tests/idris2/{ => pkg}/pkg006/expected (100%) rename tests/idris2/{ => pkg}/pkg006/run (84%) rename tests/idris2/{ => pkg}/pkg006/test1.ipkg (100%) rename tests/idris2/{ => pkg}/pkg006/test2.ipkg (100%) rename tests/idris2/{ => pkg}/pkg006/test3.ipkg (100%) rename tests/idris2/{ => pkg}/pkg006/test4.ipkg (100%) rename tests/idris2/{ => pkg}/pkg006/test5.ipkg (100%) rename tests/idris2/{ => pkg}/pkg007/A/Path/Of/Dires/First.idr (100%) rename tests/idris2/{ => pkg}/pkg007/A/Path/Of/Dires/Second.idr (100%) rename tests/idris2/{ => pkg}/pkg007/Another/Fourth.idr (100%) rename tests/idris2/{ => pkg}/pkg007/Another/One/Third.idr (100%) rename tests/idris2/{ => pkg}/pkg007/expected (100%) rename tests/idris2/{ => pkg}/pkg007/input (100%) rename tests/idris2/{ => pkg}/pkg007/input2 (100%) rename tests/idris2/{ => pkg}/pkg007/run (87%) rename tests/idris2/{ => pkg}/pkg007/src/And/A/Proof.idr (100%) rename tests/idris2/{ => pkg}/pkg007/src/Yet/Another/Path.idr (100%) rename tests/idris2/{ => pkg}/pkg008/Bar.idr (100%) rename tests/idris2/{ => pkg}/pkg008/Foo.idr (100%) rename tests/idris2/{ => pkg}/pkg008/bar.ipkg (100%) rename tests/idris2/{ => pkg}/pkg008/expected (100%) rename tests/idris2/{ => pkg}/pkg008/foo.ipkg (100%) rename tests/idris2/{ => pkg}/pkg008/run (63%) rename tests/idris2/{ => pkg}/pkg009/expected (100%) rename tests/idris2/{ => pkg}/pkg009/run (71%) rename tests/idris2/{ => pkg}/pkg009/testpkg/Main.idr (100%) rename tests/idris2/{ => pkg}/pkg009/testpkg/testpkg.ipkg (100%) rename tests/idris2/{ => pkg}/pkg010/Main.idr (100%) rename tests/idris2/{ => pkg}/pkg010/expected (100%) rename tests/idris2/{ => pkg}/pkg010/run (90%) rename tests/idris2/{ => pkg}/pkg010/testpkg.ipkg (100%) rename tests/idris2/{ => pkg}/pkg011/dot-slash-dot/Main.idr (100%) rename tests/idris2/{ => pkg}/pkg011/dot-slash-dot/testpkg.ipkg (100%) rename tests/idris2/{ => pkg}/pkg011/dot-slash-name-slash/src/Main.idr (100%) rename tests/idris2/{ => pkg}/pkg011/dot-slash-name-slash/testpkg.ipkg (100%) rename tests/idris2/{ => pkg}/pkg011/dot-slash/Main.idr (100%) rename tests/idris2/{ => pkg}/pkg011/dot-slash/testpkg.ipkg (100%) rename tests/idris2/{ => pkg}/pkg011/dot/Main.idr (100%) rename tests/idris2/{ => pkg}/pkg011/dot/testpkg.ipkg (100%) rename tests/idris2/{ => pkg}/pkg011/expected (100%) rename tests/idris2/{ => pkg}/pkg011/run (94%) rename tests/idris2/{ => pkg}/pkg011/sibling/pkg/testpkg.ipkg (100%) rename tests/idris2/{ => pkg}/pkg011/sibling/src/Main.idr (100%) rename tests/idris2/{ => pkg}/pkg012/Main.idr (100%) rename tests/idris2/{ => pkg}/pkg012/expected (100%) rename tests/idris2/{ => pkg}/pkg012/run (54%) rename tests/idris2/{ => pkg}/pkg012/testpkg.ipkg (100%) rename tests/idris2/{ => pkg}/pkg013/depends/bar-0.7.2/bar.ipkg (100%) rename tests/idris2/{ => pkg}/pkg013/depends/foo-0.1.0/foo.ipkg (100%) rename tests/idris2/{ => pkg}/pkg013/expected (100%) rename tests/idris2/{ => pkg}/pkg013/run (74%) rename tests/idris2/{ => pkg}/pkg013/test.ipkg (100%) rename tests/idris2/{ => pkg}/pkg014/depends/bar-0.1.0/bar.ipkg (100%) rename tests/idris2/{ => pkg}/pkg014/depends/bar-0.1.1/bar.ipkg (100%) rename tests/idris2/{ => pkg}/pkg014/depends/baz-0.1.0/baz.ipkg (100%) rename tests/idris2/{ => pkg}/pkg014/depends/baz-0.2.0/baz.ipkg (100%) rename tests/idris2/{ => pkg}/pkg014/depends/baz-0.3.0/baz.ipkg (100%) rename tests/idris2/{ => pkg}/pkg014/depends/foo-0.1.0/foo.ipkg (100%) rename tests/idris2/{ => pkg}/pkg014/depends/foo-0.2.0/foo.ipkg (100%) rename tests/idris2/{ => pkg}/pkg014/depends/foo-0.3.0/foo.ipkg (100%) rename tests/idris2/{ => pkg}/pkg014/expected (100%) rename tests/idris2/{ => pkg}/pkg014/run (74%) rename tests/idris2/{ => pkg}/pkg014/test.ipkg (100%) rename tests/idris2/{ => pkg}/pkg015/depends/bar-0.1.0/bar.ipkg (100%) rename tests/idris2/{ => pkg}/pkg015/depends/bar-0.1.1/bar.ipkg (100%) rename tests/idris2/{ => pkg}/pkg015/depends/baz-0.1.0/baz.ipkg (100%) rename tests/idris2/{ => pkg}/pkg015/depends/baz-0.2.0/baz.ipkg (100%) rename tests/idris2/{ => pkg}/pkg015/depends/baz-0.3.0/baz.ipkg (100%) rename tests/idris2/{ => pkg}/pkg015/depends/foo-0.1.0/foo.ipkg (100%) rename tests/idris2/{ => pkg}/pkg015/depends/foo-0.2.0/foo.ipkg (100%) rename tests/idris2/{ => pkg}/pkg015/depends/foo-0.3.0/foo.ipkg (100%) rename tests/idris2/{ => pkg}/pkg015/depends/prz-0.1.0/prz.ipkg (100%) rename tests/idris2/{ => pkg}/pkg015/expected (100%) rename tests/idris2/{ => pkg}/pkg015/run (100%) rename tests/idris2/{ => pkg}/pkg015/test.ipkg (100%) rename tests/idris2/{ => pkg}/pkg016/bar.ipkg (100%) rename tests/idris2/{ => pkg}/pkg016/bar/Bar.idr (100%) rename tests/idris2/{ => pkg}/pkg016/baz.ipkg (100%) rename tests/idris2/{ => pkg}/pkg016/baz/Baz.idr (100%) rename tests/idris2/{ => pkg}/pkg016/expected (100%) rename tests/idris2/{ => pkg}/pkg016/foo.ipkg (100%) rename tests/idris2/{ => pkg}/pkg016/foo/Foo.idr (100%) rename tests/idris2/{ => pkg}/pkg016/run (100%) rename tests/idris2/{ => pkg}/pkg016/src/Test.idr (100%) rename tests/idris2/{ => pkg}/pkg016/test.ipkg (100%) rename tests/idris2/{ => pkg}/pkg017/a1/A.idr (100%) rename tests/idris2/{ => pkg}/pkg017/a1/a1.ipkg (100%) rename tests/idris2/{ => pkg}/pkg017/a2/A.idr (100%) rename tests/idris2/{ => pkg}/pkg017/a2/a2.ipkg (100%) rename tests/idris2/{ => pkg}/pkg017/b1/b1.ipkg (100%) rename tests/idris2/{ => pkg}/pkg017/b1/src/B1.idr (100%) rename tests/idris2/{ => pkg}/pkg017/b2/b2.ipkg (100%) rename tests/idris2/{ => pkg}/pkg017/b2/src/B2.idr (100%) rename tests/idris2/{ => pkg}/pkg017/expected (100%) rename tests/idris2/{ => pkg}/pkg017/input1 (100%) rename tests/idris2/{ => pkg}/pkg017/input2 (100%) rename tests/idris2/{ => pkg}/pkg017/run (95%) delete mode 100644 tests/idris2/positivity001/run delete mode 100644 tests/idris2/positivity002/run delete mode 100644 tests/idris2/positivity003/run delete mode 100644 tests/idris2/pretty002/run delete mode 100644 tests/idris2/primloop/run delete mode 100755 tests/idris2/record003/run delete mode 100755 tests/idris2/record010/run delete mode 100755 tests/idris2/record011/run delete mode 100755 tests/idris2/record012/run delete mode 100755 tests/idris2/record014/run delete mode 100755 tests/idris2/record015/run delete mode 100755 tests/idris2/record018/run rename tests/idris2/{ => reflection}/reflection001/expected (100%) rename tests/idris2/{ => reflection}/reflection001/input (100%) rename tests/idris2/{ => reflection}/reflection001/quote.idr (100%) rename tests/idris2/{ => reflection}/reflection001/run (52%) rename tests/idris2/{ => reflection}/reflection002/expected (100%) rename tests/idris2/{ => reflection}/reflection002/input (100%) rename tests/idris2/{ => reflection}/reflection002/power.idr (100%) rename tests/idris2/{ => reflection}/reflection002/run (52%) rename tests/idris2/{ => reflection}/reflection003/expected (100%) rename tests/idris2/{ => reflection}/reflection003/refprims.idr (100%) create mode 100755 tests/idris2/reflection/reflection003/run rename tests/idris2/{ => reflection}/reflection004/expected (100%) rename tests/idris2/{ => reflection}/reflection004/input (100%) rename tests/idris2/{ => reflection}/reflection004/refdecl.idr (100%) rename tests/idris2/{ => reflection}/reflection004/run (53%) rename tests/idris2/{ => reflection}/reflection005/expected (100%) rename tests/idris2/{ => reflection}/reflection005/input (100%) rename tests/idris2/{ => reflection}/reflection005/refdecl.idr (100%) rename tests/idris2/{ => reflection}/reflection005/run (53%) rename tests/idris2/{ => reflection}/reflection006/expected (100%) rename tests/idris2/{ => reflection}/reflection006/input (100%) rename tests/idris2/{ => reflection}/reflection006/refleq.idr (100%) create mode 100755 tests/idris2/reflection/reflection006/run rename tests/idris2/{ => reflection}/reflection007/NatExpr.idr (100%) rename tests/idris2/{ => reflection}/reflection007/expected (100%) rename tests/idris2/{ => reflection}/reflection007/input (100%) rename tests/idris2/{ => reflection}/reflection007/run (53%) rename tests/idris2/{ => reflection}/reflection008/Interp.idr (100%) rename tests/idris2/{ => reflection}/reflection008/expected (100%) rename tests/idris2/{ => reflection}/reflection008/input (100%) rename tests/idris2/{ => reflection}/reflection008/run (52%) rename tests/idris2/{ => reflection}/reflection009/expected (100%) rename tests/idris2/{ => reflection}/reflection009/perf.idr (100%) create mode 100755 tests/idris2/reflection/reflection009/run rename tests/idris2/{ => reflection}/reflection010/Name.idr (100%) rename tests/idris2/{ => reflection}/reflection010/expected (100%) rename tests/idris2/{ => reflection}/reflection010/run (68%) rename tests/idris2/{ => reflection}/reflection011/expected (100%) rename tests/idris2/{ => reflection}/reflection011/input (100%) rename tests/idris2/{ => reflection}/reflection011/run (52%) rename tests/idris2/{ => reflection}/reflection011/tryref.idr (100%) rename tests/idris2/{ => reflection}/reflection012/expected (100%) rename tests/idris2/{ => reflection}/reflection012/input (100%) rename tests/idris2/{ => reflection}/reflection012/nameinfo.idr (100%) rename tests/idris2/{ => reflection}/reflection012/run (54%) rename tests/idris2/{ => reflection}/reflection013/WithUnambig.idr (100%) rename tests/idris2/{ => reflection}/reflection013/expected (100%) create mode 100755 tests/idris2/reflection/reflection013/run rename tests/idris2/{ => reflection}/reflection014/expected (100%) rename tests/idris2/{ => reflection}/reflection014/refdecl.idr (100%) create mode 100755 tests/idris2/reflection/reflection014/run rename tests/idris2/{ => reflection}/reflection015/MacroRetFunc.idr (100%) rename tests/idris2/{ => reflection}/reflection015/expected (100%) rename tests/idris2/{ => reflection}/reflection015/run (50%) rename tests/idris2/{ => reflection}/reflection016/BindElabScBug.idr (100%) rename tests/idris2/{ => reflection}/reflection016/Eta.idr (100%) rename tests/idris2/{ => reflection}/reflection016/expected (100%) rename tests/idris2/{ => reflection}/reflection016/input (100%) rename tests/idris2/{ => reflection}/reflection016/run (66%) rename tests/idris2/{ => reflection}/reflection017/CanElabType.idr (100%) rename tests/idris2/{ => reflection}/reflection017/StillCantEscape.idr (100%) rename tests/idris2/{ => reflection}/reflection017/expected (100%) rename tests/idris2/{ => reflection}/reflection017/run (67%) rename tests/idris2/{ => reflection}/reflection018/AtTypeLevel.idr (100%) rename tests/idris2/{ => reflection}/reflection018/expected (100%) create mode 100755 tests/idris2/reflection/reflection018/run rename tests/idris2/{ => reflection}/reflection019/ElabScriptWarning.idr (100%) rename tests/idris2/{ => reflection}/reflection019/expected (100%) rename tests/idris2/{ => reflection}/reflection019/input (100%) rename tests/idris2/{ => reflection}/reflection019/run (61%) rename tests/idris2/{ => reflection}/reflection019/test.ipkg (100%) rename tests/idris2/{ => reflection}/reflection020/FromDecls.idr (100%) rename tests/idris2/{ => reflection}/reflection020/FromName.idr (100%) rename tests/idris2/{ => reflection}/reflection020/FromTTImp.idr (100%) rename tests/idris2/{ => reflection}/reflection020/expected (100%) rename tests/idris2/{ => reflection}/reflection020/run (94%) rename tests/idris2/{ => reflection}/reflection021/QuoteSearch.idr (100%) rename tests/idris2/{ => reflection}/reflection021/expected (100%) create mode 100755 tests/idris2/reflection/reflection021/run delete mode 100755 tests/idris2/reflection003/run delete mode 100755 tests/idris2/reflection006/run delete mode 100755 tests/idris2/reflection009/run delete mode 100755 tests/idris2/reflection013/run delete mode 100755 tests/idris2/reflection014/run delete mode 100755 tests/idris2/reflection018/run delete mode 100755 tests/idris2/reflection021/run rename tests/idris2/{ => reg}/reg001/D.idr (100%) rename tests/idris2/{ => reg}/reg001/expected (100%) create mode 100755 tests/idris2/reg/reg001/run rename tests/idris2/{ => reg}/reg002/expected (100%) rename tests/idris2/{ => reg}/reg002/linm.idr (100%) create mode 100755 tests/idris2/reg/reg002/run rename tests/idris2/{ => reg}/reg003/Holes.idr (100%) rename tests/idris2/{ => reg}/reg003/expected (100%) rename tests/idris2/{ => reg}/reg003/input (100%) rename tests/idris2/{ => reg}/reg003/run (52%) rename tests/idris2/{ => reg}/reg004/ambig.idr (100%) rename tests/idris2/{ => reg}/reg004/expected (100%) create mode 100755 tests/idris2/reg/reg004/run rename tests/idris2/{ => reg}/reg005/expected (100%) rename tests/idris2/{ => reg}/reg005/iftype.idr (100%) create mode 100755 tests/idris2/reg/reg005/run rename tests/idris2/{ => reg}/reg006/Cmd.idr (100%) rename tests/idris2/{ => reg}/reg006/expected (100%) create mode 100755 tests/idris2/reg/reg006/run rename tests/idris2/{ => reg}/reg007/Main.idr (100%) rename tests/idris2/{ => reg}/reg007/expected (100%) create mode 100755 tests/idris2/reg/reg007/run rename tests/idris2/{ => reg}/reg008/Vending.idr (100%) rename tests/idris2/{ => reg}/reg008/expected (100%) rename tests/idris2/{ => reg}/reg008/input (100%) rename tests/idris2/{ => reg}/reg008/run (66%) rename tests/idris2/{ => reg}/reg009/Case.idr (100%) rename tests/idris2/{ => reg}/reg009/expected (100%) create mode 100755 tests/idris2/reg/reg009/run rename tests/idris2/{ => reg}/reg010/Recordname.idr (100%) rename tests/idris2/{ => reg}/reg010/expected (100%) create mode 100755 tests/idris2/reg/reg010/run rename tests/idris2/{ => reg}/reg011/expected (100%) rename tests/idris2/{ => reg}/reg011/mut.idr (100%) create mode 100755 tests/idris2/reg/reg011/run rename tests/idris2/{ => reg}/reg012/Foo.idr (100%) rename tests/idris2/{ => reg}/reg012/expected (100%) create mode 100755 tests/idris2/reg/reg012/run rename tests/idris2/{ => reg}/reg013/UnboundImplicits.idr (100%) rename tests/idris2/{ => reg}/reg013/expected (100%) rename tests/idris2/{ => reg}/reg013/run (53%) rename tests/idris2/{ => reg}/reg014/casecase.idr (100%) rename tests/idris2/{ => reg}/reg014/expected (100%) rename tests/idris2/{ => reg}/reg014/run (61%) rename tests/idris2/{ => reg}/reg015/anyfail.idr (100%) rename tests/idris2/{ => reg}/reg015/expected (100%) create mode 100755 tests/idris2/reg/reg015/run rename tests/idris2/{ => reg}/reg016/Using.idr (100%) rename tests/idris2/{ => reg}/reg016/expected (100%) create mode 100755 tests/idris2/reg/reg016/run rename tests/idris2/{ => reg}/reg017/expected (100%) rename tests/idris2/{ => reg}/reg017/lammult.idr (100%) create mode 100755 tests/idris2/reg/reg017/run rename tests/idris2/{ => reg}/reg018/cycle.idr (100%) rename tests/idris2/{ => reg}/reg018/expected (100%) create mode 100755 tests/idris2/reg/reg018/run rename tests/idris2/{ => reg}/reg019/expected (100%) rename tests/idris2/{ => reg}/reg019/lazybug.idr (100%) create mode 100755 tests/idris2/reg/reg019/run rename tests/idris2/{ => reg}/reg020/expected (100%) rename tests/idris2/{ => reg}/reg020/input (100%) rename tests/idris2/{ => reg}/reg020/matchlits.idr (100%) rename tests/idris2/{ => reg}/reg020/run (55%) rename tests/idris2/{ => reg}/reg021/case.idr (100%) rename tests/idris2/{ => reg}/reg021/expected (100%) create mode 100755 tests/idris2/reg/reg021/run rename tests/idris2/{ => reg}/reg022/case.idr (100%) rename tests/idris2/{ => reg}/reg022/expected (100%) rename tests/idris2/{ => reg}/reg022/input (100%) create mode 100755 tests/idris2/reg/reg022/run rename tests/idris2/{ => reg}/reg023/boom.idr (100%) rename tests/idris2/{ => reg}/reg023/expected (100%) create mode 100755 tests/idris2/reg/reg023/run rename tests/idris2/{ => reg}/reg024/expected (100%) rename tests/idris2/{ => reg}/reg024/input (100%) rename tests/idris2/{ => reg}/reg024/run (52%) rename tests/idris2/{ => reg}/reg024/split.idr (100%) rename tests/idris2/{ => reg}/reg025/expected (100%) rename tests/idris2/{ => reg}/reg025/input (100%) rename tests/idris2/{ => reg}/reg025/lift.idr (100%) rename tests/idris2/{ => reg}/reg025/run (51%) rename tests/idris2/{ => reg}/reg026/Meh.idr (100%) rename tests/idris2/{ => reg}/reg026/expected (100%) create mode 100755 tests/idris2/reg/reg026/run rename tests/idris2/{ => reg}/reg027/expected (100%) rename tests/idris2/{ => reg}/reg027/pwhere.idr (100%) create mode 100755 tests/idris2/reg/reg027/run rename tests/idris2/{ => reg}/reg028/Test.idr (100%) rename tests/idris2/{ => reg}/reg028/expected (100%) create mode 100755 tests/idris2/reg/reg028/run rename tests/idris2/{ => reg}/reg029/expected (100%) rename tests/idris2/{ => reg}/reg029/lqueue.idr (100%) create mode 100755 tests/idris2/reg/reg029/run rename tests/idris2/{ => reg}/reg030/A.idr (100%) rename tests/idris2/{ => reg}/reg030/B.idr (100%) rename tests/idris2/{ => reg}/reg030/C.idr (100%) rename tests/idris2/{ => reg}/reg030/expected (100%) create mode 100755 tests/idris2/reg/reg030/run rename tests/idris2/{ => reg}/reg031/dpair.idr (100%) rename tests/idris2/{ => reg}/reg031/expected (100%) create mode 100755 tests/idris2/reg/reg031/run rename tests/idris2/{ => reg}/reg032/expected (100%) rename tests/idris2/{ => reg}/reg032/recupdate.idr (100%) create mode 100755 tests/idris2/reg/reg032/run rename tests/idris2/{ => reg}/reg033/DerivingEq.idr (100%) rename tests/idris2/{ => reg}/reg033/expected (100%) create mode 100755 tests/idris2/reg/reg033/run rename tests/idris2/{ => reg}/reg033/test.idr (100%) rename tests/idris2/{ => reg}/reg034/expected (100%) create mode 100755 tests/idris2/reg/reg034/run rename tests/idris2/{ => reg}/reg034/void.idr (100%) rename tests/idris2/{ => reg}/reg035/Implicit.idr (100%) rename tests/idris2/{ => reg}/reg035/expected (100%) create mode 100755 tests/idris2/reg/reg035/run rename tests/idris2/{ => reg}/reg036/Test.idr (100%) rename tests/idris2/{ => reg}/reg036/expected (100%) create mode 100755 tests/idris2/reg/reg036/run rename tests/idris2/{ => reg}/reg037/Test.idr (100%) rename tests/idris2/{ => reg}/reg037/expected (100%) create mode 100755 tests/idris2/reg/reg037/run rename tests/idris2/{ => reg}/reg038/Test1.idr (100%) rename tests/idris2/{ => reg}/reg038/Test2.idr (100%) rename tests/idris2/{ => reg}/reg038/expected (100%) rename tests/idris2/{ => reg}/reg038/run (57%) rename tests/idris2/{ => reg}/reg039/dupdup.idr (100%) rename tests/idris2/{ => reg}/reg039/expected (100%) create mode 100755 tests/idris2/reg/reg039/run rename tests/idris2/{ => reg}/reg040/CoverBug.idr (100%) rename tests/idris2/{ => reg}/reg040/expected (100%) create mode 100755 tests/idris2/reg/reg040/run rename tests/idris2/{ => reg}/reg041/expected (100%) create mode 100755 tests/idris2/reg/reg041/run rename tests/idris2/{ => reg}/reg041/tuple.idr (100%) rename tests/idris2/{ => reg}/reg042/NatOpts.idr (100%) rename tests/idris2/{ => reg}/reg042/expected (100%) rename tests/idris2/{ => reg}/reg042/input (100%) rename tests/idris2/{ => reg}/reg042/run (53%) rename tests/idris2/{ => reg}/reg043/NotFake.idr (100%) rename tests/idris2/{ => reg}/reg043/TestFake.idr (100%) rename tests/idris2/{ => reg}/reg043/expected (100%) rename tests/idris2/{ => reg}/reg043/run (81%) rename tests/idris2/{ => reg}/reg044/Methods.idr (100%) rename tests/idris2/{ => reg}/reg044/expected (100%) create mode 100755 tests/idris2/reg/reg044/run rename tests/idris2/{ => reg}/reg045/expected (100%) create mode 100755 tests/idris2/reg/reg045/run rename tests/idris2/{ => reg}/reg045/withparams.idr (100%) rename tests/idris2/{ => reg}/reg046/Postpone.idr (100%) rename tests/idris2/{ => reg}/reg046/expected (100%) rename tests/idris2/{ => reg}/reg046/input (100%) rename tests/idris2/{ => reg}/reg046/run (54%) rename tests/idris2/{ => reg}/reg047/QualifiedDoBang.idr (100%) rename tests/idris2/{ => reg}/reg047/expected (100%) rename tests/idris2/{ => reg}/reg047/input (100%) rename tests/idris2/{ => reg}/reg047/run (60%) rename tests/idris2/{ => reg}/reg048/expected (100%) rename tests/idris2/{ => reg}/reg048/inferror.idr (100%) rename tests/idris2/{ => reg}/reg048/run (55%) rename tests/idris2/{ => reg}/reg049/expected (100%) rename tests/idris2/{ => reg}/reg049/lettype.idr (100%) rename tests/idris2/{ => reg}/reg049/run (54%) rename tests/idris2/{ => reg}/reg050/expected (100%) rename tests/idris2/{ => reg}/reg050/loopy.idr (100%) rename tests/idris2/{ => reg}/reg050/run (52%) rename tests/idris2/{ => reg}/reg051/BigFins.idr (100%) rename tests/idris2/{ => reg}/reg051/expected (100%) create mode 100755 tests/idris2/reg/reg051/run rename tests/idris2/{ => reg}/reg051/test.ipkg (100%) rename tests/idris2/{ => reg}/reg052/DPairQuote.idr (100%) rename tests/idris2/{ => reg}/reg052/expected (100%) create mode 100755 tests/idris2/reg/reg052/run rename tests/idris2/{ => reg}/reg052/test.ipkg (100%) delete mode 100755 tests/idris2/reg001/run delete mode 100755 tests/idris2/reg002/run delete mode 100755 tests/idris2/reg004/run delete mode 100755 tests/idris2/reg005/run delete mode 100755 tests/idris2/reg006/run delete mode 100755 tests/idris2/reg007/run delete mode 100755 tests/idris2/reg009/run delete mode 100755 tests/idris2/reg010/run delete mode 100755 tests/idris2/reg011/run delete mode 100755 tests/idris2/reg012/run delete mode 100755 tests/idris2/reg015/run delete mode 100755 tests/idris2/reg016/run delete mode 100755 tests/idris2/reg017/run delete mode 100755 tests/idris2/reg018/run delete mode 100755 tests/idris2/reg019/run delete mode 100755 tests/idris2/reg021/run delete mode 100755 tests/idris2/reg022/run delete mode 100755 tests/idris2/reg023/run delete mode 100755 tests/idris2/reg026/run delete mode 100755 tests/idris2/reg027/run delete mode 100755 tests/idris2/reg028/run delete mode 100755 tests/idris2/reg029/run delete mode 100755 tests/idris2/reg030/run delete mode 100755 tests/idris2/reg031/run delete mode 100755 tests/idris2/reg032/run delete mode 100755 tests/idris2/reg033/run delete mode 100755 tests/idris2/reg034/run delete mode 100755 tests/idris2/reg035/run delete mode 100755 tests/idris2/reg036/run delete mode 100755 tests/idris2/reg037/run delete mode 100755 tests/idris2/reg039/run delete mode 100755 tests/idris2/reg040/run delete mode 100755 tests/idris2/reg041/run delete mode 100755 tests/idris2/reg044/run delete mode 100755 tests/idris2/reg045/run delete mode 100755 tests/idris2/reg051/run delete mode 100755 tests/idris2/reg052/run rename tests/idris2/{ => repl}/repl001/expected (100%) rename tests/idris2/{ => repl}/repl001/input (100%) create mode 100755 tests/idris2/repl/repl001/run rename tests/idris2/{ => repl}/repl002/expected (100%) rename tests/idris2/{ => repl}/repl002/input (100%) create mode 100755 tests/idris2/repl/repl002/run rename tests/idris2/{ => repl}/repl003/expected (100%) rename tests/idris2/{ => repl}/repl003/input (100%) create mode 100755 tests/idris2/repl/repl003/run rename tests/idris2/{ => repl}/repl004/expected (100%) rename tests/idris2/{ => repl}/repl004/input (100%) create mode 100755 tests/idris2/repl/repl004/run rename tests/idris2/{ => repl}/repl005/expected (100%) rename tests/idris2/{ => repl}/repl005/input (100%) create mode 100755 tests/idris2/repl/repl005/run rename tests/idris2/{ => repl}/repl006/expected (100%) rename tests/idris2/{ => repl}/repl006/input (100%) create mode 100755 tests/idris2/repl/repl006/run delete mode 100755 tests/idris2/repl001/run delete mode 100755 tests/idris2/repl002/run delete mode 100755 tests/idris2/repl003/run delete mode 100755 tests/idris2/repl004/run delete mode 100755 tests/idris2/repl005/run delete mode 100755 tests/idris2/repl006/run delete mode 100755 tests/idris2/rewrite001/run rename tests/idris2/{ => schemeeval}/schemeeval001/expected (100%) rename tests/idris2/{ => schemeeval}/schemeeval001/input (100%) create mode 100755 tests/idris2/schemeeval/schemeeval001/run rename tests/idris2/{ => schemeeval}/schemeeval002/expected (100%) rename tests/idris2/{ => schemeeval}/schemeeval002/input (100%) create mode 100755 tests/idris2/schemeeval/schemeeval002/run rename tests/idris2/{ => schemeeval}/schemeeval003/expected (100%) rename tests/idris2/{ => schemeeval}/schemeeval003/input (100%) create mode 100755 tests/idris2/schemeeval/schemeeval003/run rename tests/idris2/{ => schemeeval}/schemeeval004/expected (100%) rename tests/idris2/{ => schemeeval}/schemeeval004/input (100%) rename tests/idris2/{ => schemeeval}/schemeeval004/list.idr (100%) rename tests/idris2/{ => schemeeval}/schemeeval004/run (51%) rename tests/idris2/{ => schemeeval}/schemeeval005/Printf.idr (100%) rename tests/idris2/{ => schemeeval}/schemeeval005/expected (100%) rename tests/idris2/{ => schemeeval}/schemeeval005/input (100%) rename tests/idris2/{ => schemeeval}/schemeeval005/run (52%) rename tests/idris2/{ => schemeeval}/schemeeval006/expected (100%) rename tests/idris2/{ => schemeeval}/schemeeval006/input (100%) create mode 100644 tests/idris2/schemeeval/schemeeval006/run delete mode 100755 tests/idris2/schemeeval001/run delete mode 100755 tests/idris2/schemeeval002/run delete mode 100755 tests/idris2/schemeeval003/run delete mode 100644 tests/idris2/schemeeval006/run rename tests/idris2/{ => termination}/termination001/AgdaIssue6059.idr (100%) rename tests/idris2/{ => termination}/termination001/expected (100%) rename tests/idris2/{ => termination}/termination001/run (54%) rename tests/idris2/{ => total}/positivity001/Issue660.idr (100%) rename tests/idris2/{ => total}/positivity001/expected (100%) create mode 100644 tests/idris2/total/positivity001/run rename tests/idris2/{ => total}/positivity002/Issue660.idr (100%) rename tests/idris2/{ => total}/positivity002/expected (100%) create mode 100644 tests/idris2/total/positivity002/run rename tests/idris2/{ => total}/positivity003/Issue660.idr (100%) rename tests/idris2/{ => total}/positivity003/expected (100%) create mode 100644 tests/idris2/total/positivity003/run rename tests/idris2/{ => total}/positivity004/Issue1771-1.idr (100%) rename tests/idris2/{ => total}/positivity004/Issue1771-2.idr (100%) rename tests/idris2/{ => total}/positivity004/Issue1771-3.idr (100%) rename tests/idris2/{ => total}/positivity004/expected (100%) rename tests/idris2/{ => total}/positivity004/run (79%) rename tests/idris2/{ => total}/total001/Total.idr (100%) rename tests/idris2/{ => total}/total001/expected (100%) rename tests/idris2/{ => total}/total001/input (100%) rename tests/idris2/{ => total}/total001/run (52%) rename tests/idris2/{ => total}/total002/Total.idr (100%) rename tests/idris2/{ => total}/total002/expected (100%) rename tests/idris2/{ => total}/total002/input (100%) rename tests/idris2/{ => total}/total002/run (52%) rename tests/idris2/{ => total}/total003/Total.idr (100%) rename tests/idris2/{ => total}/total003/expected (100%) rename tests/idris2/{ => total}/total003/input (100%) rename tests/idris2/{ => total}/total003/run (52%) rename tests/idris2/{ => total}/total004/Total.idr (100%) rename tests/idris2/{ => total}/total004/expected (100%) rename tests/idris2/{ => total}/total004/input (100%) rename tests/idris2/{ => total}/total004/run (52%) rename tests/idris2/{ => total}/total005/Total.idr (100%) rename tests/idris2/{ => total}/total005/expected (100%) rename tests/idris2/{ => total}/total005/input (100%) create mode 100755 tests/idris2/total/total005/run rename tests/idris2/{ => total}/total006/Total.idr (100%) rename tests/idris2/{ => total}/total006/expected (100%) rename tests/idris2/{ => total}/total006/input (100%) create mode 100755 tests/idris2/total/total006/run rename tests/idris2/{ => total}/total007/expected (100%) rename tests/idris2/{ => total}/total007/partial.idr (100%) create mode 100755 tests/idris2/total/total007/run rename tests/idris2/{ => total}/total008/expected (100%) rename tests/idris2/{ => total}/total008/partial.idr (100%) create mode 100755 tests/idris2/total/total008/run rename tests/idris2/{ => total}/total009/expected (100%) create mode 100755 tests/idris2/total/total009/run rename tests/idris2/{ => total}/total009/tree.idr (100%) rename tests/idris2/{ => total}/total010/PartialWith.idr (100%) rename tests/idris2/{ => total}/total010/expected (100%) create mode 100755 tests/idris2/total/total010/run rename tests/idris2/{ => total}/total011/Issue1460.idr (100%) rename tests/idris2/{ => total}/total011/Issue1782.idr (100%) rename tests/idris2/{ => total}/total011/Issue1828.idr (100%) rename tests/idris2/{ => total}/total011/Issue1859-2.idr (100%) rename tests/idris2/{ => total}/total011/Issue1859.idr (100%) rename tests/idris2/{ => total}/total011/expected (100%) rename tests/idris2/{ => total}/total011/run (81%) rename tests/idris2/{ => total}/total012/Issue1828.idr (100%) rename tests/idris2/{ => total}/total012/TotallyTotal.idr (100%) rename tests/idris2/{ => total}/total012/expected (100%) rename tests/idris2/{ => total}/total012/run (94%) rename tests/idris2/{ => total}/total013/Issue1404.idr (100%) rename tests/idris2/{ => total}/total013/expected (100%) create mode 100755 tests/idris2/total/total013/run rename tests/idris2/{ => total}/total014/FunCompTC.idr (100%) rename tests/idris2/{ => total}/total014/expected (100%) create mode 100755 tests/idris2/total/total014/run rename tests/idris2/{ => total}/total015/CoveringData.idr (100%) rename tests/idris2/{ => total}/total015/expected (100%) rename tests/idris2/{ => total}/total015/run (50%) rename tests/idris2/{ => total}/total016/AssertPositivity.idr (100%) rename tests/idris2/{ => total}/total016/LazyPositivityCheck.idr (100%) rename tests/idris2/{ => total}/total016/expected (100%) rename tests/idris2/{ => total}/total016/run (70%) rename tests/idris2/{ => total}/total017/Paper.idr (100%) rename tests/idris2/{ => total}/total017/expected (100%) create mode 100755 tests/idris2/total/total017/run rename tests/idris2/{ => total}/total018/Issue2448.idr (100%) rename tests/idris2/{ => total}/total018/expected (100%) create mode 100755 tests/idris2/total/total018/run rename tests/idris2/{ => total}/total019/Check.idr (100%) rename tests/idris2/{ => total}/total019/expected (100%) create mode 100755 tests/idris2/total/total019/run rename tests/idris2/{ => total}/total020/Check.idr (100%) rename tests/idris2/{ => total}/total020/expected (100%) create mode 100755 tests/idris2/total/total020/run rename tests/idris2/{ => total}/total020/test.ipkg (100%) delete mode 100755 tests/idris2/total005/run delete mode 100755 tests/idris2/total006/run delete mode 100755 tests/idris2/total007/run delete mode 100755 tests/idris2/total008/run delete mode 100755 tests/idris2/total009/run delete mode 100755 tests/idris2/total010/run delete mode 100755 tests/idris2/total013/run delete mode 100755 tests/idris2/total014/run delete mode 100755 tests/idris2/total017/run delete mode 100755 tests/idris2/total018/run delete mode 100755 tests/idris2/total019/run delete mode 100755 tests/idris2/total020/run delete mode 100755 tests/idris2/unification001/run rename tests/idris2/{ => warning}/warning001/Holes.idr (100%) rename tests/idris2/{ => warning}/warning001/Issue1401.idr (100%) rename tests/idris2/{ => warning}/warning001/Issue539.idr (100%) rename tests/idris2/{ => warning}/warning001/Issue621.idr (100%) rename tests/idris2/{ => warning}/warning001/PR1407.idr (100%) rename tests/idris2/{ => warning}/warning001/expected (100%) rename tests/idris2/{ => warning}/warning001/run (79%) rename tests/idris2/{ => warning}/warning002/Foo.idr (100%) rename tests/idris2/{ => warning}/warning002/Main.idr (100%) rename tests/idris2/{ => warning}/warning002/deprecated.ipkg (100%) rename tests/idris2/{ => warning}/warning002/expected (100%) rename tests/idris2/{ => warning}/warning002/run (76%) rename tests/idris2/{ => warning}/warning003/Main.idr (100%) rename tests/idris2/{ => warning}/warning003/deprecated.ipkg (100%) rename tests/idris2/{ => warning}/warning003/expected (100%) create mode 100755 tests/idris2/warning/warning003/run rename tests/idris2/{ => warning}/warning004/Lib1.idr (100%) rename tests/idris2/{ => warning}/warning004/Lib2.idr (100%) rename tests/idris2/{ => warning}/warning004/Main1.idr (100%) rename tests/idris2/{ => warning}/warning004/expected (100%) rename tests/idris2/{ => warning}/warning004/run (57%) delete mode 100755 tests/idris2/warning003/run rename tests/idris2/{ => with}/with001/Temp.idr (100%) rename tests/idris2/{ => with}/with001/expected (100%) create mode 100755 tests/idris2/with/with001/run rename tests/idris2/{ => with}/with002/Temp.idr (100%) rename tests/idris2/{ => with}/with002/expected (100%) create mode 100755 tests/idris2/with/with002/run rename tests/idris2/{ => with}/with004/Issue637-2.idr (100%) rename tests/idris2/{ => with}/with004/Issue637-3.idr (100%) rename tests/idris2/{ => with}/with004/Issue637.idr (100%) rename tests/idris2/{ => with}/with004/expected (100%) rename tests/idris2/{ => with}/with004/input (100%) rename tests/idris2/{ => with}/with004/run (74%) rename tests/idris2/{ => with}/with005/Issue893.idr (100%) rename tests/idris2/{ => with}/with005/WithProof.idr (100%) rename tests/idris2/{ => with}/with005/expected (100%) rename tests/idris2/{ => with}/with005/run (62%) rename tests/idris2/{ => with}/with006/SparseWith.idr (100%) rename tests/idris2/{ => with}/with006/expected (100%) create mode 100755 tests/idris2/with/with006/run rename tests/idris2/{ => with}/with007/With0.idr (100%) rename tests/idris2/{ => with}/with007/expected (100%) create mode 100755 tests/idris2/with/with007/run rename tests/idris2/{ => with}/with008/WithClause.idr (100%) rename tests/idris2/{ => with}/with008/expected (100%) create mode 100755 tests/idris2/with/with008/run rename tests/idris2/{ => with}/with009/WithClause.idr (100%) rename tests/idris2/{ => with}/with009/expected (100%) create mode 100755 tests/idris2/with/with009/run rename tests/idris2/{ => with}/with010/NestedWith.idr (100%) rename tests/idris2/{ => with}/with010/expected (100%) create mode 100755 tests/idris2/with/with010/run rename tests/idris2/{ => with}/with011/WithImplicits.idr (100%) rename tests/idris2/{ => with}/with011/expected (100%) rename tests/idris2/{ => with}/with011/run (60%) delete mode 100755 tests/idris2/with001/run delete mode 100755 tests/idris2/with002/run delete mode 100755 tests/idris2/with006/run delete mode 100755 tests/idris2/with007/run delete mode 100755 tests/idris2/with008/run delete mode 100755 tests/idris2/with009/run delete mode 100755 tests/idris2/with010/run delete mode 100644 tests/node/node014/Echo.idr delete mode 100644 tests/node/node014/expected delete mode 100755 tests/node/node014/run diff --git a/.github/workflows/ci-idris2.yml b/.github/workflows/ci-idris2.yml index 08fccaeaae..9fdf02e50b 100644 --- a/.github/workflows/ci-idris2.yml +++ b/.github/workflows/ci-idris2.yml @@ -483,7 +483,7 @@ jobs: run: make install-api shell: bash - name: Test API - run: cd tests/idris2/api001 && ./run idris2 + run: cd tests/idris2/api/api001 && ./run "$HOME/.idris2/bin/idris2" shell: bash ###################################################################### diff --git a/tests/Main.idr b/tests/Main.idr index e30b721df2..e4b96c2ca5 100644 --- a/tests/Main.idr +++ b/tests/Main.idr @@ -11,206 +11,74 @@ import Test.Golden ------------------------------------------------------------------------ -- Test cases -ttimpTests : TestPool -ttimpTests = MkTestPool "TTImp" [] Nothing - [ "basic001", "basic002", "basic003", "basic004", "basic005" - , "basic006" - , "coverage002" - , "dot001" - , "eta001" - , "lazy001" - , "nest001", "nest002" - , "perf001", "perf002", "perf003" - , "record001", "record002", "record003", "record004" - , "qtt001", "qtt003" - , "total001", "total002", "total003" - ] - -idrisTestsBasic : TestPool -idrisTestsBasic = MkTestPool "Fundamental language features" [] Nothing - -- Fundamental language features - ["basic001", "basic002", "basic003", "basic004", "basic005", - "basic006", "basic007", "basic008", "basic009", "basic010", - "basic011", "basic012", "basic013", "basic014", "basic015", - "basic016", "basic017", "basic018", "basic019", "basic020", - "basic021", "basic022", "basic023", "basic024", "basic025", - "basic026", "basic027", "basic028", "basic029", "basic030", - "basic031", "basic032", "basic033", "basic034", "basic035", - "basic036", "basic037", "basic038", "basic039", "basic040", - "basic041", "basic042", "basic043", "basic044", "basic045", - "basic046", "basic047", "basic049", "basic050", - "basic051", "basic052", "basic053", "basic054", "basic055", - "basic056", "basic057", "basic058", "basic059", "basic060", - "basic061", "basic062", "basic063", "basic064", "basic065", - "basic066", "basic067", "basic068", "basic069", "basic070", - "basic071", - "idiom001", - "literals001", - "dotted001", - "case001", - "rewrite001", - "interpolation001", "interpolation002", "interpolation003", - "interpolation004"] - -idrisTestsDebug : TestPool -idrisTestsDebug = MkTestPool "Debug features" [] Nothing - ["debug001"] - -idrisTestsCoverage : TestPool -idrisTestsCoverage = MkTestPool "Coverage checking" [] Nothing - -- Coverage checking - ["coverage001", "coverage002", "coverage003", "coverage004", - "coverage005", "coverage006", "coverage007", "coverage008", - "coverage009", "coverage010", "coverage011", "coverage012", - "coverage013", "coverage014", "coverage015", "coverage016", - "coverage017", "coverage018", "coverage019"] - -idrisTestsTermination : TestPool -idrisTestsTermination = MkTestPool "Termination checking" [] Nothing - -- Termination checking - ["termination001"] - -idrisTestsCasetree : TestPool -idrisTestsCasetree = MkTestPool "Case tree building" [] Nothing - -- Case tree building - ["casetree001", "casetree002", "casetree003", "casetree004"] - -idrisTestsWarning : TestPool -idrisTestsWarning = MkTestPool "Warnings" [] Nothing - ["warning001", "warning002", "warning003", "warning004"] - -idrisTestsFailing : TestPool -idrisTestsFailing = MkTestPool "Failing blocks" [] Nothing - ["failing001", "failing002", "failing003", "failing004" - ] - -idrisTestsError : TestPool -idrisTestsError = MkTestPool "Error messages" [] Nothing - -- Error messages - ["error001", "error002", "error003", "error004", "error005", - "error006", "error007", "error008", "error009", "error010", - "error011", "error012", "error013", "error014", "error015", - "error016", "error017", "error018", "error019", "error020", - "error021", "error022", "error023", "error024", "error025", - "error026", "error027", - -- Parse errors - "perror001", "perror002", "perror003", "perror004", "perror005", - "perror006", "perror007", "perror008", "perror009", "perror010", - "perror011", "perror012", "perror013", "perror014", "perror015", - "perror016", "perror017", "perror018", "perror019", "perror020", - "perror021", "perror022", "perror023", "perror024", "perror025", - "perror026", "perror027", "perror028", "perror029"] - -idrisTestsInteractive : TestPool -idrisTestsInteractive = MkTestPool "Interactive editing" [] Nothing - -- Interactive editing support - ["interactive001", "interactive002", "interactive003", "interactive004", - "interactive005", "interactive006", "interactive007", "interactive008", - "interactive009", "interactive010", "interactive011", "interactive012", - "interactive013", "interactive014", "interactive015", "interactive016", - "interactive017", "interactive018", "interactive019", "interactive020", - "interactive021", "interactive022", "interactive023", "interactive024", - "interactive025", "interactive026", "interactive027", "interactive028", - "interactive029", "interactive030", "interactive031", "interactive032", - "interactive033", "interactive034", "interactive035", "interactive036", - "interactive037", "interactive038", "interactive039", "interactive040", - "interactive041", "interactive042", "interactive043", "interactive044", - "interactive045", "interactive046" - ] +ttimpTests : IO TestPool +ttimpTests = testsInDir "ttimp" (const True) "TTImp" [] Nothing + +idrisTestsBasic : IO TestPool +idrisTestsBasic = testsInDir "idris2/basic" (const True) "Fundamental language features" [] Nothing + +idrisTestsDebug : IO TestPool +idrisTestsDebug = testsInDir "idris2/debug" (const True) "Debug features" [] Nothing + +idrisTestsCoverage : IO TestPool +idrisTestsCoverage = testsInDir "idris2/coverage" (const True) "Coverage checking" [] Nothing + +idrisTestsTermination : IO TestPool +idrisTestsTermination = testsInDir "idris2/termination" (const True) "Termination checking" [] Nothing + +idrisTestsCasetree : IO TestPool +idrisTestsCasetree = testsInDir "idris2/casetree" (const True) "Case tree building" [] Nothing + +idrisTestsWarning : IO TestPool +idrisTestsWarning = testsInDir "idris2/warning" (const True) "Warnings" [] Nothing + +idrisTestsFailing : IO TestPool +idrisTestsFailing = testsInDir "idris2/failing" (const True) "Failing blocks" [] Nothing + +||| Error messages, including parse errors ("perror") +idrisTestsError : IO TestPool +idrisTestsError = testsInDir "idris2/error" (const True) "Error messages" [] Nothing + +idrisTestsInteractive : IO TestPool +idrisTestsInteractive = testsInDir "idris2/interactive" (const True) "Interactive editing" [] Nothing + +idrisTestsInterface : IO TestPool +idrisTestsInterface = testsInDir "idris2/interface" (const True) "Interface" [] Nothing + +||| QTT and linearity related +idrisTestsLinear : IO TestPool +idrisTestsLinear = testsInDir "idris2/linear" (const True) "Quantities" [] Nothing -idrisTestsInterface : TestPool -idrisTestsInterface = MkTestPool "Interface" [] Nothing - -- Interfaces - ["interface001", "interface002", "interface003", "interface004", - "interface005", "interface006", "interface007", "interface008", - "interface009", "interface010", "interface011", "interface012", - "interface013", "interface014", "interface015", "interface016", - "interface017", "interface018", "interface019", "interface020", - "interface021", "interface022", "interface023", "interface024", - "interface025", "interface026", "interface027", "interface028", - "interface029"] - -idrisTestsLinear : TestPool -idrisTestsLinear = MkTestPool "Quantities" [] Nothing - -- QTT and linearity related - ["linear001", "linear002", "linear003", -- "linear004" -- disabled due to requiring linearity subtyping - "linear005", "linear006", "linear007", "linear008", - "linear009", "linear010", "linear011", "linear012", - "linear013", "linear014", "linear015", "linear016", - "linear017"] - -idrisTestsLiterate : TestPool -idrisTestsLiterate = MkTestPool "Literate programming" [] Nothing - -- Literate - ["literate001", "literate002", "literate003", "literate004", - "literate005", "literate006", "literate007", "literate008", - "literate009", "literate010", "literate011", "literate012", - "literate013", "literate014", "literate015", "literate016", - "literate017"] - -idrisTestsPerformance : TestPool -idrisTestsPerformance = MkTestPool "Performance" [] Nothing - -- Performance: things which have been slow in the past, or which - -- pose interesting challenges for the elaborator - ["perf001", "perf002", "perf003", "perf004", "perf005", - "perf007", "perf008", "perf009", "perf010", "perf011", - "perf012", - "perf2202"] - -idrisTestsRegression : TestPool -idrisTestsRegression = MkTestPool "Various regressions" [] Nothing - -- Miscellaneous regressions - ["reg001", "reg002", "reg003", "reg004", "reg005", "reg006", "reg007", - "reg008", "reg009", "reg010", "reg011", "reg012", "reg013", "reg014", - "reg015", "reg016", "reg017", "reg018", "reg019", "reg020", "reg021", - "reg022", "reg023", "reg024", "reg025", "reg026", "reg027", "reg028", - "reg029", "reg030", "reg031", "reg032", "reg033", "reg034", "reg035", - "reg036", "reg037", "reg038", "reg039", "reg040", "reg041", "reg042", - "reg043", "reg044", "reg045", "reg046", "reg047", "reg048", "reg049", - "reg050", "reg051", "reg052"] - -idrisTestsData : TestPool -idrisTestsData = MkTestPool "Data and record types" [] Nothing - [-- Data types - "data001", "data002", - -- Records, access and dependent update - "record001", "record002", "record003", "record004", "record005", - "record006", "record007", "record008", "record009", "record010", - "record011", "record012", "record013", "record014", "record015", - "record016", "record017", "record018", "record019" ] - -idrisTestsBuiltin : TestPool -idrisTestsBuiltin = MkTestPool "Builtin types and functions" [] Nothing - -- %builtin related tests for the frontend (type-checking) - ["builtin001", "builtin002", "builtin003", "builtin004", "builtin005", - "builtin006", "builtin007", "builtin008", "builtin009", "builtin010", - "builtin011"] - -idrisTestsEvaluator : TestPool -idrisTestsEvaluator = MkTestPool "Evaluation" [] Nothing - [ -- Evaluator - "evaluator001", "evaluator002", "evaluator003", "evaluator004", - -- Miscellaneous REPL - "interpreter001", "interpreter002", "interpreter003", "interpreter004", - "interpreter005", "interpreter006", "interpreter007", "interpreter008", - -- Specialisation - "spec001" - ] - -idrisTestsREPL : TestPool -idrisTestsREPL = MkTestPool "REPL commands and help" [] Nothing - [ "repl001", "repl002", "repl003", "repl004", "repl005" - , "repl006" - ] - -idrisTestsAllSchemes : Requirement -> TestPool -idrisTestsAllSchemes cg = MkTestPool +idrisTestsLiterate : IO TestPool +idrisTestsLiterate = testsInDir "idris2/literate" (const True) "Literate programming" [] Nothing + +||| Performance: things which have been slow in the past, or which +||| pose interesting challenges for the elaborator +idrisTestsPerformance : IO TestPool +idrisTestsPerformance = testsInDir "idris2/perf" (const True) "Performance" [] Nothing + +idrisTestsRegression : IO TestPool +idrisTestsRegression = testsInDir "idris2/reg" (const True) "Various regressions" [] Nothing + +||| Data types, including records +idrisTestsData : IO TestPool +idrisTestsData = testsInDir "idris2/data" (const True) "Data and record types" [] Nothing + +||| %builtin related tests for the frontend (type-checking) +idrisTestsBuiltin : IO TestPool +idrisTestsBuiltin = testsInDir "idris2/builtin" (const True) "Builtin types and functions" [] Nothing + +||| Evaluator, REPL, specialisation +idrisTestsEvaluator : IO TestPool +idrisTestsEvaluator = testsInDir "idris2/evaluator" (const True) "Evaluation" [] Nothing + +idrisTestsREPL : IO TestPool +idrisTestsREPL = testsInDir "idris2/repl" (const True) "REPL commands and help" [] Nothing + +idrisTestsAllSchemes : Requirement -> IO TestPool +idrisTestsAllSchemes cg = testsInDir "allschemes" (const True) ("Test across all scheme backends: " ++ show cg ++ " instance") [] (Just cg) - [ "scheme001", "scheme002" - , "channels001", "channels002", "channels003", "channels004", "channels005" - , "channels006" - ] idrisTestsAllBackends : Requirement -> TestPool idrisTestsAllBackends cg = MkTestPool @@ -228,45 +96,24 @@ idrisTestsAllBackends cg = MkTestPool "basic048", "perf006"] -idrisTestsTotality : TestPool -idrisTestsTotality = MkTestPool "Totality checking" [] Nothing - -- Positivity checking - ["positivity001", "positivity002", "positivity003", "positivity004", - -- Totality checking - "total001", "total002", "total003", "total004", "total005", - "total006", "total007", "total008", "total009", "total010", - "total011", "total012", "total013", "total014", "total015", - "total016", "total017", "total018", "total019", "total020" - ] +||| Totality checking, including positivity +idrisTestsTotality : IO TestPool +idrisTestsTotality = testsInDir "idris2/total" (const True) "Totality checking" [] Nothing -- This will only work with an Idris compiled via Chez or Racket, but at -- least for the moment we're not officially supporting self hosting any -- other way. If we do, we'll need to have a way to disable these. -idrisTestsSchemeEval : TestPool -idrisTestsSchemeEval = MkTestPool "Scheme Evaluator" [] Nothing - ["schemeeval001", "schemeeval002", "schemeeval003", "schemeeval004", - "schemeeval005", "schemeeval006"] - -idrisTestsReflection : TestPool -idrisTestsReflection = MkTestPool "Quotation and Reflection" [] Nothing - ["reflection001", "reflection002", "reflection003", "reflection004", - "reflection005", "reflection006", "reflection007", "reflection008", - "reflection009", "reflection010", "reflection011", "reflection012", - "reflection013", "reflection014", "reflection015", "reflection016", - "reflection017", "reflection018", "reflection019", "reflection020", - "reflection021"] - -idrisTestsWith : TestPool -idrisTestsWith = MkTestPool "With abstraction" [] Nothing - [ "with001", "with002", "with004", "with005", "with006", "with007", - "with008", "with009", "with010", "with011" - ] - -idrisTestsIPKG : TestPool -idrisTestsIPKG = MkTestPool "Package and .ipkg files" [] Nothing - ["pkg001", "pkg002", "pkg003", "pkg004", "pkg005", "pkg006", "pkg007", - "pkg008", "pkg009", "pkg010", "pkg011", "pkg012", "pkg013", "pkg014", - "pkg015", "pkg016", "pkg017" ] +idrisTestsSchemeEval : IO TestPool +idrisTestsSchemeEval = testsInDir "idris2/schemeeval" (const True) "Scheme Evaluator" [] Nothing + +idrisTestsReflection : IO TestPool +idrisTestsReflection = testsInDir "idris2/reflection" (const True) "Quotation and Reflection" [] Nothing + +idrisTestsWith : IO TestPool +idrisTestsWith = testsInDir "idris2/with" (const True) "With abstraction" [] Nothing + +idrisTestsIPKG : IO TestPool +idrisTestsIPKG = testsInDir "idris2/pkg" (const True) "Package and .ipkg files" [] Nothing idrisTests : TestPool idrisTests = MkTestPool "Misc" [] Nothing @@ -305,73 +152,17 @@ idrisTests = MkTestPool "Misc" [] Nothing typeddTests : IO TestPool typeddTests = testsInDir "typedd-book" (const True) "Type Driven Development" [] Nothing -chezTests : TestPool -chezTests = MkTestPool "Chez backend" [] (Just Chez) - [ "chez001", "chez002", "chez003", "chez004", "chez005", "chez006" - , "chez007", "chez008", "chez009", "chez010", "chez011", "chez012" - , "chez013", "chez014", "chez015", "chez016", "chez017", "chez018" - , "chez019", "chez020", "chez021", "chez022", "chez023", "chez024" - , "chez025", "chez026", "chez027", "chez028", "chez029", "chez030" - , "chez031", "chez032", "chez033", "chez034", "chez035", "chez036" - , "futures001" - , "bitops" - , "casts" - , "constfold", "constfold2", "constfold3" - , "memo" - , "newints" - , "integers" - , "nat2fin" - , "inlineiobind" - , "semaphores001" - , "semaphores002" - , "perf001" - , "reg001" - , "buffer001" - ] +chezTests : IO TestPool +chezTests = testsInDir "chez" (const True) "Chez backend" [] (Just Chez) refcTests : IO TestPool refcTests = testsInDir "refc" (const True) "Reference counting C backend" [] (Just C) -racketTests : TestPool -racketTests = MkTestPool "Racket backend" [] (Just Racket) - [ "forkjoin001" - , "semaphores001", "semaphores002" - , "futures001" - , "mutex001", "mutex002", "mutex003", "mutex004", "mutex005" - , "conditions001" , "conditions002" , "conditions003" , "conditions004" - , "conditions005" --- , "conditions006" --- , "conditions007" - , "ffi001" - ] - -nodeTests : TestPool -nodeTests = MkTestPool "Node backend" [] (Just Node) - [ "node001", "node002", "node003", "node004", "node005", "node006" - , "node007", "node008", "node009", "node011", "node012", "node015" - , "node017", "node018", "node019", "node020", "node021", "node022" - , "node023", "node024", "node025", "node026", "node027" - , "perf001" - -- , "node14" - , "args" - , "bitops" - , "casts" - , "memo" - , "fastConcat" - , "newints" - , "reg001", "reg002" - , "stringcast" - , "syntax001" - , "tailrec001" - , "tailrec002" - , "idiom001" - , "integers" - , "doubles" - , "fix1839" - , "tailrec_libs" - , "nomangle001", "nomangle002" - , "integer_array" - ] +racketTests : IO TestPool +racketTests = testsInDir "racket" (const True) "Racket backend" [] (Just Racket) + +nodeTests : IO TestPool +nodeTests = testsInDir "node" (const True) "Node backend" [] (Just Node) vmcodeInterpTests : IO TestPool vmcodeInterpTests = testsInDir "vmcode" (const True) "VMCode interpreter" [] Nothing @@ -402,45 +193,45 @@ codegenTests = testsInDir "codegen" (const True) "Code generation" [] Nothing main : IO () main = runner $ - [ testPaths "ttimp" ttimpTests - , testPaths "idris2" idrisTestsBasic - , testPaths "idris2" idrisTestsCoverage - , testPaths "idris2" idrisTestsTermination - , testPaths "idris2" idrisTestsCasetree - , testPaths "idris2" idrisTestsError - , testPaths "idris2" idrisTestsFailing - , testPaths "idris2" idrisTestsWarning - , testPaths "idris2" idrisTestsInteractive - , testPaths "idris2" idrisTestsInterface - , testPaths "idris2" idrisTestsLiterate - , testPaths "idris2" idrisTestsLinear - , testPaths "idris2" idrisTestsPerformance - , testPaths "idris2" idrisTestsRegression - , testPaths "idris2" idrisTestsData - , testPaths "idris2" idrisTestsBuiltin - , testPaths "idris2" idrisTestsEvaluator - , testPaths "idris2" idrisTestsREPL - , testPaths "idris2" idrisTestsTotality - , testPaths "idris2" idrisTestsSchemeEval - , testPaths "idris2" idrisTestsReflection - , testPaths "idris2" idrisTestsWith - , testPaths "idris2" idrisTestsDebug - , testPaths "idris2" idrisTestsIPKG - , testPaths "idris2" idrisTests + [ !ttimpTests + , !idrisTestsBasic + , !idrisTestsCoverage + , !idrisTestsTermination + , !idrisTestsCasetree + , !idrisTestsError + , !idrisTestsFailing + , !idrisTestsWarning + , !idrisTestsInteractive + , !idrisTestsInterface + , !idrisTestsLiterate + , !idrisTestsLinear + , !idrisTestsPerformance + , !idrisTestsRegression + , !idrisTestsData + , !idrisTestsBuiltin + , !idrisTestsEvaluator + , !idrisTestsREPL + , !idrisTestsTotality + , !idrisTestsSchemeEval + , !idrisTestsReflection + , !idrisTestsWith + , !idrisTestsDebug + , !idrisTestsIPKG + , testPaths "idris2/misc" idrisTests , !typeddTests , !ideModeTests , !preludeTests , !baseLibraryTests , !contribLibraryTests - , testPaths "chez" chezTests + , !chezTests , !refcTests - , testPaths "racket" racketTests - , testPaths "node" nodeTests + , !racketTests + , !nodeTests , !vmcodeInterpTests , !templateTests , !codegenTests ] - ++ map (testPaths "allschemes" . idrisTestsAllSchemes) [Chez, Racket] + ++ !(traverse idrisTestsAllSchemes [Chez, Racket]) ++ map (testPaths "allbackends" . idrisTestsAllBackends) [Chez, Node, Racket, C] diff --git a/tests/chez/barrier001/Main.idr b/tests/chez/barrier001/Main.idr index add64d4997..d2b27330e5 100644 --- a/tests/chez/barrier001/Main.idr +++ b/tests/chez/barrier001/Main.idr @@ -10,6 +10,6 @@ main = do putStrLn "Hello" barrierWait barrier putStrLn "Goodbye" - for threadIDs $ \threadID => + for_ threadIDs $ \threadID => threadWait threadID sleep 1 diff --git a/tests/idris2/api001/README b/tests/idris2/api/README similarity index 100% rename from tests/idris2/api001/README rename to tests/idris2/api/README diff --git a/tests/idris2/api001/Hello.idr b/tests/idris2/api/api001/Hello.idr similarity index 100% rename from tests/idris2/api001/Hello.idr rename to tests/idris2/api/api001/Hello.idr diff --git a/tests/idris2/api001/LazyCodegen.idr b/tests/idris2/api/api001/LazyCodegen.idr similarity index 100% rename from tests/idris2/api001/LazyCodegen.idr rename to tests/idris2/api/api001/LazyCodegen.idr diff --git a/tests/idris2/api001/expected b/tests/idris2/api/api001/expected similarity index 100% rename from tests/idris2/api001/expected rename to tests/idris2/api/api001/expected diff --git a/tests/idris2/api001/run b/tests/idris2/api/api001/run similarity index 88% rename from tests/idris2/api001/run rename to tests/idris2/api/api001/run index 22a6d70e80..84b3afbfcb 100755 --- a/tests/idris2/api001/run +++ b/tests/idris2/api/api001/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 -p idris2 -p contrib -p network LazyCodegen.idr -o lazy-idris2 ./build/exec/lazy-idris2 --cg lazy Hello.idr -o hello > output diff --git a/tests/idris2/basic001/Vect.idr b/tests/idris2/basic/basic001/Vect.idr similarity index 100% rename from tests/idris2/basic001/Vect.idr rename to tests/idris2/basic/basic001/Vect.idr diff --git a/tests/idris2/basic001/expected b/tests/idris2/basic/basic001/expected similarity index 100% rename from tests/idris2/basic001/expected rename to tests/idris2/basic/basic001/expected diff --git a/tests/idris2/basic001/input b/tests/idris2/basic/basic001/input similarity index 100% rename from tests/idris2/basic001/input rename to tests/idris2/basic/basic001/input diff --git a/tests/idris2/basic001/run b/tests/idris2/basic/basic001/run similarity index 61% rename from tests/idris2/basic001/run rename to tests/idris2/basic/basic001/run index a469d81530..38bfd43022 100755 --- a/tests/idris2/basic001/run +++ b/tests/idris2/basic/basic001/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 --no-prelude Vect.idr < input diff --git a/tests/idris2/basic002/Do.idr b/tests/idris2/basic/basic002/Do.idr similarity index 100% rename from tests/idris2/basic002/Do.idr rename to tests/idris2/basic/basic002/Do.idr diff --git a/tests/idris2/basic002/expected b/tests/idris2/basic/basic002/expected similarity index 100% rename from tests/idris2/basic002/expected rename to tests/idris2/basic/basic002/expected diff --git a/tests/idris2/basic002/input b/tests/idris2/basic/basic002/input similarity index 100% rename from tests/idris2/basic002/input rename to tests/idris2/basic/basic002/input diff --git a/tests/idris2/basic002/run b/tests/idris2/basic/basic002/run similarity index 60% rename from tests/idris2/basic002/run rename to tests/idris2/basic/basic002/run index 95fcc62fa0..091eb1d62d 100755 --- a/tests/idris2/basic002/run +++ b/tests/idris2/basic/basic002/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 --no-prelude Do.idr < input diff --git a/tests/idris2/basic003/Ambig1.idr b/tests/idris2/basic/basic003/Ambig1.idr similarity index 100% rename from tests/idris2/basic003/Ambig1.idr rename to tests/idris2/basic/basic003/Ambig1.idr diff --git a/tests/idris2/basic003/Ambig2.idr b/tests/idris2/basic/basic003/Ambig2.idr similarity index 100% rename from tests/idris2/basic003/Ambig2.idr rename to tests/idris2/basic/basic003/Ambig2.idr diff --git a/tests/idris2/basic003/expected b/tests/idris2/basic/basic003/expected similarity index 100% rename from tests/idris2/basic003/expected rename to tests/idris2/basic/basic003/expected diff --git a/tests/idris2/basic003/run b/tests/idris2/basic/basic003/run similarity index 78% rename from tests/idris2/basic003/run rename to tests/idris2/basic/basic003/run index e525674551..2fcbacb30d 100755 --- a/tests/idris2/basic003/run +++ b/tests/idris2/basic/basic003/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh echo ':q' | idris2 --no-prelude Ambig1.idr echo ':q' | idris2 --no-prelude Ambig2.idr diff --git a/tests/idris2/basic004/Stuff.idr b/tests/idris2/basic/basic004/Stuff.idr similarity index 100% rename from tests/idris2/basic004/Stuff.idr rename to tests/idris2/basic/basic004/Stuff.idr diff --git a/tests/idris2/basic004/Wheres.idr b/tests/idris2/basic/basic004/Wheres.idr similarity index 100% rename from tests/idris2/basic004/Wheres.idr rename to tests/idris2/basic/basic004/Wheres.idr diff --git a/tests/idris2/basic004/expected b/tests/idris2/basic/basic004/expected similarity index 100% rename from tests/idris2/basic004/expected rename to tests/idris2/basic/basic004/expected diff --git a/tests/idris2/basic004/input b/tests/idris2/basic/basic004/input similarity index 100% rename from tests/idris2/basic004/input rename to tests/idris2/basic/basic004/input diff --git a/tests/idris2/basic004/run b/tests/idris2/basic/basic004/run similarity index 62% rename from tests/idris2/basic004/run rename to tests/idris2/basic/basic004/run index 48fbd46850..3127365f3d 100755 --- a/tests/idris2/basic004/run +++ b/tests/idris2/basic/basic004/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 --no-prelude Wheres.idr < input diff --git a/tests/idris2/basic005/NoInfer.idr b/tests/idris2/basic/basic005/NoInfer.idr similarity index 100% rename from tests/idris2/basic005/NoInfer.idr rename to tests/idris2/basic/basic005/NoInfer.idr diff --git a/tests/idris2/basic005/expected b/tests/idris2/basic/basic005/expected similarity index 100% rename from tests/idris2/basic005/expected rename to tests/idris2/basic/basic005/expected diff --git a/tests/idris2/basic005/run b/tests/idris2/basic/basic005/run similarity index 65% rename from tests/idris2/basic005/run rename to tests/idris2/basic/basic005/run index 6b6218fd5b..152a7a7404 100755 --- a/tests/idris2/basic005/run +++ b/tests/idris2/basic/basic005/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh echo ':q' | idris2 --no-prelude NoInfer.idr diff --git a/tests/idris2/basic006/PMLet.idr b/tests/idris2/basic/basic006/PMLet.idr similarity index 100% rename from tests/idris2/basic006/PMLet.idr rename to tests/idris2/basic/basic006/PMLet.idr diff --git a/tests/idris2/basic006/Stuff.idr b/tests/idris2/basic/basic006/Stuff.idr similarity index 100% rename from tests/idris2/basic006/Stuff.idr rename to tests/idris2/basic/basic006/Stuff.idr diff --git a/tests/idris2/basic006/expected b/tests/idris2/basic/basic006/expected similarity index 100% rename from tests/idris2/basic006/expected rename to tests/idris2/basic/basic006/expected diff --git a/tests/idris2/basic006/input b/tests/idris2/basic/basic006/input similarity index 100% rename from tests/idris2/basic006/input rename to tests/idris2/basic/basic006/input diff --git a/tests/idris2/basic006/run b/tests/idris2/basic/basic006/run similarity index 61% rename from tests/idris2/basic006/run rename to tests/idris2/basic/basic006/run index 45af0f84c4..1dd142efe8 100755 --- a/tests/idris2/basic006/run +++ b/tests/idris2/basic/basic006/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 --no-prelude PMLet.idr < input diff --git a/tests/idris2/basic007/DoLocal.idr b/tests/idris2/basic/basic007/DoLocal.idr similarity index 100% rename from tests/idris2/basic007/DoLocal.idr rename to tests/idris2/basic/basic007/DoLocal.idr diff --git a/tests/idris2/basic007/Stuff.idr b/tests/idris2/basic/basic007/Stuff.idr similarity index 100% rename from tests/idris2/basic007/Stuff.idr rename to tests/idris2/basic/basic007/Stuff.idr diff --git a/tests/idris2/basic007/expected b/tests/idris2/basic/basic007/expected similarity index 100% rename from tests/idris2/basic007/expected rename to tests/idris2/basic/basic007/expected diff --git a/tests/idris2/basic007/input b/tests/idris2/basic/basic007/input similarity index 100% rename from tests/idris2/basic007/input rename to tests/idris2/basic/basic007/input diff --git a/tests/idris2/basic007/run b/tests/idris2/basic/basic007/run similarity index 63% rename from tests/idris2/basic007/run rename to tests/idris2/basic/basic007/run index b1a0170ead..6803f81c69 100755 --- a/tests/idris2/basic007/run +++ b/tests/idris2/basic/basic007/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 --no-prelude DoLocal.idr < input diff --git a/tests/idris2/basic008/If.idr b/tests/idris2/basic/basic008/If.idr similarity index 100% rename from tests/idris2/basic008/If.idr rename to tests/idris2/basic/basic008/If.idr diff --git a/tests/idris2/basic008/Stuff.idr b/tests/idris2/basic/basic008/Stuff.idr similarity index 100% rename from tests/idris2/basic008/Stuff.idr rename to tests/idris2/basic/basic008/Stuff.idr diff --git a/tests/idris2/basic008/expected b/tests/idris2/basic/basic008/expected similarity index 100% rename from tests/idris2/basic008/expected rename to tests/idris2/basic/basic008/expected diff --git a/tests/idris2/basic008/input b/tests/idris2/basic/basic008/input similarity index 100% rename from tests/idris2/basic008/input rename to tests/idris2/basic/basic008/input diff --git a/tests/idris2/basic008/run b/tests/idris2/basic/basic008/run similarity index 60% rename from tests/idris2/basic008/run rename to tests/idris2/basic/basic008/run index 82f762e699..b565cc6b77 100755 --- a/tests/idris2/basic008/run +++ b/tests/idris2/basic/basic008/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 --no-prelude If.idr < input diff --git a/tests/idris2/basic009/LetCase.idr b/tests/idris2/basic/basic009/LetCase.idr similarity index 100% rename from tests/idris2/basic009/LetCase.idr rename to tests/idris2/basic/basic009/LetCase.idr diff --git a/tests/idris2/basic009/Stuff.idr b/tests/idris2/basic/basic009/Stuff.idr similarity index 100% rename from tests/idris2/basic009/Stuff.idr rename to tests/idris2/basic/basic009/Stuff.idr diff --git a/tests/idris2/basic009/expected b/tests/idris2/basic/basic009/expected similarity index 100% rename from tests/idris2/basic009/expected rename to tests/idris2/basic/basic009/expected diff --git a/tests/idris2/basic009/input b/tests/idris2/basic/basic009/input similarity index 100% rename from tests/idris2/basic009/input rename to tests/idris2/basic/basic009/input diff --git a/tests/idris2/basic009/run b/tests/idris2/basic/basic009/run similarity index 63% rename from tests/idris2/basic009/run rename to tests/idris2/basic/basic009/run index 380ff1625d..23fd17f0e1 100755 --- a/tests/idris2/basic009/run +++ b/tests/idris2/basic/basic009/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 --no-prelude LetCase.idr < input diff --git a/tests/idris2/basic010/Comp.idr b/tests/idris2/basic/basic010/Comp.idr similarity index 100% rename from tests/idris2/basic010/Comp.idr rename to tests/idris2/basic/basic010/Comp.idr diff --git a/tests/idris2/basic010/expected b/tests/idris2/basic/basic010/expected similarity index 100% rename from tests/idris2/basic010/expected rename to tests/idris2/basic/basic010/expected diff --git a/tests/idris2/basic010/input b/tests/idris2/basic/basic010/input similarity index 100% rename from tests/idris2/basic010/input rename to tests/idris2/basic/basic010/input diff --git a/tests/idris2/basic010/run b/tests/idris2/basic/basic010/run similarity index 61% rename from tests/idris2/basic010/run rename to tests/idris2/basic/basic010/run index be231ce846..c5da56cd07 100755 --- a/tests/idris2/basic010/run +++ b/tests/idris2/basic/basic010/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 --no-prelude Comp.idr < input diff --git a/tests/idris2/basic011/Dots1.idr b/tests/idris2/basic/basic011/Dots1.idr similarity index 100% rename from tests/idris2/basic011/Dots1.idr rename to tests/idris2/basic/basic011/Dots1.idr diff --git a/tests/idris2/basic011/Dots2.idr b/tests/idris2/basic/basic011/Dots2.idr similarity index 100% rename from tests/idris2/basic011/Dots2.idr rename to tests/idris2/basic/basic011/Dots2.idr diff --git a/tests/idris2/basic011/Dots3.idr b/tests/idris2/basic/basic011/Dots3.idr similarity index 100% rename from tests/idris2/basic011/Dots3.idr rename to tests/idris2/basic/basic011/Dots3.idr diff --git a/tests/idris2/basic011/expected b/tests/idris2/basic/basic011/expected similarity index 100% rename from tests/idris2/basic011/expected rename to tests/idris2/basic/basic011/expected diff --git a/tests/idris2/basic011/run b/tests/idris2/basic/basic011/run similarity index 67% rename from tests/idris2/basic011/run rename to tests/idris2/basic/basic011/run index fd4d43813b..5e405be927 100755 --- a/tests/idris2/basic011/run +++ b/tests/idris2/basic/basic011/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh check Dots1.idr check Dots2.idr diff --git a/tests/idris2/basic012/VIndex.idr b/tests/idris2/basic/basic012/VIndex.idr similarity index 100% rename from tests/idris2/basic012/VIndex.idr rename to tests/idris2/basic/basic012/VIndex.idr diff --git a/tests/idris2/basic012/expected b/tests/idris2/basic/basic012/expected similarity index 100% rename from tests/idris2/basic012/expected rename to tests/idris2/basic/basic012/expected diff --git a/tests/idris2/basic/basic012/run b/tests/idris2/basic/basic012/run new file mode 100755 index 0000000000..7665442bea --- /dev/null +++ b/tests/idris2/basic/basic012/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check VIndex.idr diff --git a/tests/idris2/basic013/Implicits.idr b/tests/idris2/basic/basic013/Implicits.idr similarity index 100% rename from tests/idris2/basic013/Implicits.idr rename to tests/idris2/basic/basic013/Implicits.idr diff --git a/tests/idris2/basic013/expected b/tests/idris2/basic/basic013/expected similarity index 100% rename from tests/idris2/basic013/expected rename to tests/idris2/basic/basic013/expected diff --git a/tests/idris2/basic/basic013/run b/tests/idris2/basic/basic013/run new file mode 100755 index 0000000000..244390a1c8 --- /dev/null +++ b/tests/idris2/basic/basic013/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Implicits.idr diff --git a/tests/idris2/basic014/Rewrite.idr b/tests/idris2/basic/basic014/Rewrite.idr similarity index 100% rename from tests/idris2/basic014/Rewrite.idr rename to tests/idris2/basic/basic014/Rewrite.idr diff --git a/tests/idris2/basic014/expected b/tests/idris2/basic/basic014/expected similarity index 100% rename from tests/idris2/basic014/expected rename to tests/idris2/basic/basic014/expected diff --git a/tests/idris2/basic/basic014/run b/tests/idris2/basic/basic014/run new file mode 100755 index 0000000000..0a492040bb --- /dev/null +++ b/tests/idris2/basic/basic014/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Rewrite.idr diff --git a/tests/idris2/basic015/George.idr b/tests/idris2/basic/basic015/George.idr similarity index 100% rename from tests/idris2/basic015/George.idr rename to tests/idris2/basic/basic015/George.idr diff --git a/tests/idris2/basic015/expected b/tests/idris2/basic/basic015/expected similarity index 100% rename from tests/idris2/basic015/expected rename to tests/idris2/basic/basic015/expected diff --git a/tests/idris2/basic/basic015/run b/tests/idris2/basic/basic015/run new file mode 100755 index 0000000000..0ebd3e0d8f --- /dev/null +++ b/tests/idris2/basic/basic015/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check George.idr diff --git a/tests/idris2/basic016/Eta.idr b/tests/idris2/basic/basic016/Eta.idr similarity index 100% rename from tests/idris2/basic016/Eta.idr rename to tests/idris2/basic/basic016/Eta.idr diff --git a/tests/idris2/basic016/Eta2.idr b/tests/idris2/basic/basic016/Eta2.idr similarity index 100% rename from tests/idris2/basic016/Eta2.idr rename to tests/idris2/basic/basic016/Eta2.idr diff --git a/tests/idris2/basic016/expected b/tests/idris2/basic/basic016/expected similarity index 100% rename from tests/idris2/basic016/expected rename to tests/idris2/basic/basic016/expected diff --git a/tests/idris2/basic016/run b/tests/idris2/basic/basic016/run similarity index 55% rename from tests/idris2/basic016/run rename to tests/idris2/basic/basic016/run index 59d0d7f294..7050229028 100755 --- a/tests/idris2/basic016/run +++ b/tests/idris2/basic/basic016/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh check Eta.idr check Eta2.idr diff --git a/tests/idris2/basic017/CaseInf.idr b/tests/idris2/basic/basic017/CaseInf.idr similarity index 100% rename from tests/idris2/basic017/CaseInf.idr rename to tests/idris2/basic/basic017/CaseInf.idr diff --git a/tests/idris2/basic017/expected b/tests/idris2/basic/basic017/expected similarity index 100% rename from tests/idris2/basic017/expected rename to tests/idris2/basic/basic017/expected diff --git a/tests/idris2/basic017/input b/tests/idris2/basic/basic017/input similarity index 100% rename from tests/idris2/basic017/input rename to tests/idris2/basic/basic017/input diff --git a/tests/idris2/basic017/run b/tests/idris2/basic/basic017/run similarity index 53% rename from tests/idris2/basic017/run rename to tests/idris2/basic/basic017/run index bd22b25656..cc19a3f765 100755 --- a/tests/idris2/basic017/run +++ b/tests/idris2/basic/basic017/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 CaseInf.idr < input diff --git a/tests/idris2/basic018/Fin.idr b/tests/idris2/basic/basic018/Fin.idr similarity index 100% rename from tests/idris2/basic018/Fin.idr rename to tests/idris2/basic/basic018/Fin.idr diff --git a/tests/idris2/basic018/expected b/tests/idris2/basic/basic018/expected similarity index 100% rename from tests/idris2/basic018/expected rename to tests/idris2/basic/basic018/expected diff --git a/tests/idris2/basic/basic018/run b/tests/idris2/basic/basic018/run new file mode 100755 index 0000000000..ebad3d7e46 --- /dev/null +++ b/tests/idris2/basic/basic018/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Fin.idr diff --git a/tests/idris2/basic019/CaseBlock.idr b/tests/idris2/basic/basic019/CaseBlock.idr similarity index 100% rename from tests/idris2/basic019/CaseBlock.idr rename to tests/idris2/basic/basic019/CaseBlock.idr diff --git a/tests/idris2/basic019/expected b/tests/idris2/basic/basic019/expected similarity index 100% rename from tests/idris2/basic019/expected rename to tests/idris2/basic/basic019/expected diff --git a/tests/idris2/basic019/input b/tests/idris2/basic/basic019/input similarity index 100% rename from tests/idris2/basic019/input rename to tests/idris2/basic/basic019/input diff --git a/tests/idris2/basic019/run b/tests/idris2/basic/basic019/run similarity index 55% rename from tests/idris2/basic019/run rename to tests/idris2/basic/basic019/run index 4b9c047980..368802e634 100755 --- a/tests/idris2/basic019/run +++ b/tests/idris2/basic/basic019/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 CaseBlock.idr < input diff --git a/tests/idris2/basic020/Mut.idr b/tests/idris2/basic/basic020/Mut.idr similarity index 100% rename from tests/idris2/basic020/Mut.idr rename to tests/idris2/basic/basic020/Mut.idr diff --git a/tests/idris2/basic020/expected b/tests/idris2/basic/basic020/expected similarity index 100% rename from tests/idris2/basic020/expected rename to tests/idris2/basic/basic020/expected diff --git a/tests/idris2/basic020/input b/tests/idris2/basic/basic020/input similarity index 100% rename from tests/idris2/basic020/input rename to tests/idris2/basic/basic020/input diff --git a/tests/idris2/basic020/run b/tests/idris2/basic/basic020/run similarity index 50% rename from tests/idris2/basic020/run rename to tests/idris2/basic/basic020/run index 990791a5e1..f3cae4fb6d 100755 --- a/tests/idris2/basic020/run +++ b/tests/idris2/basic/basic020/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 Mut.idr < input diff --git a/tests/idris2/basic021/CaseDep.idr b/tests/idris2/basic/basic021/CaseDep.idr similarity index 100% rename from tests/idris2/basic021/CaseDep.idr rename to tests/idris2/basic/basic021/CaseDep.idr diff --git a/tests/idris2/basic021/expected b/tests/idris2/basic/basic021/expected similarity index 100% rename from tests/idris2/basic021/expected rename to tests/idris2/basic/basic021/expected diff --git a/tests/idris2/basic021/input b/tests/idris2/basic/basic021/input similarity index 100% rename from tests/idris2/basic021/input rename to tests/idris2/basic/basic021/input diff --git a/tests/idris2/basic021/run b/tests/idris2/basic/basic021/run similarity index 53% rename from tests/idris2/basic021/run rename to tests/idris2/basic/basic021/run index 02a25baea9..f359d35a83 100755 --- a/tests/idris2/basic021/run +++ b/tests/idris2/basic/basic021/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 CaseDep.idr < input diff --git a/tests/idris2/basic022/Erase.idr b/tests/idris2/basic/basic022/Erase.idr similarity index 100% rename from tests/idris2/basic022/Erase.idr rename to tests/idris2/basic/basic022/Erase.idr diff --git a/tests/idris2/basic022/expected b/tests/idris2/basic/basic022/expected similarity index 100% rename from tests/idris2/basic022/expected rename to tests/idris2/basic/basic022/expected diff --git a/tests/idris2/basic022/input b/tests/idris2/basic/basic022/input similarity index 100% rename from tests/idris2/basic022/input rename to tests/idris2/basic/basic022/input diff --git a/tests/idris2/basic022/run b/tests/idris2/basic/basic022/run similarity index 52% rename from tests/idris2/basic022/run rename to tests/idris2/basic/basic022/run index 2660c7e33c..d8bc439964 100755 --- a/tests/idris2/basic022/run +++ b/tests/idris2/basic/basic022/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 Erase.idr < input diff --git a/tests/idris2/basic023/Params.idr b/tests/idris2/basic/basic023/Params.idr similarity index 100% rename from tests/idris2/basic023/Params.idr rename to tests/idris2/basic/basic023/Params.idr diff --git a/tests/idris2/basic023/expected b/tests/idris2/basic/basic023/expected similarity index 100% rename from tests/idris2/basic023/expected rename to tests/idris2/basic/basic023/expected diff --git a/tests/idris2/basic023/input b/tests/idris2/basic/basic023/input similarity index 100% rename from tests/idris2/basic023/input rename to tests/idris2/basic/basic023/input diff --git a/tests/idris2/basic023/run b/tests/idris2/basic/basic023/run similarity index 52% rename from tests/idris2/basic023/run rename to tests/idris2/basic/basic023/run index d1e7b78c95..8d7e87a803 100755 --- a/tests/idris2/basic023/run +++ b/tests/idris2/basic/basic023/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 Params.idr < input diff --git a/tests/idris2/basic024/PatLam.idr b/tests/idris2/basic/basic024/PatLam.idr similarity index 100% rename from tests/idris2/basic024/PatLam.idr rename to tests/idris2/basic/basic024/PatLam.idr diff --git a/tests/idris2/basic024/expected b/tests/idris2/basic/basic024/expected similarity index 100% rename from tests/idris2/basic024/expected rename to tests/idris2/basic/basic024/expected diff --git a/tests/idris2/basic024/input b/tests/idris2/basic/basic024/input similarity index 100% rename from tests/idris2/basic024/input rename to tests/idris2/basic/basic024/input diff --git a/tests/idris2/basic024/run b/tests/idris2/basic/basic024/run similarity index 52% rename from tests/idris2/basic024/run rename to tests/idris2/basic/basic024/run index b716b4af86..6eff094927 100755 --- a/tests/idris2/basic024/run +++ b/tests/idris2/basic/basic024/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 PatLam.idr < input diff --git a/tests/idris2/basic025/expected b/tests/idris2/basic/basic025/expected similarity index 100% rename from tests/idris2/basic025/expected rename to tests/idris2/basic/basic025/expected diff --git a/tests/idris2/basic025/input b/tests/idris2/basic/basic025/input similarity index 100% rename from tests/idris2/basic025/input rename to tests/idris2/basic/basic025/input diff --git a/tests/idris2/basic/basic025/run b/tests/idris2/basic/basic025/run new file mode 100755 index 0000000000..404d7d47c6 --- /dev/null +++ b/tests/idris2/basic/basic025/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +idris2 < input diff --git a/tests/idris2/basic026/Erl.idr b/tests/idris2/basic/basic026/Erl.idr similarity index 100% rename from tests/idris2/basic026/Erl.idr rename to tests/idris2/basic/basic026/Erl.idr diff --git a/tests/idris2/basic026/expected b/tests/idris2/basic/basic026/expected similarity index 100% rename from tests/idris2/basic026/expected rename to tests/idris2/basic/basic026/expected diff --git a/tests/idris2/basic/basic026/run b/tests/idris2/basic/basic026/run new file mode 100755 index 0000000000..b37b5c33a3 --- /dev/null +++ b/tests/idris2/basic/basic026/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Erl.idr diff --git a/tests/idris2/basic027/Temp.idr b/tests/idris2/basic/basic027/Temp.idr similarity index 100% rename from tests/idris2/basic027/Temp.idr rename to tests/idris2/basic/basic027/Temp.idr diff --git a/tests/idris2/basic027/expected b/tests/idris2/basic/basic027/expected similarity index 100% rename from tests/idris2/basic027/expected rename to tests/idris2/basic/basic027/expected diff --git a/tests/idris2/basic/basic027/run b/tests/idris2/basic/basic027/run new file mode 100755 index 0000000000..eab5d203d9 --- /dev/null +++ b/tests/idris2/basic/basic027/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Temp.idr diff --git a/tests/idris2/basic028/Do.idr b/tests/idris2/basic/basic028/Do.idr similarity index 100% rename from tests/idris2/basic028/Do.idr rename to tests/idris2/basic/basic028/Do.idr diff --git a/tests/idris2/basic028/expected b/tests/idris2/basic/basic028/expected similarity index 100% rename from tests/idris2/basic028/expected rename to tests/idris2/basic/basic028/expected diff --git a/tests/idris2/basic028/input b/tests/idris2/basic/basic028/input similarity index 100% rename from tests/idris2/basic028/input rename to tests/idris2/basic/basic028/input diff --git a/tests/idris2/basic028/run b/tests/idris2/basic/basic028/run similarity index 66% rename from tests/idris2/basic028/run rename to tests/idris2/basic/basic028/run index 5122537366..8388f0f05a 100755 --- a/tests/idris2/basic028/run +++ b/tests/idris2/basic/basic028/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh unset IDRIS2_PATH diff --git a/tests/idris2/basic029/Params.idr b/tests/idris2/basic/basic029/Params.idr similarity index 100% rename from tests/idris2/basic029/Params.idr rename to tests/idris2/basic/basic029/Params.idr diff --git a/tests/idris2/basic029/expected b/tests/idris2/basic/basic029/expected similarity index 100% rename from tests/idris2/basic029/expected rename to tests/idris2/basic/basic029/expected diff --git a/tests/idris2/basic029/input b/tests/idris2/basic/basic029/input similarity index 100% rename from tests/idris2/basic029/input rename to tests/idris2/basic/basic029/input diff --git a/tests/idris2/basic029/run b/tests/idris2/basic/basic029/run similarity index 52% rename from tests/idris2/basic029/run rename to tests/idris2/basic/basic029/run index d1e7b78c95..8d7e87a803 100644 --- a/tests/idris2/basic029/run +++ b/tests/idris2/basic/basic029/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 Params.idr < input diff --git a/tests/idris2/basic030/arity.idr b/tests/idris2/basic/basic030/arity.idr similarity index 100% rename from tests/idris2/basic030/arity.idr rename to tests/idris2/basic/basic030/arity.idr diff --git a/tests/idris2/basic030/expected b/tests/idris2/basic/basic030/expected similarity index 100% rename from tests/idris2/basic030/expected rename to tests/idris2/basic/basic030/expected diff --git a/tests/idris2/basic/basic030/run b/tests/idris2/basic/basic030/run new file mode 100644 index 0000000000..7dcbdd49f3 --- /dev/null +++ b/tests/idris2/basic/basic030/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check arity.idr diff --git a/tests/idris2/basic031/erased.idr b/tests/idris2/basic/basic031/erased.idr similarity index 100% rename from tests/idris2/basic031/erased.idr rename to tests/idris2/basic/basic031/erased.idr diff --git a/tests/idris2/basic031/expected b/tests/idris2/basic/basic031/expected similarity index 100% rename from tests/idris2/basic031/expected rename to tests/idris2/basic/basic031/expected diff --git a/tests/idris2/basic/basic031/run b/tests/idris2/basic/basic031/run new file mode 100644 index 0000000000..086e01e6c9 --- /dev/null +++ b/tests/idris2/basic/basic031/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check erased.idr diff --git a/tests/idris2/basic032/Idiom.idr b/tests/idris2/basic/basic032/Idiom.idr similarity index 100% rename from tests/idris2/basic032/Idiom.idr rename to tests/idris2/basic/basic032/Idiom.idr diff --git a/tests/idris2/basic032/Idiom2.idr b/tests/idris2/basic/basic032/Idiom2.idr similarity index 100% rename from tests/idris2/basic032/Idiom2.idr rename to tests/idris2/basic/basic032/Idiom2.idr diff --git a/tests/idris2/basic032/expected b/tests/idris2/basic/basic032/expected similarity index 100% rename from tests/idris2/basic032/expected rename to tests/idris2/basic/basic032/expected diff --git a/tests/idris2/basic032/input b/tests/idris2/basic/basic032/input similarity index 100% rename from tests/idris2/basic032/input rename to tests/idris2/basic/basic032/input diff --git a/tests/idris2/basic032/run b/tests/idris2/basic/basic032/run similarity index 64% rename from tests/idris2/basic032/run rename to tests/idris2/basic/basic032/run index ffe7a5e4a8..cb102bccd1 100755 --- a/tests/idris2/basic032/run +++ b/tests/idris2/basic/basic032/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 Idiom.idr < input check Idiom2.idr diff --git a/tests/idris2/basic033/expected b/tests/idris2/basic/basic033/expected similarity index 100% rename from tests/idris2/basic033/expected rename to tests/idris2/basic/basic033/expected diff --git a/tests/idris2/basic/basic033/run b/tests/idris2/basic/basic033/run new file mode 100644 index 0000000000..88a843ffb8 --- /dev/null +++ b/tests/idris2/basic/basic033/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check unboundimps.idr diff --git a/tests/idris2/basic033/unboundimps.idr b/tests/idris2/basic/basic033/unboundimps.idr similarity index 100% rename from tests/idris2/basic033/unboundimps.idr rename to tests/idris2/basic/basic033/unboundimps.idr diff --git a/tests/idris2/basic034/expected b/tests/idris2/basic/basic034/expected similarity index 100% rename from tests/idris2/basic034/expected rename to tests/idris2/basic/basic034/expected diff --git a/tests/idris2/basic034/lets.idr b/tests/idris2/basic/basic034/lets.idr similarity index 100% rename from tests/idris2/basic034/lets.idr rename to tests/idris2/basic/basic034/lets.idr diff --git a/tests/idris2/basic/basic034/run b/tests/idris2/basic/basic034/run new file mode 100644 index 0000000000..60eb85a055 --- /dev/null +++ b/tests/idris2/basic/basic034/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check lets.idr diff --git a/tests/idris2/basic035/expected b/tests/idris2/basic/basic035/expected similarity index 100% rename from tests/idris2/basic035/expected rename to tests/idris2/basic/basic035/expected diff --git a/tests/idris2/basic035/input b/tests/idris2/basic/basic035/input similarity index 100% rename from tests/idris2/basic035/input rename to tests/idris2/basic/basic035/input diff --git a/tests/idris2/basic035/run b/tests/idris2/basic/basic035/run similarity index 52% rename from tests/idris2/basic035/run rename to tests/idris2/basic/basic035/run index 63bae6413c..43c34e185c 100755 --- a/tests/idris2/basic035/run +++ b/tests/idris2/basic/basic035/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 using.idr < input diff --git a/tests/idris2/basic035/using.idr b/tests/idris2/basic/basic035/using.idr similarity index 100% rename from tests/idris2/basic035/using.idr rename to tests/idris2/basic/basic035/using.idr diff --git a/tests/idris2/basic036/defimp.idr b/tests/idris2/basic/basic036/defimp.idr similarity index 100% rename from tests/idris2/basic036/defimp.idr rename to tests/idris2/basic/basic036/defimp.idr diff --git a/tests/idris2/basic036/expected b/tests/idris2/basic/basic036/expected similarity index 100% rename from tests/idris2/basic036/expected rename to tests/idris2/basic/basic036/expected diff --git a/tests/idris2/basic036/input b/tests/idris2/basic/basic036/input similarity index 100% rename from tests/idris2/basic036/input rename to tests/idris2/basic/basic036/input diff --git a/tests/idris2/basic036/run b/tests/idris2/basic/basic036/run similarity index 52% rename from tests/idris2/basic036/run rename to tests/idris2/basic/basic036/run index b094756dc1..e9b175ff8b 100755 --- a/tests/idris2/basic036/run +++ b/tests/idris2/basic/basic036/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 defimp.idr < input diff --git a/tests/idris2/basic037/Comments.idr b/tests/idris2/basic/basic037/Comments.idr similarity index 100% rename from tests/idris2/basic037/Comments.idr rename to tests/idris2/basic/basic037/Comments.idr diff --git a/tests/idris2/basic037/Issue279.idr b/tests/idris2/basic/basic037/Issue279.idr similarity index 100% rename from tests/idris2/basic037/Issue279.idr rename to tests/idris2/basic/basic037/Issue279.idr diff --git a/tests/idris2/basic037/expected b/tests/idris2/basic/basic037/expected similarity index 100% rename from tests/idris2/basic037/expected rename to tests/idris2/basic/basic037/expected diff --git a/tests/idris2/basic037/run b/tests/idris2/basic/basic037/run similarity index 76% rename from tests/idris2/basic037/run rename to tests/idris2/basic/basic037/run index 9675168217..8a233e3f1f 100644 --- a/tests/idris2/basic037/run +++ b/tests/idris2/basic/basic037/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh echo ':q' | idris2 --no-prelude Comments.idr echo ':q' | idris2 Issue279.idr diff --git a/tests/idris2/basic038/Resugar.idr b/tests/idris2/basic/basic038/Resugar.idr similarity index 100% rename from tests/idris2/basic038/Resugar.idr rename to tests/idris2/basic/basic038/Resugar.idr diff --git a/tests/idris2/basic038/expected b/tests/idris2/basic/basic038/expected similarity index 100% rename from tests/idris2/basic038/expected rename to tests/idris2/basic/basic038/expected diff --git a/tests/idris2/basic038/input b/tests/idris2/basic/basic038/input similarity index 100% rename from tests/idris2/basic038/input rename to tests/idris2/basic/basic038/input diff --git a/tests/idris2/basic038/run b/tests/idris2/basic/basic038/run similarity index 53% rename from tests/idris2/basic038/run rename to tests/idris2/basic/basic038/run index 695b0e59b2..30d53d2b8f 100644 --- a/tests/idris2/basic038/run +++ b/tests/idris2/basic/basic038/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 Resugar.idr < input diff --git a/tests/idris2/basic039/Main.idr b/tests/idris2/basic/basic039/Main.idr similarity index 100% rename from tests/idris2/basic039/Main.idr rename to tests/idris2/basic/basic039/Main.idr diff --git a/tests/idris2/basic039/expected b/tests/idris2/basic/basic039/expected similarity index 100% rename from tests/idris2/basic039/expected rename to tests/idris2/basic/basic039/expected diff --git a/tests/idris2/basic039/input b/tests/idris2/basic/basic039/input similarity index 100% rename from tests/idris2/basic039/input rename to tests/idris2/basic/basic039/input diff --git a/tests/idris2/basic039/run b/tests/idris2/basic/basic039/run similarity index 51% rename from tests/idris2/basic039/run rename to tests/idris2/basic/basic039/run index 37beba4789..7935482b8c 100755 --- a/tests/idris2/basic039/run +++ b/tests/idris2/basic/basic039/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 Main.idr < input diff --git a/tests/idris2/basic040/Default.idr b/tests/idris2/basic/basic040/Default.idr similarity index 100% rename from tests/idris2/basic040/Default.idr rename to tests/idris2/basic/basic040/Default.idr diff --git a/tests/idris2/basic040/expected b/tests/idris2/basic/basic040/expected similarity index 100% rename from tests/idris2/basic040/expected rename to tests/idris2/basic/basic040/expected diff --git a/tests/idris2/basic040/run b/tests/idris2/basic/basic040/run similarity index 57% rename from tests/idris2/basic040/run rename to tests/idris2/basic/basic040/run index 68be330b69..af0feb020f 100755 --- a/tests/idris2/basic040/run +++ b/tests/idris2/basic/basic040/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh echo ":q" | idris2 Default.idr diff --git a/tests/idris2/basic041/QDo.idr b/tests/idris2/basic/basic041/QDo.idr similarity index 100% rename from tests/idris2/basic041/QDo.idr rename to tests/idris2/basic/basic041/QDo.idr diff --git a/tests/idris2/basic041/expected b/tests/idris2/basic/basic041/expected similarity index 100% rename from tests/idris2/basic041/expected rename to tests/idris2/basic/basic041/expected diff --git a/tests/idris2/basic/basic041/run b/tests/idris2/basic/basic041/run new file mode 100755 index 0000000000..5c52e686b4 --- /dev/null +++ b/tests/idris2/basic/basic041/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check QDo.idr diff --git a/tests/idris2/basic042/LiteralsInteger.idr b/tests/idris2/basic/basic042/LiteralsInteger.idr similarity index 100% rename from tests/idris2/basic042/LiteralsInteger.idr rename to tests/idris2/basic/basic042/LiteralsInteger.idr diff --git a/tests/idris2/basic042/LiteralsString.idr b/tests/idris2/basic/basic042/LiteralsString.idr similarity index 100% rename from tests/idris2/basic042/LiteralsString.idr rename to tests/idris2/basic/basic042/LiteralsString.idr diff --git a/tests/idris2/basic042/expected b/tests/idris2/basic/basic042/expected similarity index 100% rename from tests/idris2/basic042/expected rename to tests/idris2/basic/basic042/expected diff --git a/tests/idris2/basic042/input b/tests/idris2/basic/basic042/input similarity index 100% rename from tests/idris2/basic042/input rename to tests/idris2/basic/basic042/input diff --git a/tests/idris2/basic042/input2 b/tests/idris2/basic/basic042/input2 similarity index 100% rename from tests/idris2/basic042/input2 rename to tests/idris2/basic/basic042/input2 diff --git a/tests/idris2/basic042/run b/tests/idris2/basic/basic042/run similarity index 74% rename from tests/idris2/basic042/run rename to tests/idris2/basic/basic042/run index 547c22b06e..af97af7fb1 100755 --- a/tests/idris2/basic042/run +++ b/tests/idris2/basic/basic042/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 LiteralsString.idr < input idris2 LiteralsInteger.idr < input2 diff --git a/tests/idris2/basic043/BitCasts.idr b/tests/idris2/basic/basic043/BitCasts.idr similarity index 100% rename from tests/idris2/basic043/BitCasts.idr rename to tests/idris2/basic/basic043/BitCasts.idr diff --git a/tests/idris2/basic043/expected b/tests/idris2/basic/basic043/expected similarity index 100% rename from tests/idris2/basic043/expected rename to tests/idris2/basic/basic043/expected diff --git a/tests/idris2/basic043/input b/tests/idris2/basic/basic043/input similarity index 100% rename from tests/idris2/basic043/input rename to tests/idris2/basic/basic043/input diff --git a/tests/idris2/basic043/run b/tests/idris2/basic/basic043/run similarity index 54% rename from tests/idris2/basic043/run rename to tests/idris2/basic/basic043/run index e4a0b380e2..4b99700232 100644 --- a/tests/idris2/basic043/run +++ b/tests/idris2/basic/basic043/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 BitCasts.idr < input diff --git a/tests/idris2/basic044/Term.idr b/tests/idris2/basic/basic044/Term.idr similarity index 100% rename from tests/idris2/basic044/Term.idr rename to tests/idris2/basic/basic044/Term.idr diff --git a/tests/idris2/basic044/Vec.idr b/tests/idris2/basic/basic044/Vec.idr similarity index 100% rename from tests/idris2/basic044/Vec.idr rename to tests/idris2/basic/basic044/Vec.idr diff --git a/tests/idris2/basic044/expected b/tests/idris2/basic/basic044/expected similarity index 100% rename from tests/idris2/basic044/expected rename to tests/idris2/basic/basic044/expected diff --git a/tests/idris2/basic044/run b/tests/idris2/basic/basic044/run similarity index 94% rename from tests/idris2/basic044/run rename to tests/idris2/basic/basic044/run index 0ce8521859..0d793db537 100644 --- a/tests/idris2/basic044/run +++ b/tests/idris2/basic/basic044/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh echo ":q" | idris2 --log unify.equal:10 --log unify:5 Term.idr \ | sed -E "s/[0-9]+}/N}/g" | sed -E "s/resolved([0-9]+)/resolvedN/g" \ diff --git a/tests/idris2/basic045/Main.idr b/tests/idris2/basic/basic045/Main.idr similarity index 100% rename from tests/idris2/basic045/Main.idr rename to tests/idris2/basic/basic045/Main.idr diff --git a/tests/idris2/basic045/expected b/tests/idris2/basic/basic045/expected similarity index 100% rename from tests/idris2/basic045/expected rename to tests/idris2/basic/basic045/expected diff --git a/tests/idris2/basic/basic045/run b/tests/idris2/basic/basic045/run new file mode 100644 index 0000000000..d49bfd87de --- /dev/null +++ b/tests/idris2/basic/basic045/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Main.idr diff --git a/tests/idris2/basic046/TupleSections.idr b/tests/idris2/basic/basic046/TupleSections.idr similarity index 100% rename from tests/idris2/basic046/TupleSections.idr rename to tests/idris2/basic/basic046/TupleSections.idr diff --git a/tests/idris2/basic046/expected b/tests/idris2/basic/basic046/expected similarity index 100% rename from tests/idris2/basic046/expected rename to tests/idris2/basic/basic046/expected diff --git a/tests/idris2/basic046/input b/tests/idris2/basic/basic046/input similarity index 100% rename from tests/idris2/basic046/input rename to tests/idris2/basic/basic046/input diff --git a/tests/idris2/basic046/run b/tests/idris2/basic/basic046/run similarity index 58% rename from tests/idris2/basic046/run rename to tests/idris2/basic/basic046/run index f6b12aa546..a3f5475c18 100644 --- a/tests/idris2/basic046/run +++ b/tests/idris2/basic/basic046/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 TupleSections.idr < input diff --git a/tests/idris2/basic047/InterleavingLets.idr b/tests/idris2/basic/basic047/InterleavingLets.idr similarity index 100% rename from tests/idris2/basic047/InterleavingLets.idr rename to tests/idris2/basic/basic047/InterleavingLets.idr diff --git a/tests/idris2/basic047/expected b/tests/idris2/basic/basic047/expected similarity index 100% rename from tests/idris2/basic047/expected rename to tests/idris2/basic/basic047/expected diff --git a/tests/idris2/basic047/input b/tests/idris2/basic/basic047/input similarity index 100% rename from tests/idris2/basic047/input rename to tests/idris2/basic/basic047/input diff --git a/tests/idris2/basic047/run b/tests/idris2/basic/basic047/run similarity index 60% rename from tests/idris2/basic047/run rename to tests/idris2/basic/basic047/run index 2c8b2324b7..dadfb3a5b9 100644 --- a/tests/idris2/basic047/run +++ b/tests/idris2/basic/basic047/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 InterleavingLets.idr < input diff --git a/tests/idris2/basic049/Fld.idr b/tests/idris2/basic/basic049/Fld.idr similarity index 100% rename from tests/idris2/basic049/Fld.idr rename to tests/idris2/basic/basic049/Fld.idr diff --git a/tests/idris2/basic049/expected b/tests/idris2/basic/basic049/expected similarity index 100% rename from tests/idris2/basic049/expected rename to tests/idris2/basic/basic049/expected diff --git a/tests/idris2/basic049/input b/tests/idris2/basic/basic049/input similarity index 100% rename from tests/idris2/basic049/input rename to tests/idris2/basic/basic049/input diff --git a/tests/idris2/basic049/run b/tests/idris2/basic/basic049/run similarity index 50% rename from tests/idris2/basic049/run rename to tests/idris2/basic/basic049/run index da8d57ecaf..a4277f19f4 100755 --- a/tests/idris2/basic049/run +++ b/tests/idris2/basic/basic049/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 Fld.idr < input diff --git a/tests/idris2/basic050/Ilc.idr b/tests/idris2/basic/basic050/Ilc.idr similarity index 100% rename from tests/idris2/basic050/Ilc.idr rename to tests/idris2/basic/basic050/Ilc.idr diff --git a/tests/idris2/basic050/expected b/tests/idris2/basic/basic050/expected similarity index 100% rename from tests/idris2/basic050/expected rename to tests/idris2/basic/basic050/expected diff --git a/tests/idris2/basic050/input b/tests/idris2/basic/basic050/input similarity index 100% rename from tests/idris2/basic050/input rename to tests/idris2/basic/basic050/input diff --git a/tests/idris2/basic050/run b/tests/idris2/basic/basic050/run similarity index 50% rename from tests/idris2/basic050/run rename to tests/idris2/basic/basic050/run index 15eb96c095..494dc1ef3c 100755 --- a/tests/idris2/basic050/run +++ b/tests/idris2/basic/basic050/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 Ilc.idr < input diff --git a/tests/idris2/basic051/Issue833.idr b/tests/idris2/basic/basic051/Issue833.idr similarity index 100% rename from tests/idris2/basic051/Issue833.idr rename to tests/idris2/basic/basic051/Issue833.idr diff --git a/tests/idris2/basic051/expected b/tests/idris2/basic/basic051/expected similarity index 100% rename from tests/idris2/basic051/expected rename to tests/idris2/basic/basic051/expected diff --git a/tests/idris2/basic/basic051/run b/tests/idris2/basic/basic051/run new file mode 100755 index 0000000000..02b533d6d8 --- /dev/null +++ b/tests/idris2/basic/basic051/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Issue833.idr diff --git a/tests/idris2/basic052/DoubleClBrace.idr b/tests/idris2/basic/basic052/DoubleClBrace.idr similarity index 100% rename from tests/idris2/basic052/DoubleClBrace.idr rename to tests/idris2/basic/basic052/DoubleClBrace.idr diff --git a/tests/idris2/basic052/expected b/tests/idris2/basic/basic052/expected similarity index 100% rename from tests/idris2/basic052/expected rename to tests/idris2/basic/basic052/expected diff --git a/tests/idris2/basic052/input b/tests/idris2/basic/basic052/input similarity index 100% rename from tests/idris2/basic052/input rename to tests/idris2/basic/basic052/input diff --git a/tests/idris2/basic052/run b/tests/idris2/basic/basic052/run similarity index 58% rename from tests/idris2/basic052/run rename to tests/idris2/basic/basic052/run index c127a186b1..4952ad9c38 100755 --- a/tests/idris2/basic052/run +++ b/tests/idris2/basic/basic052/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 DoubleClBrace.idr < input diff --git a/tests/idris2/basic053/UnderscoredIntegerLiterals.idr b/tests/idris2/basic/basic053/UnderscoredIntegerLiterals.idr similarity index 100% rename from tests/idris2/basic053/UnderscoredIntegerLiterals.idr rename to tests/idris2/basic/basic053/UnderscoredIntegerLiterals.idr diff --git a/tests/idris2/basic053/expected b/tests/idris2/basic/basic053/expected similarity index 100% rename from tests/idris2/basic053/expected rename to tests/idris2/basic/basic053/expected diff --git a/tests/idris2/basic053/input b/tests/idris2/basic/basic053/input similarity index 100% rename from tests/idris2/basic053/input rename to tests/idris2/basic/basic053/input diff --git a/tests/idris2/basic053/run b/tests/idris2/basic/basic053/run similarity index 66% rename from tests/idris2/basic053/run rename to tests/idris2/basic/basic053/run index 235e06f77b..d924d851a3 100755 --- a/tests/idris2/basic053/run +++ b/tests/idris2/basic/basic053/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 UnderscoredIntegerLiterals.idr < input diff --git a/tests/idris2/basic054/Issue1023.idr b/tests/idris2/basic/basic054/Issue1023.idr similarity index 100% rename from tests/idris2/basic054/Issue1023.idr rename to tests/idris2/basic/basic054/Issue1023.idr diff --git a/tests/idris2/basic054/expected b/tests/idris2/basic/basic054/expected similarity index 100% rename from tests/idris2/basic054/expected rename to tests/idris2/basic/basic054/expected diff --git a/tests/idris2/basic/basic054/run b/tests/idris2/basic/basic054/run new file mode 100755 index 0000000000..09fda7781e --- /dev/null +++ b/tests/idris2/basic/basic054/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +run Issue1023.idr diff --git a/tests/idris2/basic055/BitOps.idr b/tests/idris2/basic/basic055/BitOps.idr similarity index 100% rename from tests/idris2/basic055/BitOps.idr rename to tests/idris2/basic/basic055/BitOps.idr diff --git a/tests/idris2/basic055/expected b/tests/idris2/basic/basic055/expected similarity index 100% rename from tests/idris2/basic055/expected rename to tests/idris2/basic/basic055/expected diff --git a/tests/idris2/basic/basic055/run b/tests/idris2/basic/basic055/run new file mode 100644 index 0000000000..0d04a04725 --- /dev/null +++ b/tests/idris2/basic/basic055/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +run BitOps.idr diff --git a/tests/idris2/basic056/DoubleLit.idr b/tests/idris2/basic/basic056/DoubleLit.idr similarity index 100% rename from tests/idris2/basic056/DoubleLit.idr rename to tests/idris2/basic/basic056/DoubleLit.idr diff --git a/tests/idris2/basic056/expected b/tests/idris2/basic/basic056/expected similarity index 100% rename from tests/idris2/basic056/expected rename to tests/idris2/basic/basic056/expected diff --git a/tests/idris2/basic/basic056/run b/tests/idris2/basic/basic056/run new file mode 100644 index 0000000000..9798c05104 --- /dev/null +++ b/tests/idris2/basic/basic056/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +run DoubleLit.idr diff --git a/tests/idris2/basic057/LetIn.idr b/tests/idris2/basic/basic057/LetIn.idr similarity index 100% rename from tests/idris2/basic057/LetIn.idr rename to tests/idris2/basic/basic057/LetIn.idr diff --git a/tests/idris2/basic057/expected b/tests/idris2/basic/basic057/expected similarity index 100% rename from tests/idris2/basic057/expected rename to tests/idris2/basic/basic057/expected diff --git a/tests/idris2/basic/basic057/run b/tests/idris2/basic/basic057/run new file mode 100644 index 0000000000..5bd5697ca6 --- /dev/null +++ b/tests/idris2/basic/basic057/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check LetIn.idr diff --git a/tests/idris2/basic058/DataTypeOp.idr b/tests/idris2/basic/basic058/DataTypeOp.idr similarity index 100% rename from tests/idris2/basic058/DataTypeOp.idr rename to tests/idris2/basic/basic058/DataTypeOp.idr diff --git a/tests/idris2/basic058/DataTypeProj.idr b/tests/idris2/basic/basic058/DataTypeProj.idr similarity index 100% rename from tests/idris2/basic058/DataTypeProj.idr rename to tests/idris2/basic/basic058/DataTypeProj.idr diff --git a/tests/idris2/basic058/expected b/tests/idris2/basic/basic058/expected similarity index 100% rename from tests/idris2/basic058/expected rename to tests/idris2/basic/basic058/expected diff --git a/tests/idris2/basic058/run b/tests/idris2/basic/basic058/run similarity index 65% rename from tests/idris2/basic058/run rename to tests/idris2/basic/basic058/run index ca076d8403..9c69391e03 100644 --- a/tests/idris2/basic058/run +++ b/tests/idris2/basic/basic058/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh check DataTypeOp.idr check DataTypeProj.idr diff --git a/tests/idris2/basic059/MultiClaim.idr b/tests/idris2/basic/basic059/MultiClaim.idr similarity index 100% rename from tests/idris2/basic059/MultiClaim.idr rename to tests/idris2/basic/basic059/MultiClaim.idr diff --git a/tests/idris2/basic059/expected b/tests/idris2/basic/basic059/expected similarity index 100% rename from tests/idris2/basic059/expected rename to tests/idris2/basic/basic059/expected diff --git a/tests/idris2/basic/basic059/run b/tests/idris2/basic/basic059/run new file mode 100644 index 0000000000..8950dad72d --- /dev/null +++ b/tests/idris2/basic/basic059/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check MultiClaim.idr diff --git a/tests/idris2/basic060/Snoc.idr b/tests/idris2/basic/basic060/Snoc.idr similarity index 100% rename from tests/idris2/basic060/Snoc.idr rename to tests/idris2/basic/basic060/Snoc.idr diff --git a/tests/idris2/basic060/expected b/tests/idris2/basic/basic060/expected similarity index 100% rename from tests/idris2/basic060/expected rename to tests/idris2/basic/basic060/expected diff --git a/tests/idris2/basic060/input b/tests/idris2/basic/basic060/input similarity index 100% rename from tests/idris2/basic060/input rename to tests/idris2/basic/basic060/input diff --git a/tests/idris2/basic060/run b/tests/idris2/basic/basic060/run similarity index 51% rename from tests/idris2/basic060/run rename to tests/idris2/basic/basic060/run index ead0a5e691..a688ea6b18 100755 --- a/tests/idris2/basic060/run +++ b/tests/idris2/basic/basic060/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 Snoc.idr < input diff --git a/tests/idris2/basic061/IgnoreDo.idr b/tests/idris2/basic/basic061/IgnoreDo.idr similarity index 100% rename from tests/idris2/basic061/IgnoreDo.idr rename to tests/idris2/basic/basic061/IgnoreDo.idr diff --git a/tests/idris2/basic061/expected b/tests/idris2/basic/basic061/expected similarity index 100% rename from tests/idris2/basic061/expected rename to tests/idris2/basic/basic061/expected diff --git a/tests/idris2/basic061/input b/tests/idris2/basic/basic061/input similarity index 100% rename from tests/idris2/basic061/input rename to tests/idris2/basic/basic061/input diff --git a/tests/idris2/basic061/run b/tests/idris2/basic/basic061/run similarity index 54% rename from tests/idris2/basic061/run rename to tests/idris2/basic/basic061/run index fccd22fb91..5f1776797c 100755 --- a/tests/idris2/basic061/run +++ b/tests/idris2/basic/basic061/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 IgnoreDo.idr < input diff --git a/tests/idris2/basic062/Issue1943.idr b/tests/idris2/basic/basic062/Issue1943.idr similarity index 100% rename from tests/idris2/basic062/Issue1943.idr rename to tests/idris2/basic/basic062/Issue1943.idr diff --git a/tests/idris2/basic062/expected b/tests/idris2/basic/basic062/expected similarity index 100% rename from tests/idris2/basic062/expected rename to tests/idris2/basic/basic062/expected diff --git a/tests/idris2/basic/basic062/run b/tests/idris2/basic/basic062/run new file mode 100755 index 0000000000..3406c1416c --- /dev/null +++ b/tests/idris2/basic/basic062/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Issue1943.idr diff --git a/tests/idris2/basic063/NoDeclaration.idr b/tests/idris2/basic/basic063/NoDeclaration.idr similarity index 100% rename from tests/idris2/basic063/NoDeclaration.idr rename to tests/idris2/basic/basic063/NoDeclaration.idr diff --git a/tests/idris2/basic063/expected b/tests/idris2/basic/basic063/expected similarity index 100% rename from tests/idris2/basic063/expected rename to tests/idris2/basic/basic063/expected diff --git a/tests/idris2/basic063/input b/tests/idris2/basic/basic063/input similarity index 100% rename from tests/idris2/basic063/input rename to tests/idris2/basic/basic063/input diff --git a/tests/idris2/basic063/run b/tests/idris2/basic/basic063/run similarity index 58% rename from tests/idris2/basic063/run rename to tests/idris2/basic/basic063/run index d7e55ad492..bcecc378a5 100755 --- a/tests/idris2/basic063/run +++ b/tests/idris2/basic/basic063/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 NoDeclaration.idr < input diff --git a/tests/idris2/basic064/Issue2072.idr b/tests/idris2/basic/basic064/Issue2072.idr similarity index 100% rename from tests/idris2/basic064/Issue2072.idr rename to tests/idris2/basic/basic064/Issue2072.idr diff --git a/tests/idris2/basic064/expected b/tests/idris2/basic/basic064/expected similarity index 100% rename from tests/idris2/basic064/expected rename to tests/idris2/basic/basic064/expected diff --git a/tests/idris2/basic064/input b/tests/idris2/basic/basic064/input similarity index 100% rename from tests/idris2/basic064/input rename to tests/idris2/basic/basic064/input diff --git a/tests/idris2/basic064/run b/tests/idris2/basic/basic064/run similarity index 55% rename from tests/idris2/basic064/run rename to tests/idris2/basic/basic064/run index 8dd2c571f7..51ca2c4e5e 100755 --- a/tests/idris2/basic064/run +++ b/tests/idris2/basic/basic064/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 Issue2072.idr < input diff --git a/tests/idris2/basic065/Issue215.idr b/tests/idris2/basic/basic065/Issue215.idr similarity index 100% rename from tests/idris2/basic065/Issue215.idr rename to tests/idris2/basic/basic065/Issue215.idr diff --git a/tests/idris2/basic065/expected b/tests/idris2/basic/basic065/expected similarity index 100% rename from tests/idris2/basic065/expected rename to tests/idris2/basic/basic065/expected diff --git a/tests/idris2/basic/basic065/run b/tests/idris2/basic/basic065/run new file mode 100755 index 0000000000..3da76eaca9 --- /dev/null +++ b/tests/idris2/basic/basic065/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Issue215.idr diff --git a/tests/idris2/basic066/comment.idr b/tests/idris2/basic/basic066/comment.idr similarity index 100% rename from tests/idris2/basic066/comment.idr rename to tests/idris2/basic/basic066/comment.idr diff --git a/tests/idris2/basic066/expected b/tests/idris2/basic/basic066/expected similarity index 100% rename from tests/idris2/basic066/expected rename to tests/idris2/basic/basic066/expected diff --git a/tests/idris2/basic/basic066/run b/tests/idris2/basic/basic066/run new file mode 100755 index 0000000000..0b3949bb32 --- /dev/null +++ b/tests/idris2/basic/basic066/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check comment.idr diff --git a/tests/idris2/basic067/expected b/tests/idris2/basic/basic067/expected similarity index 100% rename from tests/idris2/basic067/expected rename to tests/idris2/basic/basic067/expected diff --git a/tests/idris2/basic067/input b/tests/idris2/basic/basic067/input similarity index 100% rename from tests/idris2/basic067/input rename to tests/idris2/basic/basic067/input diff --git a/tests/idris2/basic067/run b/tests/idris2/basic/basic067/run similarity index 86% rename from tests/idris2/basic067/run rename to tests/idris2/basic/basic067/run index 64c026cfa2..ae00c2d6ae 100755 --- a/tests/idris2/basic067/run +++ b/tests/idris2/basic/basic067/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh echo "unclosed1.idr" idris2 unclosed1.idr < input diff --git a/tests/idris2/basic067/unclosed1.idr b/tests/idris2/basic/basic067/unclosed1.idr similarity index 100% rename from tests/idris2/basic067/unclosed1.idr rename to tests/idris2/basic/basic067/unclosed1.idr diff --git a/tests/idris2/basic067/unclosed2.idr b/tests/idris2/basic/basic067/unclosed2.idr similarity index 100% rename from tests/idris2/basic067/unclosed2.idr rename to tests/idris2/basic/basic067/unclosed2.idr diff --git a/tests/idris2/basic067/unclosed3.idr b/tests/idris2/basic/basic067/unclosed3.idr similarity index 100% rename from tests/idris2/basic067/unclosed3.idr rename to tests/idris2/basic/basic067/unclosed3.idr diff --git a/tests/idris2/basic068/Issue2138.idr b/tests/idris2/basic/basic068/Issue2138.idr similarity index 100% rename from tests/idris2/basic068/Issue2138.idr rename to tests/idris2/basic/basic068/Issue2138.idr diff --git a/tests/idris2/basic068/expected b/tests/idris2/basic/basic068/expected similarity index 100% rename from tests/idris2/basic068/expected rename to tests/idris2/basic/basic068/expected diff --git a/tests/idris2/basic/basic068/run b/tests/idris2/basic/basic068/run new file mode 100755 index 0000000000..c813f1574d --- /dev/null +++ b/tests/idris2/basic/basic068/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +run Issue2138.idr diff --git a/tests/idris2/basic069/DebugInfo.idr b/tests/idris2/basic/basic069/DebugInfo.idr similarity index 100% rename from tests/idris2/basic069/DebugInfo.idr rename to tests/idris2/basic/basic069/DebugInfo.idr diff --git a/tests/idris2/basic069/expected b/tests/idris2/basic/basic069/expected similarity index 100% rename from tests/idris2/basic069/expected rename to tests/idris2/basic/basic069/expected diff --git a/tests/idris2/basic/basic069/run b/tests/idris2/basic/basic069/run new file mode 100755 index 0000000000..edec0db1ba --- /dev/null +++ b/tests/idris2/basic/basic069/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +run DebugInfo.idr diff --git a/tests/idris2/basic070/Issue2592.idr b/tests/idris2/basic/basic070/Issue2592.idr similarity index 100% rename from tests/idris2/basic070/Issue2592.idr rename to tests/idris2/basic/basic070/Issue2592.idr diff --git a/tests/idris2/basic070/Issue2593.idr b/tests/idris2/basic/basic070/Issue2593.idr similarity index 100% rename from tests/idris2/basic070/Issue2593.idr rename to tests/idris2/basic/basic070/Issue2593.idr diff --git a/tests/idris2/basic070/Issue2782.idr b/tests/idris2/basic/basic070/Issue2782.idr similarity index 100% rename from tests/idris2/basic070/Issue2782.idr rename to tests/idris2/basic/basic070/Issue2782.idr diff --git a/tests/idris2/basic070/Issue3016.idr b/tests/idris2/basic/basic070/Issue3016.idr similarity index 100% rename from tests/idris2/basic070/Issue3016.idr rename to tests/idris2/basic/basic070/Issue3016.idr diff --git a/tests/idris2/basic070/expected b/tests/idris2/basic/basic070/expected similarity index 100% rename from tests/idris2/basic070/expected rename to tests/idris2/basic/basic070/expected diff --git a/tests/idris2/basic070/run b/tests/idris2/basic/basic070/run similarity index 77% rename from tests/idris2/basic070/run rename to tests/idris2/basic/basic070/run index b3a8f50554..9f93779e45 100755 --- a/tests/idris2/basic070/run +++ b/tests/idris2/basic/basic070/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh check Issue3016.idr check Issue2782.idr diff --git a/tests/idris2/basic071/A.idr b/tests/idris2/basic/basic071/A.idr similarity index 100% rename from tests/idris2/basic071/A.idr rename to tests/idris2/basic/basic071/A.idr diff --git a/tests/idris2/basic071/B.idr b/tests/idris2/basic/basic071/B.idr similarity index 100% rename from tests/idris2/basic071/B.idr rename to tests/idris2/basic/basic071/B.idr diff --git a/tests/idris2/basic071/expected b/tests/idris2/basic/basic071/expected similarity index 100% rename from tests/idris2/basic071/expected rename to tests/idris2/basic/basic071/expected diff --git a/tests/idris2/basic071/run b/tests/idris2/basic/basic071/run similarity index 100% rename from tests/idris2/basic071/run rename to tests/idris2/basic/basic071/run diff --git a/tests/idris2/case001/InlineCase.idr b/tests/idris2/basic/case001/InlineCase.idr similarity index 100% rename from tests/idris2/case001/InlineCase.idr rename to tests/idris2/basic/case001/InlineCase.idr diff --git a/tests/idris2/case001/expected b/tests/idris2/basic/case001/expected similarity index 100% rename from tests/idris2/case001/expected rename to tests/idris2/basic/case001/expected diff --git a/tests/idris2/case001/run b/tests/idris2/basic/case001/run similarity index 100% rename from tests/idris2/case001/run rename to tests/idris2/basic/case001/run diff --git a/tests/idris2/dotted001/Issue2726.idr b/tests/idris2/basic/dotted001/Issue2726.idr similarity index 100% rename from tests/idris2/dotted001/Issue2726.idr rename to tests/idris2/basic/dotted001/Issue2726.idr diff --git a/tests/idris2/dotted001/expected b/tests/idris2/basic/dotted001/expected similarity index 100% rename from tests/idris2/dotted001/expected rename to tests/idris2/basic/dotted001/expected diff --git a/tests/idris2/dotted001/run b/tests/idris2/basic/dotted001/run similarity index 51% rename from tests/idris2/dotted001/run rename to tests/idris2/basic/dotted001/run index ee8de68057..623772f0a0 100755 --- a/tests/idris2/dotted001/run +++ b/tests/idris2/basic/dotted001/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 -c Issue2726.idr diff --git a/tests/idris2/idiom001/Main.idr b/tests/idris2/basic/idiom001/Main.idr similarity index 100% rename from tests/idris2/idiom001/Main.idr rename to tests/idris2/basic/idiom001/Main.idr diff --git a/tests/idris2/idiom001/expected b/tests/idris2/basic/idiom001/expected similarity index 100% rename from tests/idris2/idiom001/expected rename to tests/idris2/basic/idiom001/expected diff --git a/tests/idris2/basic/idiom001/run b/tests/idris2/basic/idiom001/run new file mode 100644 index 0000000000..2f1eae173f --- /dev/null +++ b/tests/idris2/basic/idiom001/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +run Main.idr diff --git a/tests/idris2/interpolation001/IfThenElse.idr b/tests/idris2/basic/interpolation001/IfThenElse.idr similarity index 100% rename from tests/idris2/interpolation001/IfThenElse.idr rename to tests/idris2/basic/interpolation001/IfThenElse.idr diff --git a/tests/idris2/interpolation001/expected b/tests/idris2/basic/interpolation001/expected similarity index 100% rename from tests/idris2/interpolation001/expected rename to tests/idris2/basic/interpolation001/expected diff --git a/tests/idris2/interpolation001/run b/tests/idris2/basic/interpolation001/run similarity index 64% rename from tests/idris2/interpolation001/run rename to tests/idris2/basic/interpolation001/run index bf0b168bef..595ea6c688 100755 --- a/tests/idris2/interpolation001/run +++ b/tests/idris2/basic/interpolation001/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh check --alt-error-count 10 IfThenElse.idr diff --git a/tests/idris2/interpolation002/StringLiteral.idr b/tests/idris2/basic/interpolation002/StringLiteral.idr similarity index 100% rename from tests/idris2/interpolation002/StringLiteral.idr rename to tests/idris2/basic/interpolation002/StringLiteral.idr diff --git a/tests/idris2/interpolation002/expected b/tests/idris2/basic/interpolation002/expected similarity index 100% rename from tests/idris2/interpolation002/expected rename to tests/idris2/basic/interpolation002/expected diff --git a/tests/idris2/interpolation002/run b/tests/idris2/basic/interpolation002/run similarity index 61% rename from tests/idris2/interpolation002/run rename to tests/idris2/basic/interpolation002/run index fa944260bc..912a585bc7 100755 --- a/tests/idris2/interpolation002/run +++ b/tests/idris2/basic/interpolation002/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 --exec test StringLiteral.idr diff --git a/tests/idris2/interpolation003/Test.idr b/tests/idris2/basic/interpolation003/Test.idr similarity index 100% rename from tests/idris2/interpolation003/Test.idr rename to tests/idris2/basic/interpolation003/Test.idr diff --git a/tests/idris2/interpolation003/expected b/tests/idris2/basic/interpolation003/expected similarity index 100% rename from tests/idris2/interpolation003/expected rename to tests/idris2/basic/interpolation003/expected diff --git a/tests/idris2/interpolation003/input b/tests/idris2/basic/interpolation003/input similarity index 100% rename from tests/idris2/interpolation003/input rename to tests/idris2/basic/interpolation003/input diff --git a/tests/idris2/builtin001/run b/tests/idris2/basic/interpolation003/run similarity index 51% rename from tests/idris2/builtin001/run rename to tests/idris2/basic/interpolation003/run index 441b64ba30..c7738ebe4e 100755 --- a/tests/idris2/builtin001/run +++ b/tests/idris2/basic/interpolation003/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 Test.idr < input diff --git a/tests/idris2/interpolation004/StringLiteral.idr b/tests/idris2/basic/interpolation004/StringLiteral.idr similarity index 100% rename from tests/idris2/interpolation004/StringLiteral.idr rename to tests/idris2/basic/interpolation004/StringLiteral.idr diff --git a/tests/idris2/interpolation004/expected b/tests/idris2/basic/interpolation004/expected similarity index 100% rename from tests/idris2/interpolation004/expected rename to tests/idris2/basic/interpolation004/expected diff --git a/tests/idris2/interpolation004/run b/tests/idris2/basic/interpolation004/run similarity index 51% rename from tests/idris2/interpolation004/run rename to tests/idris2/basic/interpolation004/run index b66578cdae..45c55a9aa6 100755 --- a/tests/idris2/interpolation004/run +++ b/tests/idris2/basic/interpolation004/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh check StringLiteral.idr diff --git a/tests/idris2/literals001/Test.idr b/tests/idris2/basic/literals001/Test.idr similarity index 100% rename from tests/idris2/literals001/Test.idr rename to tests/idris2/basic/literals001/Test.idr diff --git a/tests/idris2/literals001/expected b/tests/idris2/basic/literals001/expected similarity index 100% rename from tests/idris2/literals001/expected rename to tests/idris2/basic/literals001/expected diff --git a/tests/idris2/literals001/run b/tests/idris2/basic/literals001/run similarity index 100% rename from tests/idris2/literals001/run rename to tests/idris2/basic/literals001/run diff --git a/tests/idris2/rewrite001/Issue2573.idr b/tests/idris2/basic/rewrite001/Issue2573.idr similarity index 100% rename from tests/idris2/rewrite001/Issue2573.idr rename to tests/idris2/basic/rewrite001/Issue2573.idr diff --git a/tests/idris2/rewrite001/expected b/tests/idris2/basic/rewrite001/expected similarity index 100% rename from tests/idris2/rewrite001/expected rename to tests/idris2/basic/rewrite001/expected diff --git a/tests/idris2/basic/rewrite001/run b/tests/idris2/basic/rewrite001/run new file mode 100755 index 0000000000..c5e5307200 --- /dev/null +++ b/tests/idris2/basic/rewrite001/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Issue2573.idr diff --git a/tests/idris2/basic012/run b/tests/idris2/basic012/run deleted file mode 100755 index dd6ccaa67b..0000000000 --- a/tests/idris2/basic012/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check VIndex.idr diff --git a/tests/idris2/basic013/run b/tests/idris2/basic013/run deleted file mode 100755 index 34847877d7..0000000000 --- a/tests/idris2/basic013/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check Implicits.idr diff --git a/tests/idris2/basic014/run b/tests/idris2/basic014/run deleted file mode 100755 index 90447fd424..0000000000 --- a/tests/idris2/basic014/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check Rewrite.idr diff --git a/tests/idris2/basic015/run b/tests/idris2/basic015/run deleted file mode 100755 index d9f9e59d4e..0000000000 --- a/tests/idris2/basic015/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check George.idr diff --git a/tests/idris2/basic018/run b/tests/idris2/basic018/run deleted file mode 100755 index 8b2ca2341f..0000000000 --- a/tests/idris2/basic018/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check Fin.idr diff --git a/tests/idris2/basic025/run b/tests/idris2/basic025/run deleted file mode 100755 index fecfb4b107..0000000000 --- a/tests/idris2/basic025/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -idris2 < input diff --git a/tests/idris2/basic026/run b/tests/idris2/basic026/run deleted file mode 100755 index 6f656217d9..0000000000 --- a/tests/idris2/basic026/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check Erl.idr diff --git a/tests/idris2/basic027/run b/tests/idris2/basic027/run deleted file mode 100755 index 823be3e2f8..0000000000 --- a/tests/idris2/basic027/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check Temp.idr diff --git a/tests/idris2/basic030/run b/tests/idris2/basic030/run deleted file mode 100644 index 2ce64ef3af..0000000000 --- a/tests/idris2/basic030/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check arity.idr diff --git a/tests/idris2/basic031/run b/tests/idris2/basic031/run deleted file mode 100644 index c5cd987fcd..0000000000 --- a/tests/idris2/basic031/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check erased.idr diff --git a/tests/idris2/basic033/run b/tests/idris2/basic033/run deleted file mode 100644 index 9480bdeaa8..0000000000 --- a/tests/idris2/basic033/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check unboundimps.idr diff --git a/tests/idris2/basic034/run b/tests/idris2/basic034/run deleted file mode 100644 index 8fc7532fdb..0000000000 --- a/tests/idris2/basic034/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check lets.idr diff --git a/tests/idris2/basic041/run b/tests/idris2/basic041/run deleted file mode 100755 index c8d35dd935..0000000000 --- a/tests/idris2/basic041/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check QDo.idr diff --git a/tests/idris2/basic045/run b/tests/idris2/basic045/run deleted file mode 100644 index 64c80a5b21..0000000000 --- a/tests/idris2/basic045/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check Main.idr diff --git a/tests/idris2/basic051/run b/tests/idris2/basic051/run deleted file mode 100755 index 83c4419b5d..0000000000 --- a/tests/idris2/basic051/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check Issue833.idr diff --git a/tests/idris2/basic054/run b/tests/idris2/basic054/run deleted file mode 100755 index 5c11709ceb..0000000000 --- a/tests/idris2/basic054/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -run Issue1023.idr diff --git a/tests/idris2/basic055/run b/tests/idris2/basic055/run deleted file mode 100644 index 5bb842ecfb..0000000000 --- a/tests/idris2/basic055/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -run BitOps.idr diff --git a/tests/idris2/basic056/run b/tests/idris2/basic056/run deleted file mode 100644 index f64488f035..0000000000 --- a/tests/idris2/basic056/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -run DoubleLit.idr diff --git a/tests/idris2/basic057/run b/tests/idris2/basic057/run deleted file mode 100644 index eeceab8184..0000000000 --- a/tests/idris2/basic057/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check LetIn.idr diff --git a/tests/idris2/basic059/run b/tests/idris2/basic059/run deleted file mode 100644 index b2bc2105c0..0000000000 --- a/tests/idris2/basic059/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check MultiClaim.idr diff --git a/tests/idris2/basic062/run b/tests/idris2/basic062/run deleted file mode 100755 index f54b91bf3a..0000000000 --- a/tests/idris2/basic062/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check Issue1943.idr diff --git a/tests/idris2/basic065/run b/tests/idris2/basic065/run deleted file mode 100755 index b8a7cd495b..0000000000 --- a/tests/idris2/basic065/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check Issue215.idr diff --git a/tests/idris2/basic066/run b/tests/idris2/basic066/run deleted file mode 100755 index 478fd5723f..0000000000 --- a/tests/idris2/basic066/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check comment.idr diff --git a/tests/idris2/basic068/run b/tests/idris2/basic068/run deleted file mode 100755 index dc975e3814..0000000000 --- a/tests/idris2/basic068/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -run Issue2138.idr diff --git a/tests/idris2/basic069/run b/tests/idris2/basic069/run deleted file mode 100755 index d23326c00f..0000000000 --- a/tests/idris2/basic069/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -run DebugInfo.idr diff --git a/tests/idris2/builtin001/Test.idr b/tests/idris2/builtin/builtin001/Test.idr similarity index 100% rename from tests/idris2/builtin001/Test.idr rename to tests/idris2/builtin/builtin001/Test.idr diff --git a/tests/idris2/builtin001/expected b/tests/idris2/builtin/builtin001/expected similarity index 100% rename from tests/idris2/builtin001/expected rename to tests/idris2/builtin/builtin001/expected diff --git a/tests/idris2/builtin001/input b/tests/idris2/builtin/builtin001/input similarity index 100% rename from tests/idris2/builtin001/input rename to tests/idris2/builtin/builtin001/input diff --git a/tests/idris2/builtin002/run b/tests/idris2/builtin/builtin001/run similarity index 51% rename from tests/idris2/builtin002/run rename to tests/idris2/builtin/builtin001/run index 441b64ba30..c7738ebe4e 100755 --- a/tests/idris2/builtin002/run +++ b/tests/idris2/builtin/builtin001/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 Test.idr < input diff --git a/tests/idris2/builtin002/Test.idr b/tests/idris2/builtin/builtin002/Test.idr similarity index 100% rename from tests/idris2/builtin002/Test.idr rename to tests/idris2/builtin/builtin002/Test.idr diff --git a/tests/idris2/builtin002/expected b/tests/idris2/builtin/builtin002/expected similarity index 100% rename from tests/idris2/builtin002/expected rename to tests/idris2/builtin/builtin002/expected diff --git a/tests/idris2/builtin002/input b/tests/idris2/builtin/builtin002/input similarity index 100% rename from tests/idris2/builtin002/input rename to tests/idris2/builtin/builtin002/input diff --git a/tests/idris2/builtin003/run b/tests/idris2/builtin/builtin002/run similarity index 51% rename from tests/idris2/builtin003/run rename to tests/idris2/builtin/builtin002/run index 441b64ba30..c7738ebe4e 100755 --- a/tests/idris2/builtin003/run +++ b/tests/idris2/builtin/builtin002/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 Test.idr < input diff --git a/tests/idris2/builtin003/Test.idr b/tests/idris2/builtin/builtin003/Test.idr similarity index 100% rename from tests/idris2/builtin003/Test.idr rename to tests/idris2/builtin/builtin003/Test.idr diff --git a/tests/idris2/builtin003/expected b/tests/idris2/builtin/builtin003/expected similarity index 100% rename from tests/idris2/builtin003/expected rename to tests/idris2/builtin/builtin003/expected diff --git a/tests/idris2/builtin003/input b/tests/idris2/builtin/builtin003/input similarity index 100% rename from tests/idris2/builtin003/input rename to tests/idris2/builtin/builtin003/input diff --git a/tests/idris2/builtin004/run b/tests/idris2/builtin/builtin003/run similarity index 51% rename from tests/idris2/builtin004/run rename to tests/idris2/builtin/builtin003/run index 441b64ba30..c7738ebe4e 100755 --- a/tests/idris2/builtin004/run +++ b/tests/idris2/builtin/builtin003/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 Test.idr < input diff --git a/tests/idris2/builtin004/Test.idr b/tests/idris2/builtin/builtin004/Test.idr similarity index 100% rename from tests/idris2/builtin004/Test.idr rename to tests/idris2/builtin/builtin004/Test.idr diff --git a/tests/idris2/builtin004/expected b/tests/idris2/builtin/builtin004/expected similarity index 100% rename from tests/idris2/builtin004/expected rename to tests/idris2/builtin/builtin004/expected diff --git a/tests/idris2/builtin004/input b/tests/idris2/builtin/builtin004/input similarity index 100% rename from tests/idris2/builtin004/input rename to tests/idris2/builtin/builtin004/input diff --git a/tests/idris2/builtin/builtin004/run b/tests/idris2/builtin/builtin004/run new file mode 100755 index 0000000000..c7738ebe4e --- /dev/null +++ b/tests/idris2/builtin/builtin004/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +idris2 Test.idr < input diff --git a/tests/idris2/builtin005/Test.idr b/tests/idris2/builtin/builtin005/Test.idr similarity index 100% rename from tests/idris2/builtin005/Test.idr rename to tests/idris2/builtin/builtin005/Test.idr diff --git a/tests/idris2/builtin005/expected b/tests/idris2/builtin/builtin005/expected similarity index 100% rename from tests/idris2/builtin005/expected rename to tests/idris2/builtin/builtin005/expected diff --git a/tests/idris2/builtin005/input b/tests/idris2/builtin/builtin005/input similarity index 100% rename from tests/idris2/builtin005/input rename to tests/idris2/builtin/builtin005/input diff --git a/tests/idris2/builtin/builtin005/run b/tests/idris2/builtin/builtin005/run new file mode 100755 index 0000000000..c7738ebe4e --- /dev/null +++ b/tests/idris2/builtin/builtin005/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +idris2 Test.idr < input diff --git a/tests/idris2/builtin006/Test.idr b/tests/idris2/builtin/builtin006/Test.idr similarity index 100% rename from tests/idris2/builtin006/Test.idr rename to tests/idris2/builtin/builtin006/Test.idr diff --git a/tests/idris2/builtin006/expected b/tests/idris2/builtin/builtin006/expected similarity index 100% rename from tests/idris2/builtin006/expected rename to tests/idris2/builtin/builtin006/expected diff --git a/tests/idris2/builtin006/input b/tests/idris2/builtin/builtin006/input similarity index 100% rename from tests/idris2/builtin006/input rename to tests/idris2/builtin/builtin006/input diff --git a/tests/idris2/builtin/builtin006/run b/tests/idris2/builtin/builtin006/run new file mode 100755 index 0000000000..c7738ebe4e --- /dev/null +++ b/tests/idris2/builtin/builtin006/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +idris2 Test.idr < input diff --git a/tests/idris2/builtin007/Test.idr b/tests/idris2/builtin/builtin007/Test.idr similarity index 100% rename from tests/idris2/builtin007/Test.idr rename to tests/idris2/builtin/builtin007/Test.idr diff --git a/tests/idris2/builtin007/expected b/tests/idris2/builtin/builtin007/expected similarity index 100% rename from tests/idris2/builtin007/expected rename to tests/idris2/builtin/builtin007/expected diff --git a/tests/idris2/builtin007/input b/tests/idris2/builtin/builtin007/input similarity index 100% rename from tests/idris2/builtin007/input rename to tests/idris2/builtin/builtin007/input diff --git a/tests/idris2/builtin/builtin007/run b/tests/idris2/builtin/builtin007/run new file mode 100755 index 0000000000..c7738ebe4e --- /dev/null +++ b/tests/idris2/builtin/builtin007/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +idris2 Test.idr < input diff --git a/tests/idris2/builtin008/Test.idr b/tests/idris2/builtin/builtin008/Test.idr similarity index 100% rename from tests/idris2/builtin008/Test.idr rename to tests/idris2/builtin/builtin008/Test.idr diff --git a/tests/idris2/builtin008/expected b/tests/idris2/builtin/builtin008/expected similarity index 100% rename from tests/idris2/builtin008/expected rename to tests/idris2/builtin/builtin008/expected diff --git a/tests/idris2/builtin008/input b/tests/idris2/builtin/builtin008/input similarity index 100% rename from tests/idris2/builtin008/input rename to tests/idris2/builtin/builtin008/input diff --git a/tests/idris2/builtin/builtin008/run b/tests/idris2/builtin/builtin008/run new file mode 100755 index 0000000000..c7738ebe4e --- /dev/null +++ b/tests/idris2/builtin/builtin008/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +idris2 Test.idr < input diff --git a/tests/idris2/builtin009/Test.idr b/tests/idris2/builtin/builtin009/Test.idr similarity index 100% rename from tests/idris2/builtin009/Test.idr rename to tests/idris2/builtin/builtin009/Test.idr diff --git a/tests/idris2/builtin009/expected b/tests/idris2/builtin/builtin009/expected similarity index 100% rename from tests/idris2/builtin009/expected rename to tests/idris2/builtin/builtin009/expected diff --git a/tests/idris2/builtin/builtin009/run b/tests/idris2/builtin/builtin009/run new file mode 100755 index 0000000000..b687e5259c --- /dev/null +++ b/tests/idris2/builtin/builtin009/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +run Test.idr diff --git a/tests/idris2/builtin010/Test.idr b/tests/idris2/builtin/builtin010/Test.idr similarity index 100% rename from tests/idris2/builtin010/Test.idr rename to tests/idris2/builtin/builtin010/Test.idr diff --git a/tests/idris2/builtin010/expected b/tests/idris2/builtin/builtin010/expected similarity index 100% rename from tests/idris2/builtin010/expected rename to tests/idris2/builtin/builtin010/expected diff --git a/tests/idris2/builtin010/input b/tests/idris2/builtin/builtin010/input similarity index 100% rename from tests/idris2/builtin010/input rename to tests/idris2/builtin/builtin010/input diff --git a/tests/idris2/builtin/builtin010/run b/tests/idris2/builtin/builtin010/run new file mode 100755 index 0000000000..c7738ebe4e --- /dev/null +++ b/tests/idris2/builtin/builtin010/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +idris2 Test.idr < input diff --git a/tests/idris2/builtin011/Test.idr b/tests/idris2/builtin/builtin011/Test.idr similarity index 100% rename from tests/idris2/builtin011/Test.idr rename to tests/idris2/builtin/builtin011/Test.idr diff --git a/tests/idris2/builtin011/expected b/tests/idris2/builtin/builtin011/expected similarity index 100% rename from tests/idris2/builtin011/expected rename to tests/idris2/builtin/builtin011/expected diff --git a/tests/idris2/builtin/builtin011/run b/tests/idris2/builtin/builtin011/run new file mode 100755 index 0000000000..b687e5259c --- /dev/null +++ b/tests/idris2/builtin/builtin011/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +run Test.idr diff --git a/tests/idris2/builtin012/Issue1799.idr b/tests/idris2/builtin/builtin012/Issue1799.idr similarity index 100% rename from tests/idris2/builtin012/Issue1799.idr rename to tests/idris2/builtin/builtin012/Issue1799.idr diff --git a/tests/idris2/builtin012/expected b/tests/idris2/builtin/builtin012/expected similarity index 100% rename from tests/idris2/builtin012/expected rename to tests/idris2/builtin/builtin012/expected diff --git a/tests/idris2/builtin/builtin012/run b/tests/idris2/builtin/builtin012/run new file mode 100755 index 0000000000..20853b9720 --- /dev/null +++ b/tests/idris2/builtin/builtin012/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +run Issue1799.idr diff --git a/tests/idris2/builtin005/run b/tests/idris2/builtin005/run deleted file mode 100755 index 441b64ba30..0000000000 --- a/tests/idris2/builtin005/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -idris2 Test.idr < input diff --git a/tests/idris2/builtin006/run b/tests/idris2/builtin006/run deleted file mode 100755 index 441b64ba30..0000000000 --- a/tests/idris2/builtin006/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -idris2 Test.idr < input diff --git a/tests/idris2/builtin007/run b/tests/idris2/builtin007/run deleted file mode 100755 index 441b64ba30..0000000000 --- a/tests/idris2/builtin007/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -idris2 Test.idr < input diff --git a/tests/idris2/builtin008/run b/tests/idris2/builtin008/run deleted file mode 100755 index 441b64ba30..0000000000 --- a/tests/idris2/builtin008/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -idris2 Test.idr < input diff --git a/tests/idris2/builtin009/run b/tests/idris2/builtin009/run deleted file mode 100755 index ceb556ba04..0000000000 --- a/tests/idris2/builtin009/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -run Test.idr diff --git a/tests/idris2/builtin010/run b/tests/idris2/builtin010/run deleted file mode 100755 index 441b64ba30..0000000000 --- a/tests/idris2/builtin010/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -idris2 Test.idr < input diff --git a/tests/idris2/builtin011/run b/tests/idris2/builtin011/run deleted file mode 100755 index ceb556ba04..0000000000 --- a/tests/idris2/builtin011/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -run Test.idr diff --git a/tests/idris2/builtin012/run b/tests/idris2/builtin012/run deleted file mode 100755 index 192255940d..0000000000 --- a/tests/idris2/builtin012/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -run Issue1799.idr diff --git a/tests/idris2/casetree001/IsS.idr b/tests/idris2/casetree/casetree001/IsS.idr similarity index 100% rename from tests/idris2/casetree001/IsS.idr rename to tests/idris2/casetree/casetree001/IsS.idr diff --git a/tests/idris2/casetree001/Issue762.idr b/tests/idris2/casetree/casetree001/Issue762.idr similarity index 100% rename from tests/idris2/casetree001/Issue762.idr rename to tests/idris2/casetree/casetree001/Issue762.idr diff --git a/tests/idris2/casetree001/expected b/tests/idris2/casetree/casetree001/expected similarity index 100% rename from tests/idris2/casetree001/expected rename to tests/idris2/casetree/casetree001/expected diff --git a/tests/idris2/casetree001/run b/tests/idris2/casetree/casetree001/run similarity index 58% rename from tests/idris2/casetree001/run rename to tests/idris2/casetree/casetree001/run index 4d98e0f914..5a777a4f80 100755 --- a/tests/idris2/casetree001/run +++ b/tests/idris2/casetree/casetree001/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh check Issue762.idr check IsS.idr diff --git a/tests/idris2/casetree002/DefaultCases.idr b/tests/idris2/casetree/casetree002/DefaultCases.idr similarity index 100% rename from tests/idris2/casetree002/DefaultCases.idr rename to tests/idris2/casetree/casetree002/DefaultCases.idr diff --git a/tests/idris2/casetree002/Issue1079.idr b/tests/idris2/casetree/casetree002/Issue1079.idr similarity index 100% rename from tests/idris2/casetree002/Issue1079.idr rename to tests/idris2/casetree/casetree002/Issue1079.idr diff --git a/tests/idris2/casetree002/expected b/tests/idris2/casetree/casetree002/expected similarity index 100% rename from tests/idris2/casetree002/expected rename to tests/idris2/casetree/casetree002/expected diff --git a/tests/idris2/casetree002/run b/tests/idris2/casetree/casetree002/run similarity index 64% rename from tests/idris2/casetree002/run rename to tests/idris2/casetree/casetree002/run index db84ac67fb..7616491b5a 100755 --- a/tests/idris2/casetree002/run +++ b/tests/idris2/casetree/casetree002/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh check DefaultCases.idr check Issue1079.idr diff --git a/tests/idris2/casetree003/ForcedPats.idr b/tests/idris2/casetree/casetree003/ForcedPats.idr similarity index 100% rename from tests/idris2/casetree003/ForcedPats.idr rename to tests/idris2/casetree/casetree003/ForcedPats.idr diff --git a/tests/idris2/casetree003/expected b/tests/idris2/casetree/casetree003/expected similarity index 100% rename from tests/idris2/casetree003/expected rename to tests/idris2/casetree/casetree003/expected diff --git a/tests/idris2/casetree003/input b/tests/idris2/casetree/casetree003/input similarity index 100% rename from tests/idris2/casetree003/input rename to tests/idris2/casetree/casetree003/input diff --git a/tests/idris2/casetree003/run b/tests/idris2/casetree/casetree003/run similarity index 56% rename from tests/idris2/casetree003/run rename to tests/idris2/casetree/casetree003/run index 75b9db51d2..97b9bb68bd 100755 --- a/tests/idris2/casetree003/run +++ b/tests/idris2/casetree/casetree003/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 ForcedPats.idr < input diff --git a/tests/idris2/casetree004/LocalArgs.idr b/tests/idris2/casetree/casetree004/LocalArgs.idr similarity index 100% rename from tests/idris2/casetree004/LocalArgs.idr rename to tests/idris2/casetree/casetree004/LocalArgs.idr diff --git a/tests/idris2/casetree004/PiMatch.idr b/tests/idris2/casetree/casetree004/PiMatch.idr similarity index 100% rename from tests/idris2/casetree004/PiMatch.idr rename to tests/idris2/casetree/casetree004/PiMatch.idr diff --git a/tests/idris2/casetree004/expected b/tests/idris2/casetree/casetree004/expected similarity index 100% rename from tests/idris2/casetree004/expected rename to tests/idris2/casetree/casetree004/expected diff --git a/tests/idris2/casetree004/run b/tests/idris2/casetree/casetree004/run similarity index 61% rename from tests/idris2/casetree004/run rename to tests/idris2/casetree/casetree004/run index 1c7e49e8ef..ffa3badfe7 100755 --- a/tests/idris2/casetree004/run +++ b/tests/idris2/casetree/casetree004/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh check LocalArgs.idr check PiMatch.idr diff --git a/tests/idris2/coverage001/Vect.idr b/tests/idris2/coverage/coverage001/Vect.idr similarity index 100% rename from tests/idris2/coverage001/Vect.idr rename to tests/idris2/coverage/coverage001/Vect.idr diff --git a/tests/idris2/coverage001/Vect2.idr b/tests/idris2/coverage/coverage001/Vect2.idr similarity index 100% rename from tests/idris2/coverage001/Vect2.idr rename to tests/idris2/coverage/coverage001/Vect2.idr diff --git a/tests/idris2/coverage001/Vect3.idr b/tests/idris2/coverage/coverage001/Vect3.idr similarity index 100% rename from tests/idris2/coverage001/Vect3.idr rename to tests/idris2/coverage/coverage001/Vect3.idr diff --git a/tests/idris2/coverage001/expected b/tests/idris2/coverage/coverage001/expected similarity index 100% rename from tests/idris2/coverage001/expected rename to tests/idris2/coverage/coverage001/expected diff --git a/tests/idris2/coverage001/input b/tests/idris2/coverage/coverage001/input similarity index 100% rename from tests/idris2/coverage001/input rename to tests/idris2/coverage/coverage001/input diff --git a/tests/idris2/coverage001/run b/tests/idris2/coverage/coverage001/run similarity index 73% rename from tests/idris2/coverage001/run rename to tests/idris2/coverage/coverage001/run index 57aa3d817f..d08e6699ff 100755 --- a/tests/idris2/coverage001/run +++ b/tests/idris2/coverage/coverage001/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 Vect.idr < input idris2 -c Vect2.idr diff --git a/tests/idris2/coverage002/Vect.idr b/tests/idris2/coverage/coverage002/Vect.idr similarity index 100% rename from tests/idris2/coverage002/Vect.idr rename to tests/idris2/coverage/coverage002/Vect.idr diff --git a/tests/idris2/coverage002/expected b/tests/idris2/coverage/coverage002/expected similarity index 100% rename from tests/idris2/coverage002/expected rename to tests/idris2/coverage/coverage002/expected diff --git a/tests/idris2/coverage002/input b/tests/idris2/coverage/coverage002/input similarity index 100% rename from tests/idris2/coverage002/input rename to tests/idris2/coverage/coverage002/input diff --git a/tests/idris2/coverage002/run b/tests/idris2/coverage/coverage002/run similarity index 51% rename from tests/idris2/coverage002/run rename to tests/idris2/coverage/coverage002/run index fcf96b59e0..1f7ddd7109 100755 --- a/tests/idris2/coverage002/run +++ b/tests/idris2/coverage/coverage002/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 Vect.idr < input diff --git a/tests/idris2/coverage003/Cover.idr b/tests/idris2/coverage/coverage003/Cover.idr similarity index 100% rename from tests/idris2/coverage003/Cover.idr rename to tests/idris2/coverage/coverage003/Cover.idr diff --git a/tests/idris2/coverage003/expected b/tests/idris2/coverage/coverage003/expected similarity index 100% rename from tests/idris2/coverage003/expected rename to tests/idris2/coverage/coverage003/expected diff --git a/tests/idris2/coverage003/input b/tests/idris2/coverage/coverage003/input similarity index 100% rename from tests/idris2/coverage003/input rename to tests/idris2/coverage/coverage003/input diff --git a/tests/idris2/coverage003/run b/tests/idris2/coverage/coverage003/run similarity index 52% rename from tests/idris2/coverage003/run rename to tests/idris2/coverage/coverage003/run index b64bc24733..97367fc5fc 100755 --- a/tests/idris2/coverage003/run +++ b/tests/idris2/coverage/coverage003/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 Cover.idr < input diff --git a/tests/idris2/coverage004/Cover.idr b/tests/idris2/coverage/coverage004/Cover.idr similarity index 100% rename from tests/idris2/coverage004/Cover.idr rename to tests/idris2/coverage/coverage004/Cover.idr diff --git a/tests/idris2/coverage004/expected b/tests/idris2/coverage/coverage004/expected similarity index 100% rename from tests/idris2/coverage004/expected rename to tests/idris2/coverage/coverage004/expected diff --git a/tests/idris2/coverage004/input b/tests/idris2/coverage/coverage004/input similarity index 100% rename from tests/idris2/coverage004/input rename to tests/idris2/coverage/coverage004/input diff --git a/tests/idris2/coverage004/run b/tests/idris2/coverage/coverage004/run similarity index 52% rename from tests/idris2/coverage004/run rename to tests/idris2/coverage/coverage004/run index b64bc24733..97367fc5fc 100755 --- a/tests/idris2/coverage004/run +++ b/tests/idris2/coverage/coverage004/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 Cover.idr < input diff --git a/tests/idris2/coverage005/Cover.idr b/tests/idris2/coverage/coverage005/Cover.idr similarity index 100% rename from tests/idris2/coverage005/Cover.idr rename to tests/idris2/coverage/coverage005/Cover.idr diff --git a/tests/idris2/coverage005/expected b/tests/idris2/coverage/coverage005/expected similarity index 100% rename from tests/idris2/coverage005/expected rename to tests/idris2/coverage/coverage005/expected diff --git a/tests/idris2/coverage005/input b/tests/idris2/coverage/coverage005/input similarity index 100% rename from tests/idris2/coverage005/input rename to tests/idris2/coverage/coverage005/input diff --git a/tests/idris2/coverage005/run b/tests/idris2/coverage/coverage005/run similarity index 52% rename from tests/idris2/coverage005/run rename to tests/idris2/coverage/coverage005/run index b64bc24733..97367fc5fc 100755 --- a/tests/idris2/coverage005/run +++ b/tests/idris2/coverage/coverage005/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 Cover.idr < input diff --git a/tests/idris2/coverage006/expected b/tests/idris2/coverage/coverage006/expected similarity index 100% rename from tests/idris2/coverage006/expected rename to tests/idris2/coverage/coverage006/expected diff --git a/tests/idris2/coverage006/foobar.idr b/tests/idris2/coverage/coverage006/foobar.idr similarity index 100% rename from tests/idris2/coverage006/foobar.idr rename to tests/idris2/coverage/coverage006/foobar.idr diff --git a/tests/idris2/coverage006/input b/tests/idris2/coverage/coverage006/input similarity index 100% rename from tests/idris2/coverage006/input rename to tests/idris2/coverage/coverage006/input diff --git a/tests/idris2/coverage006/run b/tests/idris2/coverage/coverage006/run similarity index 52% rename from tests/idris2/coverage006/run rename to tests/idris2/coverage/coverage006/run index e18dcdb4dd..7176cda48c 100755 --- a/tests/idris2/coverage006/run +++ b/tests/idris2/coverage/coverage006/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 foobar.idr < input diff --git a/tests/idris2/coverage007/eq.idr b/tests/idris2/coverage/coverage007/eq.idr similarity index 100% rename from tests/idris2/coverage007/eq.idr rename to tests/idris2/coverage/coverage007/eq.idr diff --git a/tests/idris2/coverage007/expected b/tests/idris2/coverage/coverage007/expected similarity index 100% rename from tests/idris2/coverage007/expected rename to tests/idris2/coverage/coverage007/expected diff --git a/tests/idris2/coverage/coverage007/run b/tests/idris2/coverage/coverage007/run new file mode 100755 index 0000000000..53f6e1e854 --- /dev/null +++ b/tests/idris2/coverage/coverage007/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check eq.idr diff --git a/tests/idris2/coverage008/expected b/tests/idris2/coverage/coverage008/expected similarity index 100% rename from tests/idris2/coverage008/expected rename to tests/idris2/coverage/coverage008/expected diff --git a/tests/idris2/coverage008/input b/tests/idris2/coverage/coverage008/input similarity index 100% rename from tests/idris2/coverage008/input rename to tests/idris2/coverage/coverage008/input diff --git a/tests/idris2/coverage008/run b/tests/idris2/coverage/coverage008/run similarity index 51% rename from tests/idris2/coverage008/run rename to tests/idris2/coverage/coverage008/run index 94ae606197..24c5216fa1 100755 --- a/tests/idris2/coverage008/run +++ b/tests/idris2/coverage/coverage008/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 wcov.idr < input diff --git a/tests/idris2/coverage008/wcov.idr b/tests/idris2/coverage/coverage008/wcov.idr similarity index 100% rename from tests/idris2/coverage008/wcov.idr rename to tests/idris2/coverage/coverage008/wcov.idr diff --git a/tests/idris2/coverage009/expected b/tests/idris2/coverage/coverage009/expected similarity index 100% rename from tests/idris2/coverage009/expected rename to tests/idris2/coverage/coverage009/expected diff --git a/tests/idris2/coverage/coverage009/run b/tests/idris2/coverage/coverage009/run new file mode 100755 index 0000000000..734f53510c --- /dev/null +++ b/tests/idris2/coverage/coverage009/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check unreachable.idr diff --git a/tests/idris2/coverage009/unreachable.idr b/tests/idris2/coverage/coverage009/unreachable.idr similarity index 100% rename from tests/idris2/coverage009/unreachable.idr rename to tests/idris2/coverage/coverage009/unreachable.idr diff --git a/tests/idris2/coverage010/casetot.idr b/tests/idris2/coverage/coverage010/casetot.idr similarity index 100% rename from tests/idris2/coverage010/casetot.idr rename to tests/idris2/coverage/coverage010/casetot.idr diff --git a/tests/idris2/coverage010/expected b/tests/idris2/coverage/coverage010/expected similarity index 100% rename from tests/idris2/coverage010/expected rename to tests/idris2/coverage/coverage010/expected diff --git a/tests/idris2/coverage/coverage010/run b/tests/idris2/coverage/coverage010/run new file mode 100755 index 0000000000..04ead2b869 --- /dev/null +++ b/tests/idris2/coverage/coverage010/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check casetot.idr diff --git a/tests/idris2/coverage011/Sing.idr b/tests/idris2/coverage/coverage011/Sing.idr similarity index 100% rename from tests/idris2/coverage011/Sing.idr rename to tests/idris2/coverage/coverage011/Sing.idr diff --git a/tests/idris2/coverage011/expected b/tests/idris2/coverage/coverage011/expected similarity index 100% rename from tests/idris2/coverage011/expected rename to tests/idris2/coverage/coverage011/expected diff --git a/tests/idris2/coverage/coverage011/run b/tests/idris2/coverage/coverage011/run new file mode 100755 index 0000000000..faa84a9309 --- /dev/null +++ b/tests/idris2/coverage/coverage011/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Sing.idr diff --git a/tests/idris2/coverage012/Issue484.idr b/tests/idris2/coverage/coverage012/Issue484.idr similarity index 100% rename from tests/idris2/coverage012/Issue484.idr rename to tests/idris2/coverage/coverage012/Issue484.idr diff --git a/tests/idris2/coverage012/Issue899.idr b/tests/idris2/coverage/coverage012/Issue899.idr similarity index 100% rename from tests/idris2/coverage012/Issue899.idr rename to tests/idris2/coverage/coverage012/Issue899.idr diff --git a/tests/idris2/coverage012/expected b/tests/idris2/coverage/coverage012/expected similarity index 100% rename from tests/idris2/coverage012/expected rename to tests/idris2/coverage/coverage012/expected diff --git a/tests/idris2/coverage012/run b/tests/idris2/coverage/coverage012/run similarity index 61% rename from tests/idris2/coverage012/run rename to tests/idris2/coverage/coverage012/run index b73b3fcb12..1e6dbb6871 100755 --- a/tests/idris2/coverage012/run +++ b/tests/idris2/coverage/coverage012/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh check Issue899.idr check Issue484.idr diff --git a/tests/idris2/coverage013/Issue1022-Refl.idr b/tests/idris2/coverage/coverage013/Issue1022-Refl.idr similarity index 100% rename from tests/idris2/coverage013/Issue1022-Refl.idr rename to tests/idris2/coverage/coverage013/Issue1022-Refl.idr diff --git a/tests/idris2/coverage013/Issue1022.idr b/tests/idris2/coverage/coverage013/Issue1022.idr similarity index 100% rename from tests/idris2/coverage013/Issue1022.idr rename to tests/idris2/coverage/coverage013/Issue1022.idr diff --git a/tests/idris2/coverage013/expected b/tests/idris2/coverage/coverage013/expected similarity index 100% rename from tests/idris2/coverage013/expected rename to tests/idris2/coverage/coverage013/expected diff --git a/tests/idris2/coverage013/run b/tests/idris2/coverage/coverage013/run similarity index 65% rename from tests/idris2/coverage013/run rename to tests/idris2/coverage/coverage013/run index c922218e6c..56f8ac87f3 100755 --- a/tests/idris2/coverage013/run +++ b/tests/idris2/coverage/coverage013/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh check Issue1022.idr check Issue1022-Refl.idr diff --git a/tests/idris2/coverage014/Issue794.idr b/tests/idris2/coverage/coverage014/Issue794.idr similarity index 100% rename from tests/idris2/coverage014/Issue794.idr rename to tests/idris2/coverage/coverage014/Issue794.idr diff --git a/tests/idris2/coverage014/expected b/tests/idris2/coverage/coverage014/expected similarity index 100% rename from tests/idris2/coverage014/expected rename to tests/idris2/coverage/coverage014/expected diff --git a/tests/idris2/coverage/coverage014/run b/tests/idris2/coverage/coverage014/run new file mode 100755 index 0000000000..af83528582 --- /dev/null +++ b/tests/idris2/coverage/coverage014/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Issue794.idr diff --git a/tests/idris2/coverage015/Issue1169.idr b/tests/idris2/coverage/coverage015/Issue1169.idr similarity index 100% rename from tests/idris2/coverage015/Issue1169.idr rename to tests/idris2/coverage/coverage015/Issue1169.idr diff --git a/tests/idris2/coverage015/Issue1366.idr b/tests/idris2/coverage/coverage015/Issue1366.idr similarity index 100% rename from tests/idris2/coverage015/Issue1366.idr rename to tests/idris2/coverage/coverage015/Issue1366.idr diff --git a/tests/idris2/coverage015/expected b/tests/idris2/coverage/coverage015/expected similarity index 100% rename from tests/idris2/coverage015/expected rename to tests/idris2/coverage/coverage015/expected diff --git a/tests/idris2/coverage015/run b/tests/idris2/coverage/coverage015/run similarity index 63% rename from tests/idris2/coverage015/run rename to tests/idris2/coverage/coverage015/run index f9bf92bdc1..dc169aee14 100755 --- a/tests/idris2/coverage015/run +++ b/tests/idris2/coverage/coverage015/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh check Issue1169.idr check Issue1366.idr diff --git a/tests/idris2/coverage016/Issue633-2.idr b/tests/idris2/coverage/coverage016/Issue633-2.idr similarity index 100% rename from tests/idris2/coverage016/Issue633-2.idr rename to tests/idris2/coverage/coverage016/Issue633-2.idr diff --git a/tests/idris2/coverage016/Issue633.idr b/tests/idris2/coverage/coverage016/Issue633.idr similarity index 100% rename from tests/idris2/coverage016/Issue633.idr rename to tests/idris2/coverage/coverage016/Issue633.idr diff --git a/tests/idris2/coverage016/expected b/tests/idris2/coverage/coverage016/expected similarity index 100% rename from tests/idris2/coverage016/expected rename to tests/idris2/coverage/coverage016/expected diff --git a/tests/idris2/coverage016/run b/tests/idris2/coverage/coverage016/run similarity index 63% rename from tests/idris2/coverage016/run rename to tests/idris2/coverage/coverage016/run index 8298d0c4b0..39acb7194c 100755 --- a/tests/idris2/coverage016/run +++ b/tests/idris2/coverage/coverage016/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh check Issue633.idr check Issue633-2.idr diff --git a/tests/idris2/coverage017/Issue1421.idr b/tests/idris2/coverage/coverage017/Issue1421.idr similarity index 100% rename from tests/idris2/coverage017/Issue1421.idr rename to tests/idris2/coverage/coverage017/Issue1421.idr diff --git a/tests/idris2/coverage017/expected b/tests/idris2/coverage/coverage017/expected similarity index 100% rename from tests/idris2/coverage017/expected rename to tests/idris2/coverage/coverage017/expected diff --git a/tests/idris2/coverage/coverage017/run b/tests/idris2/coverage/coverage017/run new file mode 100755 index 0000000000..13d7cacc45 --- /dev/null +++ b/tests/idris2/coverage/coverage017/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Issue1421.idr diff --git a/tests/idris2/coverage018/Issue1831_1.idr b/tests/idris2/coverage/coverage018/Issue1831_1.idr similarity index 100% rename from tests/idris2/coverage018/Issue1831_1.idr rename to tests/idris2/coverage/coverage018/Issue1831_1.idr diff --git a/tests/idris2/coverage018/Issue1831_2.idr b/tests/idris2/coverage/coverage018/Issue1831_2.idr similarity index 100% rename from tests/idris2/coverage018/Issue1831_2.idr rename to tests/idris2/coverage/coverage018/Issue1831_2.idr diff --git a/tests/idris2/coverage018/expected b/tests/idris2/coverage/coverage018/expected similarity index 100% rename from tests/idris2/coverage018/expected rename to tests/idris2/coverage/coverage018/expected diff --git a/tests/idris2/coverage018/input b/tests/idris2/coverage/coverage018/input similarity index 100% rename from tests/idris2/coverage018/input rename to tests/idris2/coverage/coverage018/input diff --git a/tests/idris2/coverage018/run b/tests/idris2/coverage/coverage018/run similarity index 72% rename from tests/idris2/coverage018/run rename to tests/idris2/coverage/coverage018/run index 01c5ae7cb9..5c2e37570f 100755 --- a/tests/idris2/coverage018/run +++ b/tests/idris2/coverage/coverage018/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 Issue1831_1.idr < input idris2 Issue1831_2.idr < input diff --git a/tests/idris2/coverage019/Issue1632.idr b/tests/idris2/coverage/coverage019/Issue1632.idr similarity index 100% rename from tests/idris2/coverage019/Issue1632.idr rename to tests/idris2/coverage/coverage019/Issue1632.idr diff --git a/tests/idris2/coverage019/expected b/tests/idris2/coverage/coverage019/expected similarity index 100% rename from tests/idris2/coverage019/expected rename to tests/idris2/coverage/coverage019/expected diff --git a/tests/idris2/coverage/coverage019/run b/tests/idris2/coverage/coverage019/run new file mode 100755 index 0000000000..629a14ac6f --- /dev/null +++ b/tests/idris2/coverage/coverage019/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Issue1632.idr diff --git a/tests/idris2/coverage007/run b/tests/idris2/coverage007/run deleted file mode 100755 index a27943b96b..0000000000 --- a/tests/idris2/coverage007/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check eq.idr diff --git a/tests/idris2/coverage009/run b/tests/idris2/coverage009/run deleted file mode 100755 index 6628258da8..0000000000 --- a/tests/idris2/coverage009/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check unreachable.idr diff --git a/tests/idris2/coverage010/run b/tests/idris2/coverage010/run deleted file mode 100755 index 4072962b12..0000000000 --- a/tests/idris2/coverage010/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check casetot.idr diff --git a/tests/idris2/coverage011/run b/tests/idris2/coverage011/run deleted file mode 100755 index 6558cbca37..0000000000 --- a/tests/idris2/coverage011/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check Sing.idr diff --git a/tests/idris2/coverage014/run b/tests/idris2/coverage014/run deleted file mode 100755 index 8bf79ef26b..0000000000 --- a/tests/idris2/coverage014/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check Issue794.idr diff --git a/tests/idris2/coverage017/run b/tests/idris2/coverage017/run deleted file mode 100755 index 7ce70e8b4c..0000000000 --- a/tests/idris2/coverage017/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check Issue1421.idr diff --git a/tests/idris2/coverage019/run b/tests/idris2/coverage019/run deleted file mode 100755 index 4be2ab7098..0000000000 --- a/tests/idris2/coverage019/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check Issue1632.idr diff --git a/tests/idris2/data001/Test.idr b/tests/idris2/data/data001/Test.idr similarity index 100% rename from tests/idris2/data001/Test.idr rename to tests/idris2/data/data001/Test.idr diff --git a/tests/idris2/data001/TestImpl.idr b/tests/idris2/data/data001/TestImpl.idr similarity index 100% rename from tests/idris2/data001/TestImpl.idr rename to tests/idris2/data/data001/TestImpl.idr diff --git a/tests/idris2/data001/expected b/tests/idris2/data/data001/expected similarity index 100% rename from tests/idris2/data001/expected rename to tests/idris2/data/data001/expected diff --git a/tests/idris2/data/data001/run b/tests/idris2/data/data001/run new file mode 100755 index 0000000000..da953a6fdf --- /dev/null +++ b/tests/idris2/data/data001/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check TestImpl.idr diff --git a/tests/idris2/data002/Test.idr b/tests/idris2/data/data002/Test.idr similarity index 100% rename from tests/idris2/data002/Test.idr rename to tests/idris2/data/data002/Test.idr diff --git a/tests/idris2/data002/expected b/tests/idris2/data/data002/expected similarity index 100% rename from tests/idris2/data002/expected rename to tests/idris2/data/data002/expected diff --git a/tests/idris2/data002/input b/tests/idris2/data/data002/input similarity index 100% rename from tests/idris2/data002/input rename to tests/idris2/data/data002/input diff --git a/tests/idris2/data/data002/run b/tests/idris2/data/data002/run new file mode 100755 index 0000000000..c7738ebe4e --- /dev/null +++ b/tests/idris2/data/data002/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +idris2 Test.idr < input diff --git a/tests/idris2/record001/Record.idr b/tests/idris2/data/record001/Record.idr similarity index 100% rename from tests/idris2/record001/Record.idr rename to tests/idris2/data/record001/Record.idr diff --git a/tests/idris2/record001/expected b/tests/idris2/data/record001/expected similarity index 100% rename from tests/idris2/record001/expected rename to tests/idris2/data/record001/expected diff --git a/tests/idris2/record001/input b/tests/idris2/data/record001/input similarity index 100% rename from tests/idris2/record001/input rename to tests/idris2/data/record001/input diff --git a/tests/idris2/record001/run b/tests/idris2/data/record001/run similarity index 52% rename from tests/idris2/record001/run rename to tests/idris2/data/record001/run index 62e148929e..0ea59c9d45 100755 --- a/tests/idris2/record001/run +++ b/tests/idris2/data/record001/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 Record.idr < input diff --git a/tests/idris2/record002/Record.idr b/tests/idris2/data/record002/Record.idr similarity index 100% rename from tests/idris2/record002/Record.idr rename to tests/idris2/data/record002/Record.idr diff --git a/tests/idris2/record002/expected b/tests/idris2/data/record002/expected similarity index 100% rename from tests/idris2/record002/expected rename to tests/idris2/data/record002/expected diff --git a/tests/idris2/record002/input b/tests/idris2/data/record002/input similarity index 100% rename from tests/idris2/record002/input rename to tests/idris2/data/record002/input diff --git a/tests/idris2/record002/run b/tests/idris2/data/record002/run similarity index 52% rename from tests/idris2/record002/run rename to tests/idris2/data/record002/run index 62e148929e..0ea59c9d45 100755 --- a/tests/idris2/record002/run +++ b/tests/idris2/data/record002/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 Record.idr < input diff --git a/tests/idris2/record003/Record.idr b/tests/idris2/data/record003/Record.idr similarity index 100% rename from tests/idris2/record003/Record.idr rename to tests/idris2/data/record003/Record.idr diff --git a/tests/idris2/record003/expected b/tests/idris2/data/record003/expected similarity index 100% rename from tests/idris2/record003/expected rename to tests/idris2/data/record003/expected diff --git a/tests/idris2/data/record003/run b/tests/idris2/data/record003/run new file mode 100755 index 0000000000..db41bc6bd8 --- /dev/null +++ b/tests/idris2/data/record003/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +idris2 -c Record.idr diff --git a/tests/idris2/record004/Main.idr b/tests/idris2/data/record004/Main.idr similarity index 100% rename from tests/idris2/record004/Main.idr rename to tests/idris2/data/record004/Main.idr diff --git a/tests/idris2/record004/expected b/tests/idris2/data/record004/expected similarity index 100% rename from tests/idris2/record004/expected rename to tests/idris2/data/record004/expected diff --git a/tests/idris2/record004/input b/tests/idris2/data/record004/input similarity index 100% rename from tests/idris2/record004/input rename to tests/idris2/data/record004/input diff --git a/tests/idris2/record004/run b/tests/idris2/data/record004/run similarity index 51% rename from tests/idris2/record004/run rename to tests/idris2/data/record004/run index 37beba4789..7935482b8c 100755 --- a/tests/idris2/record004/run +++ b/tests/idris2/data/record004/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 Main.idr < input diff --git a/tests/idris2/record005/Fld.idr b/tests/idris2/data/record005/Fld.idr similarity index 100% rename from tests/idris2/record005/Fld.idr rename to tests/idris2/data/record005/Fld.idr diff --git a/tests/idris2/record005/expected b/tests/idris2/data/record005/expected similarity index 100% rename from tests/idris2/record005/expected rename to tests/idris2/data/record005/expected diff --git a/tests/idris2/record005/input b/tests/idris2/data/record005/input similarity index 100% rename from tests/idris2/record005/input rename to tests/idris2/data/record005/input diff --git a/tests/idris2/record005/run b/tests/idris2/data/record005/run similarity index 50% rename from tests/idris2/record005/run rename to tests/idris2/data/record005/run index da8d57ecaf..a4277f19f4 100755 --- a/tests/idris2/record005/run +++ b/tests/idris2/data/record005/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 Fld.idr < input diff --git a/tests/idris2/record006/Fld.idr b/tests/idris2/data/record006/Fld.idr similarity index 100% rename from tests/idris2/record006/Fld.idr rename to tests/idris2/data/record006/Fld.idr diff --git a/tests/idris2/record006/expected b/tests/idris2/data/record006/expected similarity index 100% rename from tests/idris2/record006/expected rename to tests/idris2/data/record006/expected diff --git a/tests/idris2/record006/input b/tests/idris2/data/record006/input similarity index 100% rename from tests/idris2/record006/input rename to tests/idris2/data/record006/input diff --git a/tests/idris2/record006/run b/tests/idris2/data/record006/run similarity index 50% rename from tests/idris2/record006/run rename to tests/idris2/data/record006/run index da8d57ecaf..a4277f19f4 100755 --- a/tests/idris2/record006/run +++ b/tests/idris2/data/record006/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 Fld.idr < input diff --git a/tests/idris2/record007/Bond.idr b/tests/idris2/data/record007/Bond.idr similarity index 100% rename from tests/idris2/record007/Bond.idr rename to tests/idris2/data/record007/Bond.idr diff --git a/tests/idris2/record007/expected b/tests/idris2/data/record007/expected similarity index 100% rename from tests/idris2/record007/expected rename to tests/idris2/data/record007/expected diff --git a/tests/idris2/error008/input b/tests/idris2/data/record007/input similarity index 100% rename from tests/idris2/error008/input rename to tests/idris2/data/record007/input diff --git a/tests/idris2/record007/run b/tests/idris2/data/record007/run similarity index 51% rename from tests/idris2/record007/run rename to tests/idris2/data/record007/run index 7b1104dc1d..336bab47e3 100755 --- a/tests/idris2/record007/run +++ b/tests/idris2/data/record007/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 Bond.idr < input diff --git a/tests/idris2/record008/Postfix.idr b/tests/idris2/data/record008/Postfix.idr similarity index 100% rename from tests/idris2/record008/Postfix.idr rename to tests/idris2/data/record008/Postfix.idr diff --git a/tests/idris2/record008/expected b/tests/idris2/data/record008/expected similarity index 100% rename from tests/idris2/record008/expected rename to tests/idris2/data/record008/expected diff --git a/tests/idris2/record008/input b/tests/idris2/data/record008/input similarity index 100% rename from tests/idris2/record008/input rename to tests/idris2/data/record008/input diff --git a/tests/idris2/record008/run b/tests/idris2/data/record008/run similarity index 53% rename from tests/idris2/record008/run rename to tests/idris2/data/record008/run index 64619990c7..d9105e38fc 100755 --- a/tests/idris2/record008/run +++ b/tests/idris2/data/record008/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 Postfix.idr < input diff --git a/tests/idris2/record009/expected b/tests/idris2/data/record009/expected similarity index 100% rename from tests/idris2/record009/expected rename to tests/idris2/data/record009/expected diff --git a/tests/idris2/record009/input b/tests/idris2/data/record009/input similarity index 100% rename from tests/idris2/record009/input rename to tests/idris2/data/record009/input diff --git a/tests/idris2/record009/record.idr b/tests/idris2/data/record009/record.idr similarity index 100% rename from tests/idris2/record009/record.idr rename to tests/idris2/data/record009/record.idr diff --git a/tests/idris2/record009/run b/tests/idris2/data/record009/run similarity index 52% rename from tests/idris2/record009/run rename to tests/idris2/data/record009/run index 23ab35bdff..a210af898c 100755 --- a/tests/idris2/record009/run +++ b/tests/idris2/data/record009/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 record.idr < input diff --git a/tests/idris2/record010/expected b/tests/idris2/data/record010/expected similarity index 100% rename from tests/idris2/record010/expected rename to tests/idris2/data/record010/expected diff --git a/tests/idris2/record010/record.idr b/tests/idris2/data/record010/record.idr similarity index 100% rename from tests/idris2/record010/record.idr rename to tests/idris2/data/record010/record.idr diff --git a/tests/idris2/data/record010/run b/tests/idris2/data/record010/run new file mode 100755 index 0000000000..33a8a4cbd7 --- /dev/null +++ b/tests/idris2/data/record010/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check record.idr diff --git a/tests/idris2/record011/Issue2095.idr b/tests/idris2/data/record011/Issue2095.idr similarity index 100% rename from tests/idris2/record011/Issue2095.idr rename to tests/idris2/data/record011/Issue2095.idr diff --git a/tests/idris2/record011/expected b/tests/idris2/data/record011/expected similarity index 100% rename from tests/idris2/record011/expected rename to tests/idris2/data/record011/expected diff --git a/tests/idris2/data/record011/run b/tests/idris2/data/record011/run new file mode 100755 index 0000000000..872488ca7c --- /dev/null +++ b/tests/idris2/data/record011/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Issue2095.idr diff --git a/tests/idris2/record012/Issue2065.idr b/tests/idris2/data/record012/Issue2065.idr similarity index 100% rename from tests/idris2/record012/Issue2065.idr rename to tests/idris2/data/record012/Issue2065.idr diff --git a/tests/idris2/record012/expected b/tests/idris2/data/record012/expected similarity index 100% rename from tests/idris2/record012/expected rename to tests/idris2/data/record012/expected diff --git a/tests/idris2/data/record012/run b/tests/idris2/data/record012/run new file mode 100755 index 0000000000..c97e4d3411 --- /dev/null +++ b/tests/idris2/data/record012/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Issue2065.idr diff --git a/tests/idris2/record013/Issue1945.idr b/tests/idris2/data/record013/Issue1945.idr similarity index 100% rename from tests/idris2/record013/Issue1945.idr rename to tests/idris2/data/record013/Issue1945.idr diff --git a/tests/idris2/record013/expected b/tests/idris2/data/record013/expected similarity index 100% rename from tests/idris2/record013/expected rename to tests/idris2/data/record013/expected diff --git a/tests/idris2/record013/input b/tests/idris2/data/record013/input similarity index 100% rename from tests/idris2/record013/input rename to tests/idris2/data/record013/input diff --git a/tests/idris2/record013/run b/tests/idris2/data/record013/run similarity index 55% rename from tests/idris2/record013/run rename to tests/idris2/data/record013/run index 4574c57288..be2554462b 100755 --- a/tests/idris2/record013/run +++ b/tests/idris2/data/record013/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 Issue1945.idr < input diff --git a/tests/idris2/record014/Issue1404.idr b/tests/idris2/data/record014/Issue1404.idr similarity index 100% rename from tests/idris2/record014/Issue1404.idr rename to tests/idris2/data/record014/Issue1404.idr diff --git a/tests/idris2/record014/expected b/tests/idris2/data/record014/expected similarity index 100% rename from tests/idris2/record014/expected rename to tests/idris2/data/record014/expected diff --git a/tests/idris2/data/record014/run b/tests/idris2/data/record014/run new file mode 100755 index 0000000000..066473c9ee --- /dev/null +++ b/tests/idris2/data/record014/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Issue1404.idr diff --git a/tests/idris2/record015/Issue2176.idr b/tests/idris2/data/record015/Issue2176.idr similarity index 100% rename from tests/idris2/record015/Issue2176.idr rename to tests/idris2/data/record015/Issue2176.idr diff --git a/tests/idris2/record015/expected b/tests/idris2/data/record015/expected similarity index 100% rename from tests/idris2/record015/expected rename to tests/idris2/data/record015/expected diff --git a/tests/idris2/data/record015/run b/tests/idris2/data/record015/run new file mode 100755 index 0000000000..234ca60cc7 --- /dev/null +++ b/tests/idris2/data/record015/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +run Issue2176.idr diff --git a/tests/idris2/record016/HoleRecord.idr b/tests/idris2/data/record016/HoleRecord.idr similarity index 100% rename from tests/idris2/record016/HoleRecord.idr rename to tests/idris2/data/record016/HoleRecord.idr diff --git a/tests/idris2/record016/expected b/tests/idris2/data/record016/expected similarity index 100% rename from tests/idris2/record016/expected rename to tests/idris2/data/record016/expected diff --git a/tests/idris2/record016/input b/tests/idris2/data/record016/input similarity index 100% rename from tests/idris2/record016/input rename to tests/idris2/data/record016/input diff --git a/tests/idris2/record016/run b/tests/idris2/data/record016/run similarity index 56% rename from tests/idris2/record016/run rename to tests/idris2/data/record016/run index e4a5065fc5..a5232bfca7 100755 --- a/tests/idris2/record016/run +++ b/tests/idris2/data/record016/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 HoleRecord.idr < input diff --git a/tests/idris2/record017/RecordOptions.idr b/tests/idris2/data/record017/RecordOptions.idr similarity index 100% rename from tests/idris2/record017/RecordOptions.idr rename to tests/idris2/data/record017/RecordOptions.idr diff --git a/tests/idris2/record017/expected b/tests/idris2/data/record017/expected similarity index 100% rename from tests/idris2/record017/expected rename to tests/idris2/data/record017/expected diff --git a/tests/idris2/record017/run b/tests/idris2/data/record017/run similarity index 54% rename from tests/idris2/record017/run rename to tests/idris2/data/record017/run index db5de30ae5..cdd0d226ba 100755 --- a/tests/idris2/record017/run +++ b/tests/idris2/data/record017/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 -c RecordOptions.idr diff --git a/tests/idris2/record018/expected b/tests/idris2/data/record018/expected similarity index 100% rename from tests/idris2/record018/expected rename to tests/idris2/data/record018/expected diff --git a/tests/idris2/record018/mut.idr b/tests/idris2/data/record018/mut.idr similarity index 100% rename from tests/idris2/record018/mut.idr rename to tests/idris2/data/record018/mut.idr diff --git a/tests/idris2/data/record018/run b/tests/idris2/data/record018/run new file mode 100755 index 0000000000..ed26311e5d --- /dev/null +++ b/tests/idris2/data/record018/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check mut.idr diff --git a/tests/idris2/record019/BindParams.idr b/tests/idris2/data/record019/BindParams.idr similarity index 100% rename from tests/idris2/record019/BindParams.idr rename to tests/idris2/data/record019/BindParams.idr diff --git a/tests/idris2/record019/expected b/tests/idris2/data/record019/expected similarity index 100% rename from tests/idris2/record019/expected rename to tests/idris2/data/record019/expected diff --git a/tests/idris2/record019/run b/tests/idris2/data/record019/run similarity index 52% rename from tests/idris2/record019/run rename to tests/idris2/data/record019/run index 20aaadf3c9..4fbc3f2056 100755 --- a/tests/idris2/record019/run +++ b/tests/idris2/data/record019/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 -c BindParams.idr diff --git a/tests/idris2/data001/run b/tests/idris2/data001/run deleted file mode 100755 index bc737deb90..0000000000 --- a/tests/idris2/data001/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check TestImpl.idr diff --git a/tests/idris2/data002/run b/tests/idris2/data002/run deleted file mode 100755 index 441b64ba30..0000000000 --- a/tests/idris2/data002/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -idris2 Test.idr < input diff --git a/tests/idris2/debug001/TypePat.idr b/tests/idris2/debug/debug001/TypePat.idr similarity index 100% rename from tests/idris2/debug001/TypePat.idr rename to tests/idris2/debug/debug001/TypePat.idr diff --git a/tests/idris2/debug001/expected b/tests/idris2/debug/debug001/expected similarity index 100% rename from tests/idris2/debug001/expected rename to tests/idris2/debug/debug001/expected diff --git a/tests/idris2/debug001/input b/tests/idris2/debug/debug001/input similarity index 100% rename from tests/idris2/debug001/input rename to tests/idris2/debug/debug001/input diff --git a/tests/idris2/debug001/run b/tests/idris2/debug/debug001/run similarity index 53% rename from tests/idris2/debug001/run rename to tests/idris2/debug/debug001/run index 177aeac57a..5afb40f84d 100755 --- a/tests/idris2/debug001/run +++ b/tests/idris2/debug/debug001/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 TypePat.idr < input diff --git a/tests/idris2/docs001/run b/tests/idris2/docs001/run deleted file mode 100755 index fecfb4b107..0000000000 --- a/tests/idris2/docs001/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -idris2 < input diff --git a/tests/idris2/docs005/run b/tests/idris2/docs005/run deleted file mode 100755 index fecfb4b107..0000000000 --- a/tests/idris2/docs005/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -idris2 < input diff --git a/tests/idris2/error001/Error.idr b/tests/idris2/error/error001/Error.idr similarity index 100% rename from tests/idris2/error001/Error.idr rename to tests/idris2/error/error001/Error.idr diff --git a/tests/idris2/error001/expected b/tests/idris2/error/error001/expected similarity index 100% rename from tests/idris2/error001/expected rename to tests/idris2/error/error001/expected diff --git a/tests/idris2/error/error001/run b/tests/idris2/error/error001/run new file mode 100755 index 0000000000..b1c49389fc --- /dev/null +++ b/tests/idris2/error/error001/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Error.idr diff --git a/tests/idris2/error002/Error.idr b/tests/idris2/error/error002/Error.idr similarity index 100% rename from tests/idris2/error002/Error.idr rename to tests/idris2/error/error002/Error.idr diff --git a/tests/idris2/error002/expected b/tests/idris2/error/error002/expected similarity index 100% rename from tests/idris2/error002/expected rename to tests/idris2/error/error002/expected diff --git a/tests/idris2/error/error002/run b/tests/idris2/error/error002/run new file mode 100755 index 0000000000..b1c49389fc --- /dev/null +++ b/tests/idris2/error/error002/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Error.idr diff --git a/tests/idris2/error003/Error.idr b/tests/idris2/error/error003/Error.idr similarity index 100% rename from tests/idris2/error003/Error.idr rename to tests/idris2/error/error003/Error.idr diff --git a/tests/idris2/error003/expected b/tests/idris2/error/error003/expected similarity index 100% rename from tests/idris2/error003/expected rename to tests/idris2/error/error003/expected diff --git a/tests/idris2/error/error003/run b/tests/idris2/error/error003/run new file mode 100755 index 0000000000..b1c49389fc --- /dev/null +++ b/tests/idris2/error/error003/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Error.idr diff --git a/tests/idris2/error004/Error1.idr b/tests/idris2/error/error004/Error1.idr similarity index 100% rename from tests/idris2/error004/Error1.idr rename to tests/idris2/error/error004/Error1.idr diff --git a/tests/idris2/error004/Error2.idr b/tests/idris2/error/error004/Error2.idr similarity index 100% rename from tests/idris2/error004/Error2.idr rename to tests/idris2/error/error004/Error2.idr diff --git a/tests/idris2/error004/expected b/tests/idris2/error/error004/expected similarity index 100% rename from tests/idris2/error004/expected rename to tests/idris2/error/error004/expected diff --git a/tests/idris2/error004/run b/tests/idris2/error/error004/run similarity index 59% rename from tests/idris2/error004/run rename to tests/idris2/error/error004/run index 7f7dc5cfe6..1c39d9f12a 100755 --- a/tests/idris2/error004/run +++ b/tests/idris2/error/error004/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh check Error1.idr check Error2.idr diff --git a/tests/idris2/error005/IfErr.idr b/tests/idris2/error/error005/IfErr.idr similarity index 100% rename from tests/idris2/error005/IfErr.idr rename to tests/idris2/error/error005/IfErr.idr diff --git a/tests/idris2/error005/expected b/tests/idris2/error/error005/expected similarity index 100% rename from tests/idris2/error005/expected rename to tests/idris2/error/error005/expected diff --git a/tests/idris2/error/error005/run b/tests/idris2/error/error005/run new file mode 100755 index 0000000000..1d71cac307 --- /dev/null +++ b/tests/idris2/error/error005/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check IfErr.idr diff --git a/tests/idris2/error006/IfErr.idr b/tests/idris2/error/error006/IfErr.idr similarity index 100% rename from tests/idris2/error006/IfErr.idr rename to tests/idris2/error/error006/IfErr.idr diff --git a/tests/idris2/error006/expected b/tests/idris2/error/error006/expected similarity index 100% rename from tests/idris2/error006/expected rename to tests/idris2/error/error006/expected diff --git a/tests/idris2/error/error006/run b/tests/idris2/error/error006/run new file mode 100755 index 0000000000..1d71cac307 --- /dev/null +++ b/tests/idris2/error/error006/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check IfErr.idr diff --git a/tests/idris2/error007/CongErr.idr b/tests/idris2/error/error007/CongErr.idr similarity index 100% rename from tests/idris2/error007/CongErr.idr rename to tests/idris2/error/error007/CongErr.idr diff --git a/tests/idris2/error007/expected b/tests/idris2/error/error007/expected similarity index 100% rename from tests/idris2/error007/expected rename to tests/idris2/error/error007/expected diff --git a/tests/idris2/error/error007/run b/tests/idris2/error/error007/run new file mode 100755 index 0000000000..e3393020b0 --- /dev/null +++ b/tests/idris2/error/error007/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check CongErr.idr diff --git a/tests/idris2/error008/expected b/tests/idris2/error/error008/expected similarity index 100% rename from tests/idris2/error008/expected rename to tests/idris2/error/error008/expected diff --git a/tests/idris2/error009/input b/tests/idris2/error/error008/input similarity index 100% rename from tests/idris2/error009/input rename to tests/idris2/error/error008/input diff --git a/tests/idris2/error008/run b/tests/idris2/error/error008/run similarity index 57% rename from tests/idris2/error008/run rename to tests/idris2/error/error008/run index 1f05596214..ce8b5d6288 100755 --- a/tests/idris2/error008/run +++ b/tests/idris2/error/error008/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 DoesntExist.idr < input diff --git a/tests/idris2/error009/Exists.idr b/tests/idris2/error/error009/Exists.idr similarity index 100% rename from tests/idris2/error009/Exists.idr rename to tests/idris2/error/error009/Exists.idr diff --git a/tests/idris2/error009/expected b/tests/idris2/error/error009/expected similarity index 100% rename from tests/idris2/error009/expected rename to tests/idris2/error/error009/expected diff --git a/tests/idris2/interface025/input b/tests/idris2/error/error009/input similarity index 100% rename from tests/idris2/interface025/input rename to tests/idris2/error/error009/input diff --git a/tests/idris2/error009/run b/tests/idris2/error/error009/run similarity index 52% rename from tests/idris2/error009/run rename to tests/idris2/error/error009/run index 190587fa85..6cd937c0e9 100755 --- a/tests/idris2/error009/run +++ b/tests/idris2/error/error009/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 Exists.idr < input diff --git a/tests/idris2/error010/Loop.idr b/tests/idris2/error/error010/Loop.idr similarity index 100% rename from tests/idris2/error010/Loop.idr rename to tests/idris2/error/error010/Loop.idr diff --git a/tests/idris2/error010/expected b/tests/idris2/error/error010/expected similarity index 100% rename from tests/idris2/error010/expected rename to tests/idris2/error/error010/expected diff --git a/tests/idris2/error/error010/run b/tests/idris2/error/error010/run new file mode 100755 index 0000000000..64e356539e --- /dev/null +++ b/tests/idris2/error/error010/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Loop.idr diff --git a/tests/idris2/error011/ConstructorDuplicate.idr b/tests/idris2/error/error011/ConstructorDuplicate.idr similarity index 100% rename from tests/idris2/error011/ConstructorDuplicate.idr rename to tests/idris2/error/error011/ConstructorDuplicate.idr diff --git a/tests/idris2/error011/expected b/tests/idris2/error/error011/expected similarity index 100% rename from tests/idris2/error011/expected rename to tests/idris2/error/error011/expected diff --git a/tests/idris2/error011/run b/tests/idris2/error/error011/run similarity index 57% rename from tests/idris2/error011/run rename to tests/idris2/error/error011/run index 20825f8720..3dda6246a0 100755 --- a/tests/idris2/error011/run +++ b/tests/idris2/error/error011/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh check ConstructorDuplicate.idr diff --git a/tests/idris2/error012/expected b/tests/idris2/error/error012/expected similarity index 100% rename from tests/idris2/error012/expected rename to tests/idris2/error/error012/expected diff --git a/tests/idris2/error012/run b/tests/idris2/error/error012/run similarity index 100% rename from tests/idris2/error012/run rename to tests/idris2/error/error012/run diff --git a/tests/idris2/error013/Issue361.idr b/tests/idris2/error/error013/Issue361.idr similarity index 100% rename from tests/idris2/error013/Issue361.idr rename to tests/idris2/error/error013/Issue361.idr diff --git a/tests/idris2/error013/expected b/tests/idris2/error/error013/expected similarity index 100% rename from tests/idris2/error013/expected rename to tests/idris2/error/error013/expected diff --git a/tests/idris2/error/error013/run b/tests/idris2/error/error013/run new file mode 100755 index 0000000000..7c76542f48 --- /dev/null +++ b/tests/idris2/error/error013/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Issue361.idr diff --git a/tests/idris2/error014/Issue735.idr b/tests/idris2/error/error014/Issue735.idr similarity index 100% rename from tests/idris2/error014/Issue735.idr rename to tests/idris2/error/error014/Issue735.idr diff --git a/tests/idris2/error014/expected b/tests/idris2/error/error014/expected similarity index 100% rename from tests/idris2/error014/expected rename to tests/idris2/error/error014/expected diff --git a/tests/idris2/error/error014/run b/tests/idris2/error/error014/run new file mode 100755 index 0000000000..2198d33524 --- /dev/null +++ b/tests/idris2/error/error014/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Issue735.idr diff --git a/tests/idris2/error015/Issue110.idr b/tests/idris2/error/error015/Issue110.idr similarity index 100% rename from tests/idris2/error015/Issue110.idr rename to tests/idris2/error/error015/Issue110.idr diff --git a/tests/idris2/error015/expected b/tests/idris2/error/error015/expected similarity index 100% rename from tests/idris2/error015/expected rename to tests/idris2/error/error015/expected diff --git a/tests/idris2/error/error015/run b/tests/idris2/error/error015/run new file mode 100755 index 0000000000..9d929ca233 --- /dev/null +++ b/tests/idris2/error/error015/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Issue110.idr diff --git a/tests/idris2/error016/Issue1230.idr b/tests/idris2/error/error016/Issue1230.idr similarity index 100% rename from tests/idris2/error016/Issue1230.idr rename to tests/idris2/error/error016/Issue1230.idr diff --git a/tests/idris2/error016/expected b/tests/idris2/error/error016/expected similarity index 100% rename from tests/idris2/error016/expected rename to tests/idris2/error/error016/expected diff --git a/tests/idris2/error016/input b/tests/idris2/error/error016/input similarity index 100% rename from tests/idris2/error016/input rename to tests/idris2/error/error016/input diff --git a/tests/idris2/error016/run b/tests/idris2/error/error016/run similarity index 60% rename from tests/idris2/error016/run rename to tests/idris2/error/error016/run index b1b3a6b727..9cacac0a1d 100755 --- a/tests/idris2/error016/run +++ b/tests/idris2/error/error016/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh check Issue1230.idr idris2 < input diff --git a/tests/idris2/error017/Issue962-case.idr b/tests/idris2/error/error017/Issue962-case.idr similarity index 100% rename from tests/idris2/error017/Issue962-case.idr rename to tests/idris2/error/error017/Issue962-case.idr diff --git a/tests/idris2/error017/Issue962.idr b/tests/idris2/error/error017/Issue962.idr similarity index 100% rename from tests/idris2/error017/Issue962.idr rename to tests/idris2/error/error017/Issue962.idr diff --git a/tests/idris2/error017/expected b/tests/idris2/error/error017/expected similarity index 100% rename from tests/idris2/error017/expected rename to tests/idris2/error/error017/expected diff --git a/tests/idris2/error017/run b/tests/idris2/error/error017/run similarity index 64% rename from tests/idris2/error017/run rename to tests/idris2/error/error017/run index f5422b8d46..c464964d4c 100755 --- a/tests/idris2/error017/run +++ b/tests/idris2/error/error017/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh check Issue962.idr check Issue962-case.idr diff --git a/tests/idris2/error018/Issue1031-2.idr b/tests/idris2/error/error018/Issue1031-2.idr similarity index 100% rename from tests/idris2/error018/Issue1031-2.idr rename to tests/idris2/error/error018/Issue1031-2.idr diff --git a/tests/idris2/error018/Issue1031-3.idr b/tests/idris2/error/error018/Issue1031-3.idr similarity index 100% rename from tests/idris2/error018/Issue1031-3.idr rename to tests/idris2/error/error018/Issue1031-3.idr diff --git a/tests/idris2/error018/Issue1031-4.idr b/tests/idris2/error/error018/Issue1031-4.idr similarity index 100% rename from tests/idris2/error018/Issue1031-4.idr rename to tests/idris2/error/error018/Issue1031-4.idr diff --git a/tests/idris2/error018/Issue1031.idr b/tests/idris2/error/error018/Issue1031.idr similarity index 100% rename from tests/idris2/error018/Issue1031.idr rename to tests/idris2/error/error018/Issue1031.idr diff --git a/tests/idris2/error018/expected b/tests/idris2/error/error018/expected similarity index 100% rename from tests/idris2/error018/expected rename to tests/idris2/error/error018/expected diff --git a/tests/idris2/error018/run b/tests/idris2/error/error018/run similarity index 78% rename from tests/idris2/error018/run rename to tests/idris2/error/error018/run index 19747197aa..ed1ff9282b 100755 --- a/tests/idris2/error018/run +++ b/tests/idris2/error/error018/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh check Issue1031.idr check Issue1031-2.idr diff --git a/tests/idris2/error019/Error.idr b/tests/idris2/error/error019/Error.idr similarity index 100% rename from tests/idris2/error019/Error.idr rename to tests/idris2/error/error019/Error.idr diff --git a/tests/idris2/error019/expected b/tests/idris2/error/error019/expected similarity index 100% rename from tests/idris2/error019/expected rename to tests/idris2/error/error019/expected diff --git a/tests/idris2/error019/run b/tests/idris2/error/error019/run similarity index 51% rename from tests/idris2/error019/run rename to tests/idris2/error/error019/run index 043b1a25e1..9c6d71931c 100644 --- a/tests/idris2/error019/run +++ b/tests/idris2/error/error019/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh check -Werror Error.idr diff --git a/tests/idris2/error020/Error.idr b/tests/idris2/error/error020/Error.idr similarity index 100% rename from tests/idris2/error020/Error.idr rename to tests/idris2/error/error020/Error.idr diff --git a/tests/idris2/error020/expected b/tests/idris2/error/error020/expected similarity index 100% rename from tests/idris2/error020/expected rename to tests/idris2/error/error020/expected diff --git a/tests/idris2/error020/run b/tests/idris2/error/error020/run similarity index 51% rename from tests/idris2/error020/run rename to tests/idris2/error/error020/run index 043b1a25e1..9c6d71931c 100644 --- a/tests/idris2/error020/run +++ b/tests/idris2/error/error020/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh check -Werror Error.idr diff --git a/tests/idris2/error021/DeepAmbig.idr b/tests/idris2/error/error021/DeepAmbig.idr similarity index 100% rename from tests/idris2/error021/DeepAmbig.idr rename to tests/idris2/error/error021/DeepAmbig.idr diff --git a/tests/idris2/error021/expected b/tests/idris2/error/error021/expected similarity index 100% rename from tests/idris2/error021/expected rename to tests/idris2/error/error021/expected diff --git a/tests/idris2/error021/run b/tests/idris2/error/error021/run similarity index 54% rename from tests/idris2/error021/run rename to tests/idris2/error/error021/run index 25a8e8c361..0c1cc11a10 100644 --- a/tests/idris2/error021/run +++ b/tests/idris2/error/error021/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh check -Werror DeepAmbig.idr diff --git a/tests/idris2/error022/UpdateLoc.idr b/tests/idris2/error/error022/UpdateLoc.idr similarity index 100% rename from tests/idris2/error022/UpdateLoc.idr rename to tests/idris2/error/error022/UpdateLoc.idr diff --git a/tests/idris2/error022/expected b/tests/idris2/error/error022/expected similarity index 100% rename from tests/idris2/error022/expected rename to tests/idris2/error/error022/expected diff --git a/tests/idris2/error/error022/run b/tests/idris2/error/error022/run new file mode 100644 index 0000000000..a581b57b84 --- /dev/null +++ b/tests/idris2/error/error022/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check UpdateLoc.idr diff --git a/tests/idris2/error023/Error1.idr b/tests/idris2/error/error023/Error1.idr similarity index 100% rename from tests/idris2/error023/Error1.idr rename to tests/idris2/error/error023/Error1.idr diff --git a/tests/idris2/error023/Error2.idr b/tests/idris2/error/error023/Error2.idr similarity index 100% rename from tests/idris2/error023/Error2.idr rename to tests/idris2/error/error023/Error2.idr diff --git a/tests/idris2/error023/expected b/tests/idris2/error/error023/expected similarity index 100% rename from tests/idris2/error023/expected rename to tests/idris2/error/error023/expected diff --git a/tests/idris2/error023/run b/tests/idris2/error/error023/run similarity index 59% rename from tests/idris2/error023/run rename to tests/idris2/error/error023/run index 7f7dc5cfe6..1c39d9f12a 100755 --- a/tests/idris2/error023/run +++ b/tests/idris2/error/error023/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh check Error1.idr check Error2.idr diff --git a/tests/idris2/error024/Error1.idr b/tests/idris2/error/error024/Error1.idr similarity index 100% rename from tests/idris2/error024/Error1.idr rename to tests/idris2/error/error024/Error1.idr diff --git a/tests/idris2/error024/expected b/tests/idris2/error/error024/expected similarity index 100% rename from tests/idris2/error024/expected rename to tests/idris2/error/error024/expected diff --git a/tests/idris2/error/error024/run b/tests/idris2/error/error024/run new file mode 100755 index 0000000000..33ef6f8c13 --- /dev/null +++ b/tests/idris2/error/error024/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Error1.idr diff --git a/tests/idris2/error025/IAlternativePrints.idr b/tests/idris2/error/error025/IAlternativePrints.idr similarity index 100% rename from tests/idris2/error025/IAlternativePrints.idr rename to tests/idris2/error/error025/IAlternativePrints.idr diff --git a/tests/idris2/error025/expected b/tests/idris2/error/error025/expected similarity index 100% rename from tests/idris2/error025/expected rename to tests/idris2/error/error025/expected diff --git a/tests/idris2/error025/run b/tests/idris2/error/error025/run similarity index 55% rename from tests/idris2/error025/run rename to tests/idris2/error/error025/run index 99d726ae76..8ffd35c382 100755 --- a/tests/idris2/error025/run +++ b/tests/idris2/error/error025/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh check IAlternativePrints.idr diff --git a/tests/idris2/error026/DoBlockFC.idr b/tests/idris2/error/error026/DoBlockFC.idr similarity index 100% rename from tests/idris2/error026/DoBlockFC.idr rename to tests/idris2/error/error026/DoBlockFC.idr diff --git a/tests/idris2/error026/expected b/tests/idris2/error/error026/expected similarity index 100% rename from tests/idris2/error026/expected rename to tests/idris2/error/error026/expected diff --git a/tests/idris2/error/error026/run b/tests/idris2/error/error026/run new file mode 100755 index 0000000000..7c7b974344 --- /dev/null +++ b/tests/idris2/error/error026/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check DoBlockFC.idr diff --git a/tests/idris2/error027/Issue2950.idr b/tests/idris2/error/error027/Issue2950.idr similarity index 100% rename from tests/idris2/error027/Issue2950.idr rename to tests/idris2/error/error027/Issue2950.idr diff --git a/tests/idris2/error027/expected b/tests/idris2/error/error027/expected similarity index 100% rename from tests/idris2/error027/expected rename to tests/idris2/error/error027/expected diff --git a/tests/idris2/error/error027/run b/tests/idris2/error/error027/run new file mode 100755 index 0000000000..6acc871a82 --- /dev/null +++ b/tests/idris2/error/error027/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +run Issue2950.idr diff --git a/tests/idris2/perror001/PError.idr b/tests/idris2/error/perror001/PError.idr similarity index 100% rename from tests/idris2/perror001/PError.idr rename to tests/idris2/error/perror001/PError.idr diff --git a/tests/idris2/perror001/expected b/tests/idris2/error/perror001/expected similarity index 100% rename from tests/idris2/perror001/expected rename to tests/idris2/error/perror001/expected diff --git a/tests/idris2/error/perror001/run b/tests/idris2/error/perror001/run new file mode 100755 index 0000000000..e0aa4e19f5 --- /dev/null +++ b/tests/idris2/error/perror001/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check PError.idr diff --git a/tests/idris2/perror002/PError.idr b/tests/idris2/error/perror002/PError.idr similarity index 100% rename from tests/idris2/perror002/PError.idr rename to tests/idris2/error/perror002/PError.idr diff --git a/tests/idris2/perror002/expected b/tests/idris2/error/perror002/expected similarity index 100% rename from tests/idris2/perror002/expected rename to tests/idris2/error/perror002/expected diff --git a/tests/idris2/error/perror002/run b/tests/idris2/error/perror002/run new file mode 100755 index 0000000000..e0aa4e19f5 --- /dev/null +++ b/tests/idris2/error/perror002/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check PError.idr diff --git a/tests/idris2/perror003/PError.idr b/tests/idris2/error/perror003/PError.idr similarity index 100% rename from tests/idris2/perror003/PError.idr rename to tests/idris2/error/perror003/PError.idr diff --git a/tests/idris2/perror003/PError2.idr b/tests/idris2/error/perror003/PError2.idr similarity index 100% rename from tests/idris2/perror003/PError2.idr rename to tests/idris2/error/perror003/PError2.idr diff --git a/tests/idris2/perror003/expected b/tests/idris2/error/perror003/expected similarity index 100% rename from tests/idris2/perror003/expected rename to tests/idris2/error/perror003/expected diff --git a/tests/idris2/perror003/run b/tests/idris2/error/perror003/run similarity index 60% rename from tests/idris2/perror003/run rename to tests/idris2/error/perror003/run index 222a9b1b14..0d34530944 100755 --- a/tests/idris2/perror003/run +++ b/tests/idris2/error/perror003/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh check PError.idr check PError2.idr diff --git a/tests/idris2/perror004/PError.idr b/tests/idris2/error/perror004/PError.idr similarity index 100% rename from tests/idris2/perror004/PError.idr rename to tests/idris2/error/perror004/PError.idr diff --git a/tests/idris2/perror004/expected b/tests/idris2/error/perror004/expected similarity index 100% rename from tests/idris2/perror004/expected rename to tests/idris2/error/perror004/expected diff --git a/tests/idris2/error/perror004/run b/tests/idris2/error/perror004/run new file mode 100755 index 0000000000..e0aa4e19f5 --- /dev/null +++ b/tests/idris2/error/perror004/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check PError.idr diff --git a/tests/idris2/perror005/PError.idr b/tests/idris2/error/perror005/PError.idr similarity index 100% rename from tests/idris2/perror005/PError.idr rename to tests/idris2/error/perror005/PError.idr diff --git a/tests/idris2/perror005/expected b/tests/idris2/error/perror005/expected similarity index 100% rename from tests/idris2/perror005/expected rename to tests/idris2/error/perror005/expected diff --git a/tests/idris2/error/perror005/run b/tests/idris2/error/perror005/run new file mode 100755 index 0000000000..e0aa4e19f5 --- /dev/null +++ b/tests/idris2/error/perror005/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check PError.idr diff --git a/tests/idris2/perror006/PError.idr b/tests/idris2/error/perror006/PError.idr similarity index 100% rename from tests/idris2/perror006/PError.idr rename to tests/idris2/error/perror006/PError.idr diff --git a/tests/idris2/perror006/expected b/tests/idris2/error/perror006/expected similarity index 100% rename from tests/idris2/perror006/expected rename to tests/idris2/error/perror006/expected diff --git a/tests/idris2/error/perror006/run b/tests/idris2/error/perror006/run new file mode 100755 index 0000000000..e0aa4e19f5 --- /dev/null +++ b/tests/idris2/error/perror006/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check PError.idr diff --git a/tests/idris2/perror007/StrError1.idr b/tests/idris2/error/perror007/StrError1.idr similarity index 100% rename from tests/idris2/perror007/StrError1.idr rename to tests/idris2/error/perror007/StrError1.idr diff --git a/tests/idris2/perror007/StrError10.idr b/tests/idris2/error/perror007/StrError10.idr similarity index 100% rename from tests/idris2/perror007/StrError10.idr rename to tests/idris2/error/perror007/StrError10.idr diff --git a/tests/idris2/perror007/StrError11.idr b/tests/idris2/error/perror007/StrError11.idr similarity index 100% rename from tests/idris2/perror007/StrError11.idr rename to tests/idris2/error/perror007/StrError11.idr diff --git a/tests/idris2/perror007/StrError12.idr b/tests/idris2/error/perror007/StrError12.idr similarity index 100% rename from tests/idris2/perror007/StrError12.idr rename to tests/idris2/error/perror007/StrError12.idr diff --git a/tests/idris2/perror007/StrError2.idr b/tests/idris2/error/perror007/StrError2.idr similarity index 100% rename from tests/idris2/perror007/StrError2.idr rename to tests/idris2/error/perror007/StrError2.idr diff --git a/tests/idris2/perror007/StrError3.idr b/tests/idris2/error/perror007/StrError3.idr similarity index 100% rename from tests/idris2/perror007/StrError3.idr rename to tests/idris2/error/perror007/StrError3.idr diff --git a/tests/idris2/perror007/StrError4.idr b/tests/idris2/error/perror007/StrError4.idr similarity index 100% rename from tests/idris2/perror007/StrError4.idr rename to tests/idris2/error/perror007/StrError4.idr diff --git a/tests/idris2/perror007/StrError5.idr b/tests/idris2/error/perror007/StrError5.idr similarity index 100% rename from tests/idris2/perror007/StrError5.idr rename to tests/idris2/error/perror007/StrError5.idr diff --git a/tests/idris2/perror007/StrError6.idr b/tests/idris2/error/perror007/StrError6.idr similarity index 100% rename from tests/idris2/perror007/StrError6.idr rename to tests/idris2/error/perror007/StrError6.idr diff --git a/tests/idris2/perror007/StrError7.idr b/tests/idris2/error/perror007/StrError7.idr similarity index 100% rename from tests/idris2/perror007/StrError7.idr rename to tests/idris2/error/perror007/StrError7.idr diff --git a/tests/idris2/perror007/StrError8.idr b/tests/idris2/error/perror007/StrError8.idr similarity index 100% rename from tests/idris2/perror007/StrError8.idr rename to tests/idris2/error/perror007/StrError8.idr diff --git a/tests/idris2/perror007/StrError9.idr b/tests/idris2/error/perror007/StrError9.idr similarity index 100% rename from tests/idris2/perror007/StrError9.idr rename to tests/idris2/error/perror007/StrError9.idr diff --git a/tests/idris2/perror007/expected b/tests/idris2/error/perror007/expected similarity index 100% rename from tests/idris2/perror007/expected rename to tests/idris2/error/perror007/expected diff --git a/tests/idris2/perror007/run b/tests/idris2/error/perror007/run similarity index 93% rename from tests/idris2/perror007/run rename to tests/idris2/error/perror007/run index 9b41cb8191..f9f99edda4 100755 --- a/tests/idris2/perror007/run +++ b/tests/idris2/error/perror007/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh check StrError1.idr || true check StrError2.idr || true diff --git a/tests/idris2/perror008/Issue1224a.idr b/tests/idris2/error/perror008/Issue1224a.idr similarity index 100% rename from tests/idris2/perror008/Issue1224a.idr rename to tests/idris2/error/perror008/Issue1224a.idr diff --git a/tests/idris2/perror008/Issue1224b.idr b/tests/idris2/error/perror008/Issue1224b.idr similarity index 100% rename from tests/idris2/perror008/Issue1224b.idr rename to tests/idris2/error/perror008/Issue1224b.idr diff --git a/tests/idris2/perror008/Issue710a.idr b/tests/idris2/error/perror008/Issue710a.idr similarity index 100% rename from tests/idris2/perror008/Issue710a.idr rename to tests/idris2/error/perror008/Issue710a.idr diff --git a/tests/idris2/perror008/Issue710b.idr b/tests/idris2/error/perror008/Issue710b.idr similarity index 100% rename from tests/idris2/perror008/Issue710b.idr rename to tests/idris2/error/perror008/Issue710b.idr diff --git a/tests/idris2/perror008/Issue710c.idr b/tests/idris2/error/perror008/Issue710c.idr similarity index 100% rename from tests/idris2/perror008/Issue710c.idr rename to tests/idris2/error/perror008/Issue710c.idr diff --git a/tests/idris2/perror008/Issue710d.idr b/tests/idris2/error/perror008/Issue710d.idr similarity index 100% rename from tests/idris2/perror008/Issue710d.idr rename to tests/idris2/error/perror008/Issue710d.idr diff --git a/tests/idris2/perror008/Issue710e.idr b/tests/idris2/error/perror008/Issue710e.idr similarity index 100% rename from tests/idris2/perror008/Issue710e.idr rename to tests/idris2/error/perror008/Issue710e.idr diff --git a/tests/idris2/perror008/Issue710f.idr b/tests/idris2/error/perror008/Issue710f.idr similarity index 100% rename from tests/idris2/perror008/Issue710f.idr rename to tests/idris2/error/perror008/Issue710f.idr diff --git a/tests/idris2/perror008/expected b/tests/idris2/error/perror008/expected similarity index 100% rename from tests/idris2/perror008/expected rename to tests/idris2/error/perror008/expected diff --git a/tests/idris2/perror008/run b/tests/idris2/error/perror008/run similarity index 90% rename from tests/idris2/perror008/run rename to tests/idris2/error/perror008/run index 86453dab6f..2adb354935 100755 --- a/tests/idris2/perror008/run +++ b/tests/idris2/error/perror008/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh check Issue710a.idr || true check Issue710b.idr || true diff --git a/tests/idris2/perror009/Error1.idr b/tests/idris2/error/perror009/Error1.idr similarity index 100% rename from tests/idris2/perror009/Error1.idr rename to tests/idris2/error/perror009/Error1.idr diff --git a/tests/idris2/perror009/expected b/tests/idris2/error/perror009/expected similarity index 100% rename from tests/idris2/perror009/expected rename to tests/idris2/error/perror009/expected diff --git a/tests/idris2/perror009/run b/tests/idris2/error/perror009/run similarity index 52% rename from tests/idris2/perror009/run rename to tests/idris2/error/perror009/run index 99b7a73375..314f398e8c 100644 --- a/tests/idris2/perror009/run +++ b/tests/idris2/error/perror009/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh check -Werror Error1.idr diff --git a/tests/idris2/perror010/NamedReturn1.idr b/tests/idris2/error/perror010/NamedReturn1.idr similarity index 100% rename from tests/idris2/perror010/NamedReturn1.idr rename to tests/idris2/error/perror010/NamedReturn1.idr diff --git a/tests/idris2/perror010/NamedReturn2.idr b/tests/idris2/error/perror010/NamedReturn2.idr similarity index 100% rename from tests/idris2/perror010/NamedReturn2.idr rename to tests/idris2/error/perror010/NamedReturn2.idr diff --git a/tests/idris2/perror010/NamedReturn3.idr b/tests/idris2/error/perror010/NamedReturn3.idr similarity index 100% rename from tests/idris2/perror010/NamedReturn3.idr rename to tests/idris2/error/perror010/NamedReturn3.idr diff --git a/tests/idris2/perror010/NamedReturn4.idr b/tests/idris2/error/perror010/NamedReturn4.idr similarity index 100% rename from tests/idris2/perror010/NamedReturn4.idr rename to tests/idris2/error/perror010/NamedReturn4.idr diff --git a/tests/idris2/perror010/TrailingLam.idr b/tests/idris2/error/perror010/TrailingLam.idr similarity index 100% rename from tests/idris2/perror010/TrailingLam.idr rename to tests/idris2/error/perror010/TrailingLam.idr diff --git a/tests/idris2/perror010/expected b/tests/idris2/error/perror010/expected similarity index 100% rename from tests/idris2/perror010/expected rename to tests/idris2/error/perror010/expected diff --git a/tests/idris2/perror010/input b/tests/idris2/error/perror010/input similarity index 100% rename from tests/idris2/perror010/input rename to tests/idris2/error/perror010/input diff --git a/tests/idris2/perror010/run b/tests/idris2/error/perror010/run similarity index 86% rename from tests/idris2/perror010/run rename to tests/idris2/error/perror010/run index a8b988361a..66dc47de3b 100755 --- a/tests/idris2/perror010/run +++ b/tests/idris2/error/perror010/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh check NamedReturn1.idr || true check NamedReturn2.idr || true diff --git a/tests/idris2/perror011/Issue1345.idr b/tests/idris2/error/perror011/Issue1345.idr similarity index 100% rename from tests/idris2/perror011/Issue1345.idr rename to tests/idris2/error/perror011/Issue1345.idr diff --git a/tests/idris2/perror011/Issue1496-1.idr b/tests/idris2/error/perror011/Issue1496-1.idr similarity index 100% rename from tests/idris2/perror011/Issue1496-1.idr rename to tests/idris2/error/perror011/Issue1496-1.idr diff --git a/tests/idris2/perror011/Issue1496-2.idr b/tests/idris2/error/perror011/Issue1496-2.idr similarity index 100% rename from tests/idris2/perror011/Issue1496-2.idr rename to tests/idris2/error/perror011/Issue1496-2.idr diff --git a/tests/idris2/perror011/Main.idr b/tests/idris2/error/perror011/Main.idr similarity index 100% rename from tests/idris2/perror011/Main.idr rename to tests/idris2/error/perror011/Main.idr diff --git a/tests/idris2/perror011/Pretty.idr b/tests/idris2/error/perror011/Pretty.idr similarity index 100% rename from tests/idris2/perror011/Pretty.idr rename to tests/idris2/error/perror011/Pretty.idr diff --git a/tests/idris2/perror011/expected b/tests/idris2/error/perror011/expected similarity index 100% rename from tests/idris2/perror011/expected rename to tests/idris2/error/perror011/expected diff --git a/tests/idris2/perror011/foo.ipkg b/tests/idris2/error/perror011/foo.ipkg similarity index 100% rename from tests/idris2/perror011/foo.ipkg rename to tests/idris2/error/perror011/foo.ipkg diff --git a/tests/idris2/perror011/run b/tests/idris2/error/perror011/run similarity index 82% rename from tests/idris2/perror011/run rename to tests/idris2/error/perror011/run index 7d65c2dd7d..112e27c363 100755 --- a/tests/idris2/perror011/run +++ b/tests/idris2/error/perror011/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh check Issue1345.idr || true check Issue1496-1.idr || true diff --git a/tests/idris2/perror012/CaseParseError.idr b/tests/idris2/error/perror012/CaseParseError.idr similarity index 100% rename from tests/idris2/perror012/CaseParseError.idr rename to tests/idris2/error/perror012/CaseParseError.idr diff --git a/tests/idris2/perror012/LamParseError.idr b/tests/idris2/error/perror012/LamParseError.idr similarity index 100% rename from tests/idris2/perror012/LamParseError.idr rename to tests/idris2/error/perror012/LamParseError.idr diff --git a/tests/idris2/perror012/expected b/tests/idris2/error/perror012/expected similarity index 100% rename from tests/idris2/perror012/expected rename to tests/idris2/error/perror012/expected diff --git a/tests/idris2/perror012/run b/tests/idris2/error/perror012/run similarity index 67% rename from tests/idris2/perror012/run rename to tests/idris2/error/perror012/run index 3ff2996284..f823a09bdd 100755 --- a/tests/idris2/perror012/run +++ b/tests/idris2/error/perror012/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh check LamParseError.idr check CaseParseError.idr diff --git a/tests/idris2/perror013/EmptyFailing.idr b/tests/idris2/error/perror013/EmptyFailing.idr similarity index 100% rename from tests/idris2/perror013/EmptyFailing.idr rename to tests/idris2/error/perror013/EmptyFailing.idr diff --git a/tests/idris2/perror013/EmptyMutual.idr b/tests/idris2/error/perror013/EmptyMutual.idr similarity index 100% rename from tests/idris2/perror013/EmptyMutual.idr rename to tests/idris2/error/perror013/EmptyMutual.idr diff --git a/tests/idris2/perror013/EmptyParameters.idr b/tests/idris2/error/perror013/EmptyParameters.idr similarity index 100% rename from tests/idris2/perror013/EmptyParameters.idr rename to tests/idris2/error/perror013/EmptyParameters.idr diff --git a/tests/idris2/perror013/EmptyUsing.idr b/tests/idris2/error/perror013/EmptyUsing.idr similarity index 100% rename from tests/idris2/perror013/EmptyUsing.idr rename to tests/idris2/error/perror013/EmptyUsing.idr diff --git a/tests/idris2/perror013/expected b/tests/idris2/error/perror013/expected similarity index 100% rename from tests/idris2/perror013/expected rename to tests/idris2/error/perror013/expected diff --git a/tests/idris2/perror013/run b/tests/idris2/error/perror013/run similarity index 83% rename from tests/idris2/perror013/run rename to tests/idris2/error/perror013/run index a0f3d3abdc..65dc4a4bf5 100755 --- a/tests/idris2/perror013/run +++ b/tests/idris2/error/perror013/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh check EmptyFailing.idr || true check EmptyMutual.idr || true diff --git a/tests/idris2/perror013/x.ipkg b/tests/idris2/error/perror013/x.ipkg similarity index 100% rename from tests/idris2/perror013/x.ipkg rename to tests/idris2/error/perror013/x.ipkg diff --git a/tests/idris2/perror014/ParseList.idr b/tests/idris2/error/perror014/ParseList.idr similarity index 100% rename from tests/idris2/perror014/ParseList.idr rename to tests/idris2/error/perror014/ParseList.idr diff --git a/tests/idris2/perror014/expected b/tests/idris2/error/perror014/expected similarity index 100% rename from tests/idris2/perror014/expected rename to tests/idris2/error/perror014/expected diff --git a/tests/idris2/perror014/run b/tests/idris2/error/perror014/run similarity index 54% rename from tests/idris2/perror014/run rename to tests/idris2/error/perror014/run index cb78507392..ad5f83ee58 100755 --- a/tests/idris2/perror014/run +++ b/tests/idris2/error/perror014/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh check ParseList.idr || true diff --git a/tests/idris2/perror015/ParseWith.idr b/tests/idris2/error/perror015/ParseWith.idr similarity index 100% rename from tests/idris2/perror015/ParseWith.idr rename to tests/idris2/error/perror015/ParseWith.idr diff --git a/tests/idris2/perror015/expected b/tests/idris2/error/perror015/expected similarity index 100% rename from tests/idris2/perror015/expected rename to tests/idris2/error/perror015/expected diff --git a/tests/idris2/perror015/run b/tests/idris2/error/perror015/run similarity index 54% rename from tests/idris2/perror015/run rename to tests/idris2/error/perror015/run index 9588fd16a7..d7f27147bf 100755 --- a/tests/idris2/perror015/run +++ b/tests/idris2/error/perror015/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh check ParseWith.idr || true diff --git a/tests/idris2/perror016/ParseIf.idr b/tests/idris2/error/perror016/ParseIf.idr similarity index 100% rename from tests/idris2/perror016/ParseIf.idr rename to tests/idris2/error/perror016/ParseIf.idr diff --git a/tests/idris2/perror016/ParseIf2.idr b/tests/idris2/error/perror016/ParseIf2.idr similarity index 100% rename from tests/idris2/perror016/ParseIf2.idr rename to tests/idris2/error/perror016/ParseIf2.idr diff --git a/tests/idris2/perror016/ParseIf3.idr b/tests/idris2/error/perror016/ParseIf3.idr similarity index 100% rename from tests/idris2/perror016/ParseIf3.idr rename to tests/idris2/error/perror016/ParseIf3.idr diff --git a/tests/idris2/perror016/ParseIf4.idr b/tests/idris2/error/perror016/ParseIf4.idr similarity index 100% rename from tests/idris2/perror016/ParseIf4.idr rename to tests/idris2/error/perror016/ParseIf4.idr diff --git a/tests/idris2/perror016/expected b/tests/idris2/error/perror016/expected similarity index 100% rename from tests/idris2/perror016/expected rename to tests/idris2/error/perror016/expected diff --git a/tests/idris2/perror016/run b/tests/idris2/error/perror016/run similarity index 81% rename from tests/idris2/perror016/run rename to tests/idris2/error/perror016/run index c35719fae6..8ce5dd723c 100755 --- a/tests/idris2/perror016/run +++ b/tests/idris2/error/perror016/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh check ParseIf.idr || true check ParseIf2.idr || true diff --git a/tests/idris2/perror017/ParseImpl.idr b/tests/idris2/error/perror017/ParseImpl.idr similarity index 100% rename from tests/idris2/perror017/ParseImpl.idr rename to tests/idris2/error/perror017/ParseImpl.idr diff --git a/tests/idris2/perror017/expected b/tests/idris2/error/perror017/expected similarity index 100% rename from tests/idris2/perror017/expected rename to tests/idris2/error/perror017/expected diff --git a/tests/idris2/perror017/run b/tests/idris2/error/perror017/run similarity index 54% rename from tests/idris2/perror017/run rename to tests/idris2/error/perror017/run index 03b6a7cb94..089074ebce 100755 --- a/tests/idris2/perror017/run +++ b/tests/idris2/error/perror017/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh check ParseImpl.idr || true diff --git a/tests/idris2/perror018/ParseRecord.idr b/tests/idris2/error/perror018/ParseRecord.idr similarity index 100% rename from tests/idris2/perror018/ParseRecord.idr rename to tests/idris2/error/perror018/ParseRecord.idr diff --git a/tests/idris2/perror018/ParseRecord2.idr b/tests/idris2/error/perror018/ParseRecord2.idr similarity index 100% rename from tests/idris2/perror018/ParseRecord2.idr rename to tests/idris2/error/perror018/ParseRecord2.idr diff --git a/tests/idris2/perror018/ParseRecord3.idr b/tests/idris2/error/perror018/ParseRecord3.idr similarity index 100% rename from tests/idris2/perror018/ParseRecord3.idr rename to tests/idris2/error/perror018/ParseRecord3.idr diff --git a/tests/idris2/perror018/expected b/tests/idris2/error/perror018/expected similarity index 100% rename from tests/idris2/perror018/expected rename to tests/idris2/error/perror018/expected diff --git a/tests/idris2/perror018/run b/tests/idris2/error/perror018/run similarity index 79% rename from tests/idris2/perror018/run rename to tests/idris2/error/perror018/run index 42e0ccfebb..8914bcc8e6 100755 --- a/tests/idris2/perror018/run +++ b/tests/idris2/error/perror018/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh check ParseRecord.idr || true check ParseRecord2.idr || true diff --git a/tests/idris2/perror019/ImplError.idr b/tests/idris2/error/perror019/ImplError.idr similarity index 100% rename from tests/idris2/perror019/ImplError.idr rename to tests/idris2/error/perror019/ImplError.idr diff --git a/tests/idris2/perror019/expected b/tests/idris2/error/perror019/expected similarity index 100% rename from tests/idris2/perror019/expected rename to tests/idris2/error/perror019/expected diff --git a/tests/idris2/error/perror019/run b/tests/idris2/error/perror019/run new file mode 100644 index 0000000000..9c858ad83e --- /dev/null +++ b/tests/idris2/error/perror019/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check ImplError.idr diff --git a/tests/idris2/perror020/Issue2769.idr b/tests/idris2/error/perror020/Issue2769.idr similarity index 100% rename from tests/idris2/perror020/Issue2769.idr rename to tests/idris2/error/perror020/Issue2769.idr diff --git a/tests/idris2/perror020/Issue2769b.idr b/tests/idris2/error/perror020/Issue2769b.idr similarity index 100% rename from tests/idris2/perror020/Issue2769b.idr rename to tests/idris2/error/perror020/Issue2769b.idr diff --git a/tests/idris2/perror020/expected b/tests/idris2/error/perror020/expected similarity index 100% rename from tests/idris2/perror020/expected rename to tests/idris2/error/perror020/expected diff --git a/tests/idris2/perror020/run b/tests/idris2/error/perror020/run similarity index 63% rename from tests/idris2/perror020/run rename to tests/idris2/error/perror020/run index 0847511ad1..d38ced64b9 100644 --- a/tests/idris2/perror020/run +++ b/tests/idris2/error/perror020/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh check Issue2769.idr check Issue2769b.idr diff --git a/tests/idris2/perror021/Implicit.idr b/tests/idris2/error/perror021/Implicit.idr similarity index 100% rename from tests/idris2/perror021/Implicit.idr rename to tests/idris2/error/perror021/Implicit.idr diff --git a/tests/idris2/perror021/expected b/tests/idris2/error/perror021/expected similarity index 100% rename from tests/idris2/perror021/expected rename to tests/idris2/error/perror021/expected diff --git a/tests/idris2/error/perror021/run b/tests/idris2/error/perror021/run new file mode 100644 index 0000000000..c88212c0ed --- /dev/null +++ b/tests/idris2/error/perror021/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Implicit.idr diff --git a/tests/idris2/perror022/Indent.idr b/tests/idris2/error/perror022/Indent.idr similarity index 100% rename from tests/idris2/perror022/Indent.idr rename to tests/idris2/error/perror022/Indent.idr diff --git a/tests/idris2/perror022/expected b/tests/idris2/error/perror022/expected similarity index 100% rename from tests/idris2/perror022/expected rename to tests/idris2/error/perror022/expected diff --git a/tests/idris2/error/perror022/run b/tests/idris2/error/perror022/run new file mode 100644 index 0000000000..8fe5891b80 --- /dev/null +++ b/tests/idris2/error/perror022/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Indent.idr diff --git a/tests/idris2/perror023/ParseError.idr b/tests/idris2/error/perror023/ParseError.idr similarity index 100% rename from tests/idris2/perror023/ParseError.idr rename to tests/idris2/error/perror023/ParseError.idr diff --git a/tests/idris2/perror023/expected b/tests/idris2/error/perror023/expected similarity index 100% rename from tests/idris2/perror023/expected rename to tests/idris2/error/perror023/expected diff --git a/tests/idris2/error/perror023/run b/tests/idris2/error/perror023/run new file mode 100755 index 0000000000..7b566425d3 --- /dev/null +++ b/tests/idris2/error/perror023/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check ParseError.idr diff --git a/tests/idris2/perror024/ParseError.idr b/tests/idris2/error/perror024/ParseError.idr similarity index 100% rename from tests/idris2/perror024/ParseError.idr rename to tests/idris2/error/perror024/ParseError.idr diff --git a/tests/idris2/perror024/expected b/tests/idris2/error/perror024/expected similarity index 100% rename from tests/idris2/perror024/expected rename to tests/idris2/error/perror024/expected diff --git a/tests/idris2/error/perror024/run b/tests/idris2/error/perror024/run new file mode 100755 index 0000000000..7b566425d3 --- /dev/null +++ b/tests/idris2/error/perror024/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check ParseError.idr diff --git a/tests/idris2/perror025/DataWhere.idr b/tests/idris2/error/perror025/DataWhere.idr similarity index 100% rename from tests/idris2/perror025/DataWhere.idr rename to tests/idris2/error/perror025/DataWhere.idr diff --git a/tests/idris2/perror025/expected b/tests/idris2/error/perror025/expected similarity index 100% rename from tests/idris2/perror025/expected rename to tests/idris2/error/perror025/expected diff --git a/tests/idris2/error/perror025/run b/tests/idris2/error/perror025/run new file mode 100755 index 0000000000..5a73ee3631 --- /dev/null +++ b/tests/idris2/error/perror025/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check DataWhere.idr diff --git a/tests/idris2/perror026/Micro.idr b/tests/idris2/error/perror026/Micro.idr similarity index 100% rename from tests/idris2/perror026/Micro.idr rename to tests/idris2/error/perror026/Micro.idr diff --git a/tests/idris2/perror026/expected b/tests/idris2/error/perror026/expected similarity index 100% rename from tests/idris2/perror026/expected rename to tests/idris2/error/perror026/expected diff --git a/tests/idris2/error/perror026/run b/tests/idris2/error/perror026/run new file mode 100755 index 0000000000..867e2abe20 --- /dev/null +++ b/tests/idris2/error/perror026/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Micro.idr diff --git a/tests/idris2/perror027/Outdent.idr b/tests/idris2/error/perror027/Outdent.idr similarity index 100% rename from tests/idris2/perror027/Outdent.idr rename to tests/idris2/error/perror027/Outdent.idr diff --git a/tests/idris2/perror027/expected b/tests/idris2/error/perror027/expected similarity index 100% rename from tests/idris2/perror027/expected rename to tests/idris2/error/perror027/expected diff --git a/tests/idris2/error/perror027/run b/tests/idris2/error/perror027/run new file mode 100755 index 0000000000..d835fff79c --- /dev/null +++ b/tests/idris2/error/perror027/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Outdent.idr diff --git a/tests/idris2/perror028/LetInDo.idr b/tests/idris2/error/perror028/LetInDo.idr similarity index 100% rename from tests/idris2/perror028/LetInDo.idr rename to tests/idris2/error/perror028/LetInDo.idr diff --git a/tests/idris2/perror028/expected b/tests/idris2/error/perror028/expected similarity index 100% rename from tests/idris2/perror028/expected rename to tests/idris2/error/perror028/expected diff --git a/tests/idris2/error/perror028/run b/tests/idris2/error/perror028/run new file mode 100644 index 0000000000..5e2d211d1c --- /dev/null +++ b/tests/idris2/error/perror028/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check LetInDo.idr diff --git a/tests/idris2/perror029/DelayParse.idr b/tests/idris2/error/perror029/DelayParse.idr similarity index 100% rename from tests/idris2/perror029/DelayParse.idr rename to tests/idris2/error/perror029/DelayParse.idr diff --git a/tests/idris2/perror029/expected b/tests/idris2/error/perror029/expected similarity index 100% rename from tests/idris2/perror029/expected rename to tests/idris2/error/perror029/expected diff --git a/tests/idris2/error/perror029/run b/tests/idris2/error/perror029/run new file mode 100644 index 0000000000..b5bab6bcbf --- /dev/null +++ b/tests/idris2/error/perror029/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check DelayParse.idr diff --git a/tests/idris2/error001/run b/tests/idris2/error001/run deleted file mode 100755 index ee1151b23a..0000000000 --- a/tests/idris2/error001/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check Error.idr diff --git a/tests/idris2/error002/run b/tests/idris2/error002/run deleted file mode 100755 index ee1151b23a..0000000000 --- a/tests/idris2/error002/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check Error.idr diff --git a/tests/idris2/error003/run b/tests/idris2/error003/run deleted file mode 100755 index ee1151b23a..0000000000 --- a/tests/idris2/error003/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check Error.idr diff --git a/tests/idris2/error005/run b/tests/idris2/error005/run deleted file mode 100755 index c788397c11..0000000000 --- a/tests/idris2/error005/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check IfErr.idr diff --git a/tests/idris2/error006/run b/tests/idris2/error006/run deleted file mode 100755 index c788397c11..0000000000 --- a/tests/idris2/error006/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check IfErr.idr diff --git a/tests/idris2/error007/run b/tests/idris2/error007/run deleted file mode 100755 index c1a31e8348..0000000000 --- a/tests/idris2/error007/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check CongErr.idr diff --git a/tests/idris2/error010/run b/tests/idris2/error010/run deleted file mode 100755 index 62d1d3a6b3..0000000000 --- a/tests/idris2/error010/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check Loop.idr diff --git a/tests/idris2/error013/run b/tests/idris2/error013/run deleted file mode 100755 index a7f294e9f2..0000000000 --- a/tests/idris2/error013/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check Issue361.idr diff --git a/tests/idris2/error014/run b/tests/idris2/error014/run deleted file mode 100755 index 4d51f545fc..0000000000 --- a/tests/idris2/error014/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check Issue735.idr diff --git a/tests/idris2/error015/run b/tests/idris2/error015/run deleted file mode 100755 index e659beb822..0000000000 --- a/tests/idris2/error015/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check Issue110.idr diff --git a/tests/idris2/error022/run b/tests/idris2/error022/run deleted file mode 100644 index 6a9cbb2793..0000000000 --- a/tests/idris2/error022/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check UpdateLoc.idr diff --git a/tests/idris2/error024/run b/tests/idris2/error024/run deleted file mode 100755 index 2f0cc21131..0000000000 --- a/tests/idris2/error024/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check Error1.idr diff --git a/tests/idris2/error026/run b/tests/idris2/error026/run deleted file mode 100755 index b2ce794b2f..0000000000 --- a/tests/idris2/error026/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check DoBlockFC.idr diff --git a/tests/idris2/error027/run b/tests/idris2/error027/run deleted file mode 100755 index ce22b719e5..0000000000 --- a/tests/idris2/error027/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -run Issue2950.idr diff --git a/tests/idris2/eta001/run b/tests/idris2/eta001/run deleted file mode 100755 index 48614412aa..0000000000 --- a/tests/idris2/eta001/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check Issue1370.idr diff --git a/tests/idris2/evaluator001/Issue650.idr b/tests/idris2/evaluator/evaluator001/Issue650.idr similarity index 100% rename from tests/idris2/evaluator001/Issue650.idr rename to tests/idris2/evaluator/evaluator001/Issue650.idr diff --git a/tests/idris2/evaluator001/expected b/tests/idris2/evaluator/evaluator001/expected similarity index 100% rename from tests/idris2/evaluator001/expected rename to tests/idris2/evaluator/evaluator001/expected diff --git a/tests/idris2/evaluator/evaluator001/run b/tests/idris2/evaluator/evaluator001/run new file mode 100755 index 0000000000..f441737cbd --- /dev/null +++ b/tests/idris2/evaluator/evaluator001/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Issue650.idr diff --git a/tests/idris2/evaluator002/Lib.idr b/tests/idris2/evaluator/evaluator002/Lib.idr similarity index 100% rename from tests/idris2/evaluator002/Lib.idr rename to tests/idris2/evaluator/evaluator002/Lib.idr diff --git a/tests/idris2/evaluator002/Main.idr b/tests/idris2/evaluator/evaluator002/Main.idr similarity index 100% rename from tests/idris2/evaluator002/Main.idr rename to tests/idris2/evaluator/evaluator002/Main.idr diff --git a/tests/idris2/evaluator002/expected b/tests/idris2/evaluator/evaluator002/expected similarity index 100% rename from tests/idris2/evaluator002/expected rename to tests/idris2/evaluator/evaluator002/expected diff --git a/tests/idris2/evaluator002/input b/tests/idris2/evaluator/evaluator002/input similarity index 100% rename from tests/idris2/evaluator002/input rename to tests/idris2/evaluator/evaluator002/input diff --git a/tests/idris2/evaluator002/run b/tests/idris2/evaluator/evaluator002/run similarity index 64% rename from tests/idris2/evaluator002/run rename to tests/idris2/evaluator/evaluator002/run index 2ab3dca6ed..6d7c87c8f7 100644 --- a/tests/idris2/evaluator002/run +++ b/tests/idris2/evaluator/evaluator002/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 --log eval.stuck:5 Main.idr < input diff --git a/tests/idris2/evaluator003/Issue705.idr b/tests/idris2/evaluator/evaluator003/Issue705.idr similarity index 100% rename from tests/idris2/evaluator003/Issue705.idr rename to tests/idris2/evaluator/evaluator003/Issue705.idr diff --git a/tests/idris2/evaluator003/expected b/tests/idris2/evaluator/evaluator003/expected similarity index 100% rename from tests/idris2/evaluator003/expected rename to tests/idris2/evaluator/evaluator003/expected diff --git a/tests/idris2/evaluator003/input b/tests/idris2/evaluator/evaluator003/input similarity index 100% rename from tests/idris2/evaluator003/input rename to tests/idris2/evaluator/evaluator003/input diff --git a/tests/idris2/evaluator003/run b/tests/idris2/evaluator/evaluator003/run similarity index 54% rename from tests/idris2/evaluator003/run rename to tests/idris2/evaluator/evaluator003/run index 8c9eec4418..e806722d82 100644 --- a/tests/idris2/evaluator003/run +++ b/tests/idris2/evaluator/evaluator003/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 Issue705.idr < input diff --git a/tests/idris2/evaluator004/Issue1282.idr b/tests/idris2/evaluator/evaluator004/Issue1282.idr similarity index 100% rename from tests/idris2/evaluator004/Issue1282.idr rename to tests/idris2/evaluator/evaluator004/Issue1282.idr diff --git a/tests/idris2/evaluator004/expected b/tests/idris2/evaluator/evaluator004/expected similarity index 100% rename from tests/idris2/evaluator004/expected rename to tests/idris2/evaluator/evaluator004/expected diff --git a/tests/idris2/evaluator004/input b/tests/idris2/evaluator/evaluator004/input similarity index 100% rename from tests/idris2/evaluator004/input rename to tests/idris2/evaluator/evaluator004/input diff --git a/tests/idris2/evaluator004/run b/tests/idris2/evaluator/evaluator004/run similarity index 55% rename from tests/idris2/evaluator004/run rename to tests/idris2/evaluator/evaluator004/run index a5bc53cb97..004ef2a5ed 100644 --- a/tests/idris2/evaluator004/run +++ b/tests/idris2/evaluator/evaluator004/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 Issue1282.idr < input diff --git a/tests/idris2/interpreter001/expected b/tests/idris2/evaluator/interpreter001/expected similarity index 100% rename from tests/idris2/interpreter001/expected rename to tests/idris2/evaluator/interpreter001/expected diff --git a/tests/idris2/interpreter001/input b/tests/idris2/evaluator/interpreter001/input similarity index 100% rename from tests/idris2/interpreter001/input rename to tests/idris2/evaluator/interpreter001/input diff --git a/tests/idris2/evaluator/interpreter001/run b/tests/idris2/evaluator/interpreter001/run new file mode 100755 index 0000000000..404d7d47c6 --- /dev/null +++ b/tests/idris2/evaluator/interpreter001/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +idris2 < input diff --git a/tests/idris2/interpreter002/expected b/tests/idris2/evaluator/interpreter002/expected similarity index 100% rename from tests/idris2/interpreter002/expected rename to tests/idris2/evaluator/interpreter002/expected diff --git a/tests/idris2/interpreter002/input b/tests/idris2/evaluator/interpreter002/input similarity index 100% rename from tests/idris2/interpreter002/input rename to tests/idris2/evaluator/interpreter002/input diff --git a/tests/idris2/interpreter002/run b/tests/idris2/evaluator/interpreter002/run similarity index 54% rename from tests/idris2/interpreter002/run rename to tests/idris2/evaluator/interpreter002/run index 72182e4d4e..0586563b86 100755 --- a/tests/idris2/interpreter002/run +++ b/tests/idris2/evaluator/interpreter002/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 --no-prelude < input diff --git a/tests/idris2/interpreter003/expected b/tests/idris2/evaluator/interpreter003/expected similarity index 100% rename from tests/idris2/interpreter003/expected rename to tests/idris2/evaluator/interpreter003/expected diff --git a/tests/idris2/interpreter003/input b/tests/idris2/evaluator/interpreter003/input similarity index 100% rename from tests/idris2/interpreter003/input rename to tests/idris2/evaluator/interpreter003/input diff --git a/tests/idris2/evaluator/interpreter003/run b/tests/idris2/evaluator/interpreter003/run new file mode 100755 index 0000000000..404d7d47c6 --- /dev/null +++ b/tests/idris2/evaluator/interpreter003/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +idris2 < input diff --git a/tests/idris2/interpreter004/expected b/tests/idris2/evaluator/interpreter004/expected similarity index 100% rename from tests/idris2/interpreter004/expected rename to tests/idris2/evaluator/interpreter004/expected diff --git a/tests/idris2/interpreter004/input b/tests/idris2/evaluator/interpreter004/input similarity index 100% rename from tests/idris2/interpreter004/input rename to tests/idris2/evaluator/interpreter004/input diff --git a/tests/idris2/evaluator/interpreter004/run b/tests/idris2/evaluator/interpreter004/run new file mode 100755 index 0000000000..404d7d47c6 --- /dev/null +++ b/tests/idris2/evaluator/interpreter004/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +idris2 < input diff --git a/tests/idris2/interpreter005/Issue37.idr b/tests/idris2/evaluator/interpreter005/Issue37.idr similarity index 100% rename from tests/idris2/interpreter005/Issue37.idr rename to tests/idris2/evaluator/interpreter005/Issue37.idr diff --git a/tests/idris2/interpreter005/Issue37.lidr b/tests/idris2/evaluator/interpreter005/Issue37.lidr similarity index 100% rename from tests/idris2/interpreter005/Issue37.lidr rename to tests/idris2/evaluator/interpreter005/Issue37.lidr diff --git a/tests/idris2/interpreter005/expected b/tests/idris2/evaluator/interpreter005/expected similarity index 100% rename from tests/idris2/interpreter005/expected rename to tests/idris2/evaluator/interpreter005/expected diff --git a/tests/idris2/interpreter005/input b/tests/idris2/evaluator/interpreter005/input similarity index 100% rename from tests/idris2/interpreter005/input rename to tests/idris2/evaluator/interpreter005/input diff --git a/tests/idris2/interpreter005/run b/tests/idris2/evaluator/interpreter005/run similarity index 54% rename from tests/idris2/interpreter005/run rename to tests/idris2/evaluator/interpreter005/run index efb1e49621..200981ccb3 100755 --- a/tests/idris2/interpreter005/run +++ b/tests/idris2/evaluator/interpreter005/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 Issue37.lidr < input diff --git a/tests/idris2/interpreter006/expected b/tests/idris2/evaluator/interpreter006/expected similarity index 100% rename from tests/idris2/interpreter006/expected rename to tests/idris2/evaluator/interpreter006/expected diff --git a/tests/idris2/interpreter006/input b/tests/idris2/evaluator/interpreter006/input similarity index 100% rename from tests/idris2/interpreter006/input rename to tests/idris2/evaluator/interpreter006/input diff --git a/tests/idris2/evaluator/interpreter006/run b/tests/idris2/evaluator/interpreter006/run new file mode 100755 index 0000000000..404d7d47c6 --- /dev/null +++ b/tests/idris2/evaluator/interpreter006/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +idris2 < input diff --git a/tests/idris2/interpreter007/expected b/tests/idris2/evaluator/interpreter007/expected similarity index 100% rename from tests/idris2/interpreter007/expected rename to tests/idris2/evaluator/interpreter007/expected diff --git a/tests/idris2/interpreter007/input b/tests/idris2/evaluator/interpreter007/input similarity index 100% rename from tests/idris2/interpreter007/input rename to tests/idris2/evaluator/interpreter007/input diff --git a/tests/idris2/evaluator/interpreter007/run b/tests/idris2/evaluator/interpreter007/run new file mode 100755 index 0000000000..404d7d47c6 --- /dev/null +++ b/tests/idris2/evaluator/interpreter007/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +idris2 < input diff --git a/tests/idris2/interpreter008/Issue2041.idr b/tests/idris2/evaluator/interpreter008/Issue2041.idr similarity index 100% rename from tests/idris2/interpreter008/Issue2041.idr rename to tests/idris2/evaluator/interpreter008/Issue2041.idr diff --git a/tests/idris2/interpreter008/expected b/tests/idris2/evaluator/interpreter008/expected similarity index 100% rename from tests/idris2/interpreter008/expected rename to tests/idris2/evaluator/interpreter008/expected diff --git a/tests/idris2/interpreter008/input b/tests/idris2/evaluator/interpreter008/input similarity index 100% rename from tests/idris2/interpreter008/input rename to tests/idris2/evaluator/interpreter008/input diff --git a/tests/idris2/interpreter008/run b/tests/idris2/evaluator/interpreter008/run similarity index 55% rename from tests/idris2/interpreter008/run rename to tests/idris2/evaluator/interpreter008/run index d2d8c793c9..e110c8dedb 100755 --- a/tests/idris2/interpreter008/run +++ b/tests/idris2/evaluator/interpreter008/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 Issue2041.idr < input diff --git a/tests/idris2/spec001/Desc.idr b/tests/idris2/evaluator/spec001/Desc.idr similarity index 100% rename from tests/idris2/spec001/Desc.idr rename to tests/idris2/evaluator/spec001/Desc.idr diff --git a/tests/idris2/spec001/Desc2.idr b/tests/idris2/evaluator/spec001/Desc2.idr similarity index 100% rename from tests/idris2/spec001/Desc2.idr rename to tests/idris2/evaluator/spec001/Desc2.idr diff --git a/tests/idris2/spec001/Identity.idr b/tests/idris2/evaluator/spec001/Identity.idr similarity index 100% rename from tests/idris2/spec001/Identity.idr rename to tests/idris2/evaluator/spec001/Identity.idr diff --git a/tests/idris2/spec001/Mult3.idr b/tests/idris2/evaluator/spec001/Mult3.idr similarity index 100% rename from tests/idris2/spec001/Mult3.idr rename to tests/idris2/evaluator/spec001/Mult3.idr diff --git a/tests/idris2/spec001/expected b/tests/idris2/evaluator/spec001/expected similarity index 100% rename from tests/idris2/spec001/expected rename to tests/idris2/evaluator/spec001/expected diff --git a/tests/idris2/spec001/run b/tests/idris2/evaluator/spec001/run similarity index 91% rename from tests/idris2/spec001/run rename to tests/idris2/evaluator/spec001/run index 7306f953da..11aaa6a137 100644 --- a/tests/idris2/spec001/run +++ b/tests/idris2/evaluator/spec001/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 --log specialise:5 -c Mult3.idr idris2 --cg node -o mult3.js -c Mult3.idr diff --git a/tests/idris2/evaluator001/run b/tests/idris2/evaluator001/run deleted file mode 100755 index f04a7b6e84..0000000000 --- a/tests/idris2/evaluator001/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check Issue650.idr diff --git a/tests/idris2/failing001/Fail.idr b/tests/idris2/failing/failing001/Fail.idr similarity index 100% rename from tests/idris2/failing001/Fail.idr rename to tests/idris2/failing/failing001/Fail.idr diff --git a/tests/idris2/failing001/expected b/tests/idris2/failing/failing001/expected similarity index 100% rename from tests/idris2/failing001/expected rename to tests/idris2/failing/failing001/expected diff --git a/tests/idris2/failing/failing001/run b/tests/idris2/failing/failing001/run new file mode 100755 index 0000000000..d66b29bef7 --- /dev/null +++ b/tests/idris2/failing/failing001/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Fail.idr diff --git a/tests/idris2/failing002/FailingBug.idr b/tests/idris2/failing/failing002/FailingBug.idr similarity index 100% rename from tests/idris2/failing002/FailingBug.idr rename to tests/idris2/failing/failing002/FailingBug.idr diff --git a/tests/idris2/failing002/expected b/tests/idris2/failing/failing002/expected similarity index 100% rename from tests/idris2/failing002/expected rename to tests/idris2/failing/failing002/expected diff --git a/tests/idris2/failing002/run b/tests/idris2/failing/failing002/run similarity index 58% rename from tests/idris2/failing002/run rename to tests/idris2/failing/failing002/run index 78810e46a9..96370c1440 100755 --- a/tests/idris2/failing002/run +++ b/tests/idris2/failing/failing002/run @@ -1,8 +1,8 @@ -. ../../testutils.sh +. ../../../testutils.sh check FailingBug.idr -. ../../testutils.sh +. ../../../testutils.sh check --console-width 20 FailingBug.idr -. ../../testutils.sh +. ../../../testutils.sh check --console-width 80 FailingBug.idr diff --git a/tests/idris2/failing003/FailingTotality.idr b/tests/idris2/failing/failing003/FailingTotality.idr similarity index 100% rename from tests/idris2/failing003/FailingTotality.idr rename to tests/idris2/failing/failing003/FailingTotality.idr diff --git a/tests/idris2/failing003/expected b/tests/idris2/failing/failing003/expected similarity index 100% rename from tests/idris2/failing003/expected rename to tests/idris2/failing/failing003/expected diff --git a/tests/idris2/failing003/run b/tests/idris2/failing/failing003/run similarity index 52% rename from tests/idris2/failing003/run rename to tests/idris2/failing/failing003/run index 4c16dd5097..2745393837 100755 --- a/tests/idris2/failing003/run +++ b/tests/idris2/failing/failing003/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh check FailingTotality.idr diff --git a/tests/idris2/failing004/Issue2821.idr b/tests/idris2/failing/failing004/Issue2821.idr similarity index 100% rename from tests/idris2/failing004/Issue2821.idr rename to tests/idris2/failing/failing004/Issue2821.idr diff --git a/tests/idris2/failing004/expected b/tests/idris2/failing/failing004/expected similarity index 100% rename from tests/idris2/failing004/expected rename to tests/idris2/failing/failing004/expected diff --git a/tests/idris2/failing/failing004/run b/tests/idris2/failing/failing004/run new file mode 100755 index 0000000000..3c6b7a162d --- /dev/null +++ b/tests/idris2/failing/failing004/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Issue2821.idr diff --git a/tests/idris2/failing001/run b/tests/idris2/failing001/run deleted file mode 100755 index c296850235..0000000000 --- a/tests/idris2/failing001/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check Fail.idr diff --git a/tests/idris2/failing004/run b/tests/idris2/failing004/run deleted file mode 100755 index fe450082c0..0000000000 --- a/tests/idris2/failing004/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check Issue2821.idr diff --git a/tests/idris2/idiom001/run b/tests/idris2/idiom001/run deleted file mode 100644 index c6e6f2ab79..0000000000 --- a/tests/idris2/idiom001/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -run Main.idr diff --git a/tests/idris2/import005/run b/tests/idris2/import005/run deleted file mode 100644 index 867d00c171..0000000000 --- a/tests/idris2/import005/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -idris2 As.idr < input diff --git a/tests/idris2/import007/run b/tests/idris2/import007/run deleted file mode 100644 index fc9dabdf6a..0000000000 --- a/tests/idris2/import007/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check Mod.idr diff --git a/tests/idris2/interactive001/LocType.idr b/tests/idris2/interactive/interactive001/LocType.idr similarity index 100% rename from tests/idris2/interactive001/LocType.idr rename to tests/idris2/interactive/interactive001/LocType.idr diff --git a/tests/idris2/interactive001/expected b/tests/idris2/interactive/interactive001/expected similarity index 100% rename from tests/idris2/interactive001/expected rename to tests/idris2/interactive/interactive001/expected diff --git a/tests/idris2/interactive001/input b/tests/idris2/interactive/interactive001/input similarity index 100% rename from tests/idris2/interactive001/input rename to tests/idris2/interactive/interactive001/input diff --git a/tests/idris2/interactive001/run b/tests/idris2/interactive/interactive001/run similarity index 53% rename from tests/idris2/interactive001/run rename to tests/idris2/interactive/interactive001/run index 6d4602837f..ccfdefc4e6 100755 --- a/tests/idris2/interactive001/run +++ b/tests/idris2/interactive/interactive001/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 LocType.idr < input diff --git a/tests/idris2/interactive002/IEdit.idr b/tests/idris2/interactive/interactive002/IEdit.idr similarity index 100% rename from tests/idris2/interactive002/IEdit.idr rename to tests/idris2/interactive/interactive002/IEdit.idr diff --git a/tests/idris2/interactive002/expected b/tests/idris2/interactive/interactive002/expected similarity index 100% rename from tests/idris2/interactive002/expected rename to tests/idris2/interactive/interactive002/expected diff --git a/tests/idris2/interactive002/input b/tests/idris2/interactive/interactive002/input similarity index 100% rename from tests/idris2/interactive002/input rename to tests/idris2/interactive/interactive002/input diff --git a/tests/idris2/interactive002/run b/tests/idris2/interactive/interactive002/run similarity index 52% rename from tests/idris2/interactive002/run rename to tests/idris2/interactive/interactive002/run index ec050c77b1..737446ffd5 100755 --- a/tests/idris2/interactive002/run +++ b/tests/idris2/interactive/interactive002/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 IEdit.idr < input diff --git a/tests/idris2/interactive003/IEdit.idr b/tests/idris2/interactive/interactive003/IEdit.idr similarity index 100% rename from tests/idris2/interactive003/IEdit.idr rename to tests/idris2/interactive/interactive003/IEdit.idr diff --git a/tests/idris2/interactive003/IEdit2.idr b/tests/idris2/interactive/interactive003/IEdit2.idr similarity index 100% rename from tests/idris2/interactive003/IEdit2.idr rename to tests/idris2/interactive/interactive003/IEdit2.idr diff --git a/tests/idris2/interactive003/expected b/tests/idris2/interactive/interactive003/expected similarity index 100% rename from tests/idris2/interactive003/expected rename to tests/idris2/interactive/interactive003/expected diff --git a/tests/idris2/interactive003/input b/tests/idris2/interactive/interactive003/input similarity index 100% rename from tests/idris2/interactive003/input rename to tests/idris2/interactive/interactive003/input diff --git a/tests/idris2/interactive003/input2 b/tests/idris2/interactive/interactive003/input2 similarity index 100% rename from tests/idris2/interactive003/input2 rename to tests/idris2/interactive/interactive003/input2 diff --git a/tests/idris2/interactive003/run b/tests/idris2/interactive/interactive003/run similarity index 68% rename from tests/idris2/interactive003/run rename to tests/idris2/interactive/interactive003/run index 45d7c58b5b..07412ca3dd 100755 --- a/tests/idris2/interactive003/run +++ b/tests/idris2/interactive/interactive003/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 IEdit.idr < input idris2 IEdit2.idr < input2 diff --git a/tests/idris2/interactive004/IEdit.idr b/tests/idris2/interactive/interactive004/IEdit.idr similarity index 100% rename from tests/idris2/interactive004/IEdit.idr rename to tests/idris2/interactive/interactive004/IEdit.idr diff --git a/tests/idris2/interactive004/expected b/tests/idris2/interactive/interactive004/expected similarity index 100% rename from tests/idris2/interactive004/expected rename to tests/idris2/interactive/interactive004/expected diff --git a/tests/idris2/interactive004/input b/tests/idris2/interactive/interactive004/input similarity index 100% rename from tests/idris2/interactive004/input rename to tests/idris2/interactive/interactive004/input diff --git a/tests/idris2/interactive006/run b/tests/idris2/interactive/interactive004/run similarity index 52% rename from tests/idris2/interactive006/run rename to tests/idris2/interactive/interactive004/run index ec050c77b1..737446ffd5 100755 --- a/tests/idris2/interactive006/run +++ b/tests/idris2/interactive/interactive004/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 IEdit.idr < input diff --git a/tests/idris2/interactive005/IEdit.idr b/tests/idris2/interactive/interactive005/IEdit.idr similarity index 100% rename from tests/idris2/interactive005/IEdit.idr rename to tests/idris2/interactive/interactive005/IEdit.idr diff --git a/tests/idris2/interactive005/expected b/tests/idris2/interactive/interactive005/expected similarity index 100% rename from tests/idris2/interactive005/expected rename to tests/idris2/interactive/interactive005/expected diff --git a/tests/idris2/interactive005/input b/tests/idris2/interactive/interactive005/input similarity index 100% rename from tests/idris2/interactive005/input rename to tests/idris2/interactive/interactive005/input diff --git a/tests/idris2/interactive004/run b/tests/idris2/interactive/interactive005/run similarity index 52% rename from tests/idris2/interactive004/run rename to tests/idris2/interactive/interactive005/run index ec050c77b1..737446ffd5 100755 --- a/tests/idris2/interactive004/run +++ b/tests/idris2/interactive/interactive005/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 IEdit.idr < input diff --git a/tests/idris2/interactive006/IEdit.idr b/tests/idris2/interactive/interactive006/IEdit.idr similarity index 100% rename from tests/idris2/interactive006/IEdit.idr rename to tests/idris2/interactive/interactive006/IEdit.idr diff --git a/tests/idris2/interactive006/expected b/tests/idris2/interactive/interactive006/expected similarity index 100% rename from tests/idris2/interactive006/expected rename to tests/idris2/interactive/interactive006/expected diff --git a/tests/idris2/interactive006/input b/tests/idris2/interactive/interactive006/input similarity index 100% rename from tests/idris2/interactive006/input rename to tests/idris2/interactive/interactive006/input diff --git a/tests/idris2/interactive005/run b/tests/idris2/interactive/interactive006/run similarity index 52% rename from tests/idris2/interactive005/run rename to tests/idris2/interactive/interactive006/run index ec050c77b1..737446ffd5 100755 --- a/tests/idris2/interactive005/run +++ b/tests/idris2/interactive/interactive006/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 IEdit.idr < input diff --git a/tests/idris2/interactive007/IEdit.idr b/tests/idris2/interactive/interactive007/IEdit.idr similarity index 100% rename from tests/idris2/interactive007/IEdit.idr rename to tests/idris2/interactive/interactive007/IEdit.idr diff --git a/tests/idris2/interactive007/expected b/tests/idris2/interactive/interactive007/expected similarity index 100% rename from tests/idris2/interactive007/expected rename to tests/idris2/interactive/interactive007/expected diff --git a/tests/idris2/interactive007/input b/tests/idris2/interactive/interactive007/input similarity index 100% rename from tests/idris2/interactive007/input rename to tests/idris2/interactive/interactive007/input diff --git a/tests/idris2/interactive/interactive007/run b/tests/idris2/interactive/interactive007/run new file mode 100755 index 0000000000..737446ffd5 --- /dev/null +++ b/tests/idris2/interactive/interactive007/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +idris2 IEdit.idr < input diff --git a/tests/idris2/interactive008/IEdit.idr b/tests/idris2/interactive/interactive008/IEdit.idr similarity index 100% rename from tests/idris2/interactive008/IEdit.idr rename to tests/idris2/interactive/interactive008/IEdit.idr diff --git a/tests/idris2/interactive008/expected b/tests/idris2/interactive/interactive008/expected similarity index 100% rename from tests/idris2/interactive008/expected rename to tests/idris2/interactive/interactive008/expected diff --git a/tests/idris2/interactive008/input b/tests/idris2/interactive/interactive008/input similarity index 100% rename from tests/idris2/interactive008/input rename to tests/idris2/interactive/interactive008/input diff --git a/tests/idris2/interactive/interactive008/run b/tests/idris2/interactive/interactive008/run new file mode 100755 index 0000000000..737446ffd5 --- /dev/null +++ b/tests/idris2/interactive/interactive008/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +idris2 IEdit.idr < input diff --git a/tests/idris2/interactive009/Door.idr b/tests/idris2/interactive/interactive009/Door.idr similarity index 100% rename from tests/idris2/interactive009/Door.idr rename to tests/idris2/interactive/interactive009/Door.idr diff --git a/tests/idris2/interactive009/expected b/tests/idris2/interactive/interactive009/expected similarity index 100% rename from tests/idris2/interactive009/expected rename to tests/idris2/interactive/interactive009/expected diff --git a/tests/idris2/interactive009/input b/tests/idris2/interactive/interactive009/input similarity index 100% rename from tests/idris2/interactive009/input rename to tests/idris2/interactive/interactive009/input diff --git a/tests/idris2/interactive009/run b/tests/idris2/interactive/interactive009/run similarity index 51% rename from tests/idris2/interactive009/run rename to tests/idris2/interactive/interactive009/run index e6c63e9986..402b977549 100755 --- a/tests/idris2/interactive009/run +++ b/tests/idris2/interactive/interactive009/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 Door.idr < input diff --git a/tests/idris2/interactive010/IEdit.idr b/tests/idris2/interactive/interactive010/IEdit.idr similarity index 100% rename from tests/idris2/interactive010/IEdit.idr rename to tests/idris2/interactive/interactive010/IEdit.idr diff --git a/tests/idris2/interactive010/expected b/tests/idris2/interactive/interactive010/expected similarity index 100% rename from tests/idris2/interactive010/expected rename to tests/idris2/interactive/interactive010/expected diff --git a/tests/idris2/interactive010/input b/tests/idris2/interactive/interactive010/input similarity index 100% rename from tests/idris2/interactive010/input rename to tests/idris2/interactive/interactive010/input diff --git a/tests/idris2/interactive/interactive010/run b/tests/idris2/interactive/interactive010/run new file mode 100755 index 0000000000..737446ffd5 --- /dev/null +++ b/tests/idris2/interactive/interactive010/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +idris2 IEdit.idr < input diff --git a/tests/idris2/interactive011/IEdit.idr b/tests/idris2/interactive/interactive011/IEdit.idr similarity index 100% rename from tests/idris2/interactive011/IEdit.idr rename to tests/idris2/interactive/interactive011/IEdit.idr diff --git a/tests/idris2/interactive011/expected b/tests/idris2/interactive/interactive011/expected similarity index 100% rename from tests/idris2/interactive011/expected rename to tests/idris2/interactive/interactive011/expected diff --git a/tests/idris2/interactive011/input b/tests/idris2/interactive/interactive011/input similarity index 100% rename from tests/idris2/interactive011/input rename to tests/idris2/interactive/interactive011/input diff --git a/tests/idris2/interactive/interactive011/run b/tests/idris2/interactive/interactive011/run new file mode 100755 index 0000000000..737446ffd5 --- /dev/null +++ b/tests/idris2/interactive/interactive011/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +idris2 IEdit.idr < input diff --git a/tests/idris2/interactive012/WithLift.idr b/tests/idris2/interactive/interactive012/WithLift.idr similarity index 100% rename from tests/idris2/interactive012/WithLift.idr rename to tests/idris2/interactive/interactive012/WithLift.idr diff --git a/tests/idris2/interactive012/expected b/tests/idris2/interactive/interactive012/expected similarity index 100% rename from tests/idris2/interactive012/expected rename to tests/idris2/interactive/interactive012/expected diff --git a/tests/idris2/interactive012/input b/tests/idris2/interactive/interactive012/input similarity index 100% rename from tests/idris2/interactive012/input rename to tests/idris2/interactive/interactive012/input diff --git a/tests/idris2/interactive012/run b/tests/idris2/interactive/interactive012/run similarity index 54% rename from tests/idris2/interactive012/run rename to tests/idris2/interactive/interactive012/run index bf4884fe0e..52309c60da 100755 --- a/tests/idris2/interactive012/run +++ b/tests/idris2/interactive/interactive012/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 WithLift.idr < input diff --git a/tests/idris2/interactive013/Spacing.idr b/tests/idris2/interactive/interactive013/Spacing.idr similarity index 100% rename from tests/idris2/interactive013/Spacing.idr rename to tests/idris2/interactive/interactive013/Spacing.idr diff --git a/tests/idris2/interactive013/expected b/tests/idris2/interactive/interactive013/expected similarity index 100% rename from tests/idris2/interactive013/expected rename to tests/idris2/interactive/interactive013/expected diff --git a/tests/idris2/interactive013/input b/tests/idris2/interactive/interactive013/input similarity index 100% rename from tests/idris2/interactive013/input rename to tests/idris2/interactive/interactive013/input diff --git a/tests/idris2/interactive013/run b/tests/idris2/interactive/interactive013/run similarity index 53% rename from tests/idris2/interactive013/run rename to tests/idris2/interactive/interactive013/run index 8f9217658b..35b250fe44 100755 --- a/tests/idris2/interactive013/run +++ b/tests/idris2/interactive/interactive013/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 Spacing.idr < input diff --git a/tests/idris2/interactive014/case.idr b/tests/idris2/interactive/interactive014/case.idr similarity index 100% rename from tests/idris2/interactive014/case.idr rename to tests/idris2/interactive/interactive014/case.idr diff --git a/tests/idris2/interactive014/expected b/tests/idris2/interactive/interactive014/expected similarity index 100% rename from tests/idris2/interactive014/expected rename to tests/idris2/interactive/interactive014/expected diff --git a/tests/idris2/interactive014/input b/tests/idris2/interactive/interactive014/input similarity index 100% rename from tests/idris2/interactive014/input rename to tests/idris2/interactive/interactive014/input diff --git a/tests/idris2/interactive014/run b/tests/idris2/interactive/interactive014/run similarity index 51% rename from tests/idris2/interactive014/run rename to tests/idris2/interactive/interactive014/run index f48eb78c45..5c8611c14e 100755 --- a/tests/idris2/interactive014/run +++ b/tests/idris2/interactive/interactive014/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 case.idr < input diff --git a/tests/idris2/interactive015/IEdit.idr b/tests/idris2/interactive/interactive015/IEdit.idr similarity index 100% rename from tests/idris2/interactive015/IEdit.idr rename to tests/idris2/interactive/interactive015/IEdit.idr diff --git a/tests/idris2/interactive015/expected b/tests/idris2/interactive/interactive015/expected similarity index 100% rename from tests/idris2/interactive015/expected rename to tests/idris2/interactive/interactive015/expected diff --git a/tests/idris2/interactive015/input b/tests/idris2/interactive/interactive015/input similarity index 100% rename from tests/idris2/interactive015/input rename to tests/idris2/interactive/interactive015/input diff --git a/tests/idris2/interactive/interactive015/run b/tests/idris2/interactive/interactive015/run new file mode 100755 index 0000000000..737446ffd5 --- /dev/null +++ b/tests/idris2/interactive/interactive015/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +idris2 IEdit.idr < input diff --git a/tests/idris2/interactive016/Cont.idr b/tests/idris2/interactive/interactive016/Cont.idr similarity index 100% rename from tests/idris2/interactive016/Cont.idr rename to tests/idris2/interactive/interactive016/Cont.idr diff --git a/tests/idris2/interactive016/expected b/tests/idris2/interactive/interactive016/expected similarity index 100% rename from tests/idris2/interactive016/expected rename to tests/idris2/interactive/interactive016/expected diff --git a/tests/idris2/interactive016/input b/tests/idris2/interactive/interactive016/input similarity index 100% rename from tests/idris2/interactive016/input rename to tests/idris2/interactive/interactive016/input diff --git a/tests/idris2/interactive016/run b/tests/idris2/interactive/interactive016/run similarity index 51% rename from tests/idris2/interactive016/run rename to tests/idris2/interactive/interactive016/run index bd22d7ef88..4da9af0ffd 100755 --- a/tests/idris2/interactive016/run +++ b/tests/idris2/interactive/interactive016/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 Cont.idr < input diff --git a/tests/idris2/interactive017/RLE.idr b/tests/idris2/interactive/interactive017/RLE.idr similarity index 100% rename from tests/idris2/interactive017/RLE.idr rename to tests/idris2/interactive/interactive017/RLE.idr diff --git a/tests/idris2/interactive017/expected b/tests/idris2/interactive/interactive017/expected similarity index 100% rename from tests/idris2/interactive017/expected rename to tests/idris2/interactive/interactive017/expected diff --git a/tests/idris2/interactive017/input b/tests/idris2/interactive/interactive017/input similarity index 100% rename from tests/idris2/interactive017/input rename to tests/idris2/interactive/interactive017/input diff --git a/tests/idris2/interactive017/run b/tests/idris2/interactive/interactive017/run similarity index 50% rename from tests/idris2/interactive017/run rename to tests/idris2/interactive/interactive017/run index ddc5e342c0..486ac530ab 100755 --- a/tests/idris2/interactive017/run +++ b/tests/idris2/interactive/interactive017/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 RLE.idr < input diff --git a/tests/idris2/interactive018/PlusPrf.idr b/tests/idris2/interactive/interactive018/PlusPrf.idr similarity index 100% rename from tests/idris2/interactive018/PlusPrf.idr rename to tests/idris2/interactive/interactive018/PlusPrf.idr diff --git a/tests/idris2/interactive018/expected b/tests/idris2/interactive/interactive018/expected similarity index 100% rename from tests/idris2/interactive018/expected rename to tests/idris2/interactive/interactive018/expected diff --git a/tests/idris2/interactive018/input b/tests/idris2/interactive/interactive018/input similarity index 100% rename from tests/idris2/interactive018/input rename to tests/idris2/interactive/interactive018/input diff --git a/tests/idris2/interactive018/run b/tests/idris2/interactive/interactive018/run similarity index 53% rename from tests/idris2/interactive018/run rename to tests/idris2/interactive/interactive018/run index 7ae00127b5..d127b87e28 100755 --- a/tests/idris2/interactive018/run +++ b/tests/idris2/interactive/interactive018/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 PlusPrf.idr < input diff --git a/tests/idris2/interactive019/TypeSearch.idr b/tests/idris2/interactive/interactive019/TypeSearch.idr similarity index 100% rename from tests/idris2/interactive019/TypeSearch.idr rename to tests/idris2/interactive/interactive019/TypeSearch.idr diff --git a/tests/idris2/interactive019/expected b/tests/idris2/interactive/interactive019/expected similarity index 100% rename from tests/idris2/interactive019/expected rename to tests/idris2/interactive/interactive019/expected diff --git a/tests/idris2/interactive019/input b/tests/idris2/interactive/interactive019/input similarity index 100% rename from tests/idris2/interactive019/input rename to tests/idris2/interactive/interactive019/input diff --git a/tests/idris2/interactive019/run b/tests/idris2/interactive/interactive019/run similarity index 56% rename from tests/idris2/interactive019/run rename to tests/idris2/interactive/interactive019/run index 2a211674dd..f3d38b2c9b 100755 --- a/tests/idris2/interactive019/run +++ b/tests/idris2/interactive/interactive019/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 TypeSearch.idr < input diff --git a/tests/idris2/interactive020/Issue835.idr b/tests/idris2/interactive/interactive020/Issue835.idr similarity index 100% rename from tests/idris2/interactive020/Issue835.idr rename to tests/idris2/interactive/interactive020/Issue835.idr diff --git a/tests/idris2/interactive020/expected b/tests/idris2/interactive/interactive020/expected similarity index 100% rename from tests/idris2/interactive020/expected rename to tests/idris2/interactive/interactive020/expected diff --git a/tests/idris2/interactive020/input b/tests/idris2/interactive/interactive020/input similarity index 100% rename from tests/idris2/interactive020/input rename to tests/idris2/interactive/interactive020/input diff --git a/tests/idris2/interactive020/run b/tests/idris2/interactive/interactive020/run similarity index 54% rename from tests/idris2/interactive020/run rename to tests/idris2/interactive/interactive020/run index a86589613d..af8815711a 100755 --- a/tests/idris2/interactive020/run +++ b/tests/idris2/interactive/interactive020/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 Issue835.idr < input diff --git a/tests/idris2/interactive021/TypeAtDoNotation.idr b/tests/idris2/interactive/interactive021/TypeAtDoNotation.idr similarity index 100% rename from tests/idris2/interactive021/TypeAtDoNotation.idr rename to tests/idris2/interactive/interactive021/TypeAtDoNotation.idr diff --git a/tests/idris2/interactive021/expected b/tests/idris2/interactive/interactive021/expected similarity index 100% rename from tests/idris2/interactive021/expected rename to tests/idris2/interactive/interactive021/expected diff --git a/tests/idris2/interactive021/input b/tests/idris2/interactive/interactive021/input similarity index 100% rename from tests/idris2/interactive021/input rename to tests/idris2/interactive/interactive021/input diff --git a/tests/idris2/interactive021/run b/tests/idris2/interactive/interactive021/run similarity index 60% rename from tests/idris2/interactive021/run rename to tests/idris2/interactive/interactive021/run index 86c2a5d98b..9973896f3c 100644 --- a/tests/idris2/interactive021/run +++ b/tests/idris2/interactive/interactive021/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 TypeAtDoNotation.idr < input diff --git a/tests/idris2/interactive022/TypeAtBangSyntax.idr b/tests/idris2/interactive/interactive022/TypeAtBangSyntax.idr similarity index 100% rename from tests/idris2/interactive022/TypeAtBangSyntax.idr rename to tests/idris2/interactive/interactive022/TypeAtBangSyntax.idr diff --git a/tests/idris2/interactive022/expected b/tests/idris2/interactive/interactive022/expected similarity index 100% rename from tests/idris2/interactive022/expected rename to tests/idris2/interactive/interactive022/expected diff --git a/tests/idris2/interactive022/input b/tests/idris2/interactive/interactive022/input similarity index 100% rename from tests/idris2/interactive022/input rename to tests/idris2/interactive/interactive022/input diff --git a/tests/idris2/interactive022/run b/tests/idris2/interactive/interactive022/run similarity index 60% rename from tests/idris2/interactive022/run rename to tests/idris2/interactive/interactive022/run index d928f12b7e..a2b76f26f9 100644 --- a/tests/idris2/interactive022/run +++ b/tests/idris2/interactive/interactive022/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 TypeAtBangSyntax.idr < input diff --git a/tests/idris2/interactive023/TypeAtLambda.idr b/tests/idris2/interactive/interactive023/TypeAtLambda.idr similarity index 100% rename from tests/idris2/interactive023/TypeAtLambda.idr rename to tests/idris2/interactive/interactive023/TypeAtLambda.idr diff --git a/tests/idris2/interactive023/expected b/tests/idris2/interactive/interactive023/expected similarity index 100% rename from tests/idris2/interactive023/expected rename to tests/idris2/interactive/interactive023/expected diff --git a/tests/idris2/interactive023/input b/tests/idris2/interactive/interactive023/input similarity index 100% rename from tests/idris2/interactive023/input rename to tests/idris2/interactive/interactive023/input diff --git a/tests/idris2/interactive023/run b/tests/idris2/interactive/interactive023/run similarity index 57% rename from tests/idris2/interactive023/run rename to tests/idris2/interactive/interactive023/run index ea627ff5fd..ec5c880434 100644 --- a/tests/idris2/interactive023/run +++ b/tests/idris2/interactive/interactive023/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 TypeAtLambda.idr < input diff --git a/tests/idris2/interactive024/TypeAtAsPatterns.idr b/tests/idris2/interactive/interactive024/TypeAtAsPatterns.idr similarity index 100% rename from tests/idris2/interactive024/TypeAtAsPatterns.idr rename to tests/idris2/interactive/interactive024/TypeAtAsPatterns.idr diff --git a/tests/idris2/interactive024/expected b/tests/idris2/interactive/interactive024/expected similarity index 100% rename from tests/idris2/interactive024/expected rename to tests/idris2/interactive/interactive024/expected diff --git a/tests/idris2/interactive024/input b/tests/idris2/interactive/interactive024/input similarity index 100% rename from tests/idris2/interactive024/input rename to tests/idris2/interactive/interactive024/input diff --git a/tests/idris2/interactive024/run b/tests/idris2/interactive/interactive024/run similarity index 60% rename from tests/idris2/interactive024/run rename to tests/idris2/interactive/interactive024/run index 171f2a86e0..230948ec4c 100644 --- a/tests/idris2/interactive024/run +++ b/tests/idris2/interactive/interactive024/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 TypeAtAsPatterns.idr < input diff --git a/tests/idris2/interactive025/TypeAtInterfaces.idr b/tests/idris2/interactive/interactive025/TypeAtInterfaces.idr similarity index 100% rename from tests/idris2/interactive025/TypeAtInterfaces.idr rename to tests/idris2/interactive/interactive025/TypeAtInterfaces.idr diff --git a/tests/idris2/interactive025/expected b/tests/idris2/interactive/interactive025/expected similarity index 100% rename from tests/idris2/interactive025/expected rename to tests/idris2/interactive/interactive025/expected diff --git a/tests/idris2/interactive025/input b/tests/idris2/interactive/interactive025/input similarity index 100% rename from tests/idris2/interactive025/input rename to tests/idris2/interactive/interactive025/input diff --git a/tests/idris2/interactive025/run b/tests/idris2/interactive/interactive025/run similarity index 60% rename from tests/idris2/interactive025/run rename to tests/idris2/interactive/interactive025/run index c0bae1656e..d766c16858 100644 --- a/tests/idris2/interactive025/run +++ b/tests/idris2/interactive/interactive025/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 TypeAtInterfaces.idr < input diff --git a/tests/idris2/interactive026/TypeAtRecords.idr b/tests/idris2/interactive/interactive026/TypeAtRecords.idr similarity index 100% rename from tests/idris2/interactive026/TypeAtRecords.idr rename to tests/idris2/interactive/interactive026/TypeAtRecords.idr diff --git a/tests/idris2/interactive026/expected b/tests/idris2/interactive/interactive026/expected similarity index 100% rename from tests/idris2/interactive026/expected rename to tests/idris2/interactive/interactive026/expected diff --git a/tests/idris2/interactive026/input b/tests/idris2/interactive/interactive026/input similarity index 100% rename from tests/idris2/interactive026/input rename to tests/idris2/interactive/interactive026/input diff --git a/tests/idris2/interactive026/run b/tests/idris2/interactive/interactive026/run similarity index 58% rename from tests/idris2/interactive026/run rename to tests/idris2/interactive/interactive026/run index 2f390b9fa7..34ceee8d4d 100644 --- a/tests/idris2/interactive026/run +++ b/tests/idris2/interactive/interactive026/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 TypeAtRecords.idr < input diff --git a/tests/idris2/interactive027/TypeAtLocalVars.idr b/tests/idris2/interactive/interactive027/TypeAtLocalVars.idr similarity index 100% rename from tests/idris2/interactive027/TypeAtLocalVars.idr rename to tests/idris2/interactive/interactive027/TypeAtLocalVars.idr diff --git a/tests/idris2/interactive027/expected b/tests/idris2/interactive/interactive027/expected similarity index 100% rename from tests/idris2/interactive027/expected rename to tests/idris2/interactive/interactive027/expected diff --git a/tests/idris2/interactive027/input b/tests/idris2/interactive/interactive027/input similarity index 100% rename from tests/idris2/interactive027/input rename to tests/idris2/interactive/interactive027/input diff --git a/tests/idris2/interactive027/run b/tests/idris2/interactive/interactive027/run similarity index 60% rename from tests/idris2/interactive027/run rename to tests/idris2/interactive/interactive027/run index 0d48d91c3d..2727316df4 100644 --- a/tests/idris2/interactive027/run +++ b/tests/idris2/interactive/interactive027/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 TypeAtLocalVars.idr < input diff --git a/tests/idris2/interactive028/expected b/tests/idris2/interactive/interactive028/expected similarity index 100% rename from tests/idris2/interactive028/expected rename to tests/idris2/interactive/interactive028/expected diff --git a/tests/idris2/interactive028/input b/tests/idris2/interactive/interactive028/input similarity index 100% rename from tests/idris2/interactive028/input rename to tests/idris2/interactive/interactive028/input diff --git a/tests/idris2/interactive/interactive028/run b/tests/idris2/interactive/interactive028/run new file mode 100644 index 0000000000..404d7d47c6 --- /dev/null +++ b/tests/idris2/interactive/interactive028/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +idris2 < input diff --git a/tests/idris2/interactive029/Issue834.idr b/tests/idris2/interactive/interactive029/Issue834.idr similarity index 100% rename from tests/idris2/interactive029/Issue834.idr rename to tests/idris2/interactive/interactive029/Issue834.idr diff --git a/tests/idris2/interactive029/expected b/tests/idris2/interactive/interactive029/expected similarity index 100% rename from tests/idris2/interactive029/expected rename to tests/idris2/interactive/interactive029/expected diff --git a/tests/idris2/interactive029/input b/tests/idris2/interactive/interactive029/input similarity index 100% rename from tests/idris2/interactive029/input rename to tests/idris2/interactive/interactive029/input diff --git a/tests/idris2/interactive029/run b/tests/idris2/interactive/interactive029/run similarity index 54% rename from tests/idris2/interactive029/run rename to tests/idris2/interactive/interactive029/run index a56b732f5f..1fc27ff866 100644 --- a/tests/idris2/interactive029/run +++ b/tests/idris2/interactive/interactive029/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 Issue834.idr < input diff --git a/tests/idris2/interactive030/expected b/tests/idris2/interactive/interactive030/expected similarity index 100% rename from tests/idris2/interactive030/expected rename to tests/idris2/interactive/interactive030/expected diff --git a/tests/idris2/interactive030/input b/tests/idris2/interactive/interactive030/input similarity index 100% rename from tests/idris2/interactive030/input rename to tests/idris2/interactive/interactive030/input diff --git a/tests/idris2/interactive/interactive030/run b/tests/idris2/interactive/interactive030/run new file mode 100644 index 0000000000..404d7d47c6 --- /dev/null +++ b/tests/idris2/interactive/interactive030/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +idris2 < input diff --git a/tests/idris2/interactive031/Signatures.idr b/tests/idris2/interactive/interactive031/Signatures.idr similarity index 100% rename from tests/idris2/interactive031/Signatures.idr rename to tests/idris2/interactive/interactive031/Signatures.idr diff --git a/tests/idris2/interactive031/expected b/tests/idris2/interactive/interactive031/expected similarity index 100% rename from tests/idris2/interactive031/expected rename to tests/idris2/interactive/interactive031/expected diff --git a/tests/idris2/interactive031/input b/tests/idris2/interactive/interactive031/input similarity index 100% rename from tests/idris2/interactive031/input rename to tests/idris2/interactive/interactive031/input diff --git a/tests/idris2/interactive031/run b/tests/idris2/interactive/interactive031/run similarity index 56% rename from tests/idris2/interactive031/run rename to tests/idris2/interactive/interactive031/run index 9242884368..89dd98f2bc 100644 --- a/tests/idris2/interactive031/run +++ b/tests/idris2/interactive/interactive031/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 Signatures.idr < input diff --git a/tests/idris2/interactive032/Uninh.idr b/tests/idris2/interactive/interactive032/Uninh.idr similarity index 100% rename from tests/idris2/interactive032/Uninh.idr rename to tests/idris2/interactive/interactive032/Uninh.idr diff --git a/tests/idris2/interactive032/expected b/tests/idris2/interactive/interactive032/expected similarity index 100% rename from tests/idris2/interactive032/expected rename to tests/idris2/interactive/interactive032/expected diff --git a/tests/idris2/interactive032/input b/tests/idris2/interactive/interactive032/input similarity index 100% rename from tests/idris2/interactive032/input rename to tests/idris2/interactive/interactive032/input diff --git a/tests/idris2/interactive032/run b/tests/idris2/interactive/interactive032/run similarity index 52% rename from tests/idris2/interactive032/run rename to tests/idris2/interactive/interactive032/run index 5a5252bddb..64d92c62a1 100755 --- a/tests/idris2/interactive032/run +++ b/tests/idris2/interactive/interactive032/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 Uninh.idr < input diff --git a/tests/idris2/interactive033/UninhIndent.idr b/tests/idris2/interactive/interactive033/UninhIndent.idr similarity index 100% rename from tests/idris2/interactive033/UninhIndent.idr rename to tests/idris2/interactive/interactive033/UninhIndent.idr diff --git a/tests/idris2/interactive033/expected b/tests/idris2/interactive/interactive033/expected similarity index 100% rename from tests/idris2/interactive033/expected rename to tests/idris2/interactive/interactive033/expected diff --git a/tests/idris2/interactive033/input b/tests/idris2/interactive/interactive033/input similarity index 100% rename from tests/idris2/interactive033/input rename to tests/idris2/interactive/interactive033/input diff --git a/tests/idris2/interactive033/run b/tests/idris2/interactive/interactive033/run similarity index 57% rename from tests/idris2/interactive033/run rename to tests/idris2/interactive/interactive033/run index 1dfbad70a1..5d8bfd389c 100755 --- a/tests/idris2/interactive033/run +++ b/tests/idris2/interactive/interactive033/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 UninhIndent.idr < input diff --git a/tests/idris2/interactive034/expected b/tests/idris2/interactive/interactive034/expected similarity index 100% rename from tests/idris2/interactive034/expected rename to tests/idris2/interactive/interactive034/expected diff --git a/tests/idris2/interactive034/input b/tests/idris2/interactive/interactive034/input similarity index 100% rename from tests/idris2/interactive034/input rename to tests/idris2/interactive/interactive034/input diff --git a/tests/idris2/interactive034/run b/tests/idris2/interactive/interactive034/run similarity index 53% rename from tests/idris2/interactive034/run rename to tests/idris2/interactive/interactive034/run index 7758f3e19e..a0f52b2358 100755 --- a/tests/idris2/interactive034/run +++ b/tests/idris2/interactive/interactive034/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 timeout.idr < input diff --git a/tests/idris2/interactive034/timeout.idr b/tests/idris2/interactive/interactive034/timeout.idr similarity index 100% rename from tests/idris2/interactive034/timeout.idr rename to tests/idris2/interactive/interactive034/timeout.idr diff --git a/tests/idris2/interactive035/expected b/tests/idris2/interactive/interactive035/expected similarity index 100% rename from tests/idris2/interactive035/expected rename to tests/idris2/interactive/interactive035/expected diff --git a/tests/idris2/interactive035/input b/tests/idris2/interactive/interactive035/input similarity index 100% rename from tests/idris2/interactive035/input rename to tests/idris2/interactive/interactive035/input diff --git a/tests/idris2/interactive035/run b/tests/idris2/interactive/interactive035/run similarity index 52% rename from tests/idris2/interactive035/run rename to tests/idris2/interactive/interactive035/run index 654715325d..fd0b67bcd1 100755 --- a/tests/idris2/interactive035/run +++ b/tests/idris2/interactive/interactive035/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 unify.idr < input diff --git a/tests/idris2/interactive035/unify.idr b/tests/idris2/interactive/interactive035/unify.idr similarity index 100% rename from tests/idris2/interactive035/unify.idr rename to tests/idris2/interactive/interactive035/unify.idr diff --git a/tests/idris2/interactive036/casefn.idr b/tests/idris2/interactive/interactive036/casefn.idr similarity index 100% rename from tests/idris2/interactive036/casefn.idr rename to tests/idris2/interactive/interactive036/casefn.idr diff --git a/tests/idris2/interactive036/expected b/tests/idris2/interactive/interactive036/expected similarity index 100% rename from tests/idris2/interactive036/expected rename to tests/idris2/interactive/interactive036/expected diff --git a/tests/idris2/interactive036/input b/tests/idris2/interactive/interactive036/input similarity index 100% rename from tests/idris2/interactive036/input rename to tests/idris2/interactive/interactive036/input diff --git a/tests/idris2/interactive036/run b/tests/idris2/interactive/interactive036/run similarity index 52% rename from tests/idris2/interactive036/run rename to tests/idris2/interactive/interactive036/run index 29c5f319c3..a2e3ddb932 100755 --- a/tests/idris2/interactive036/run +++ b/tests/idris2/interactive/interactive036/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 casefn.idr < input diff --git a/tests/idris2/interactive037/Holes.idr b/tests/idris2/interactive/interactive037/Holes.idr similarity index 100% rename from tests/idris2/interactive037/Holes.idr rename to tests/idris2/interactive/interactive037/Holes.idr diff --git a/tests/idris2/interactive037/expected b/tests/idris2/interactive/interactive037/expected similarity index 100% rename from tests/idris2/interactive037/expected rename to tests/idris2/interactive/interactive037/expected diff --git a/tests/idris2/interactive037/input b/tests/idris2/interactive/interactive037/input similarity index 100% rename from tests/idris2/interactive037/input rename to tests/idris2/interactive/interactive037/input diff --git a/tests/idris2/interactive037/run b/tests/idris2/interactive/interactive037/run similarity index 52% rename from tests/idris2/interactive037/run rename to tests/idris2/interactive/interactive037/run index e33955b1a0..33d7384221 100755 --- a/tests/idris2/interactive037/run +++ b/tests/idris2/interactive/interactive037/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 Holes.idr < input diff --git a/tests/idris2/interactive038/IEdit.idr b/tests/idris2/interactive/interactive038/IEdit.idr similarity index 100% rename from tests/idris2/interactive038/IEdit.idr rename to tests/idris2/interactive/interactive038/IEdit.idr diff --git a/tests/idris2/interactive038/expected b/tests/idris2/interactive/interactive038/expected similarity index 100% rename from tests/idris2/interactive038/expected rename to tests/idris2/interactive/interactive038/expected diff --git a/tests/idris2/interactive038/input b/tests/idris2/interactive/interactive038/input similarity index 100% rename from tests/idris2/interactive038/input rename to tests/idris2/interactive/interactive038/input diff --git a/tests/idris2/interactive/interactive038/run b/tests/idris2/interactive/interactive038/run new file mode 100755 index 0000000000..737446ffd5 --- /dev/null +++ b/tests/idris2/interactive/interactive038/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +idris2 IEdit.idr < input diff --git a/tests/idris2/interactive039/CS_Syntax.idr b/tests/idris2/interactive/interactive039/CS_Syntax.idr similarity index 100% rename from tests/idris2/interactive039/CS_Syntax.idr rename to tests/idris2/interactive/interactive039/CS_Syntax.idr diff --git a/tests/idris2/interactive039/expected b/tests/idris2/interactive/interactive039/expected similarity index 100% rename from tests/idris2/interactive039/expected rename to tests/idris2/interactive/interactive039/expected diff --git a/tests/idris2/interactive039/input b/tests/idris2/interactive/interactive039/input similarity index 100% rename from tests/idris2/interactive039/input rename to tests/idris2/interactive/interactive039/input diff --git a/tests/idris2/interactive039/run b/tests/idris2/interactive/interactive039/run similarity index 55% rename from tests/idris2/interactive039/run rename to tests/idris2/interactive/interactive039/run index f0660f1109..057d2aa6f0 100755 --- a/tests/idris2/interactive039/run +++ b/tests/idris2/interactive/interactive039/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 CS_Syntax.idr < input diff --git a/tests/idris2/interactive040/expected b/tests/idris2/interactive/interactive040/expected similarity index 100% rename from tests/idris2/interactive040/expected rename to tests/idris2/interactive/interactive040/expected diff --git a/tests/idris2/interactive040/input b/tests/idris2/interactive/interactive040/input similarity index 100% rename from tests/idris2/interactive040/input rename to tests/idris2/interactive/interactive040/input diff --git a/tests/idris2/interactive/interactive040/run b/tests/idris2/interactive/interactive040/run new file mode 100755 index 0000000000..404d7d47c6 --- /dev/null +++ b/tests/idris2/interactive/interactive040/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +idris2 < input diff --git a/tests/idris2/interactive041/Issue1741.idr b/tests/idris2/interactive/interactive041/Issue1741.idr similarity index 100% rename from tests/idris2/interactive041/Issue1741.idr rename to tests/idris2/interactive/interactive041/Issue1741.idr diff --git a/tests/idris2/interactive041/expected b/tests/idris2/interactive/interactive041/expected similarity index 100% rename from tests/idris2/interactive041/expected rename to tests/idris2/interactive/interactive041/expected diff --git a/tests/idris2/interactive041/input b/tests/idris2/interactive/interactive041/input similarity index 100% rename from tests/idris2/interactive041/input rename to tests/idris2/interactive/interactive041/input diff --git a/tests/idris2/interactive041/run b/tests/idris2/interactive/interactive041/run similarity index 55% rename from tests/idris2/interactive041/run rename to tests/idris2/interactive/interactive041/run index db6d86a927..7abf02fd93 100755 --- a/tests/idris2/interactive041/run +++ b/tests/idris2/interactive/interactive041/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 Issue1741.idr < input diff --git a/tests/idris2/interactive042/Issue35-2.idr b/tests/idris2/interactive/interactive042/Issue35-2.idr similarity index 100% rename from tests/idris2/interactive042/Issue35-2.idr rename to tests/idris2/interactive/interactive042/Issue35-2.idr diff --git a/tests/idris2/interactive042/Issue35.idr b/tests/idris2/interactive/interactive042/Issue35.idr similarity index 100% rename from tests/idris2/interactive042/Issue35.idr rename to tests/idris2/interactive/interactive042/Issue35.idr diff --git a/tests/idris2/interactive042/expected b/tests/idris2/interactive/interactive042/expected similarity index 100% rename from tests/idris2/interactive042/expected rename to tests/idris2/interactive/interactive042/expected diff --git a/tests/idris2/interactive042/input b/tests/idris2/interactive/interactive042/input similarity index 100% rename from tests/idris2/interactive042/input rename to tests/idris2/interactive/interactive042/input diff --git a/tests/idris2/interactive042/run b/tests/idris2/interactive/interactive042/run similarity index 87% rename from tests/idris2/interactive042/run rename to tests/idris2/interactive/interactive042/run index 323393e8f4..8b0b293c23 100755 --- a/tests/idris2/interactive042/run +++ b/tests/idris2/interactive/interactive042/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 Issue35.idr < input check Issue35-2.idr || true diff --git a/tests/idris2/interactive043/ImplicitSplits.idr b/tests/idris2/interactive/interactive043/ImplicitSplits.idr similarity index 100% rename from tests/idris2/interactive043/ImplicitSplits.idr rename to tests/idris2/interactive/interactive043/ImplicitSplits.idr diff --git a/tests/idris2/interactive043/expected b/tests/idris2/interactive/interactive043/expected similarity index 100% rename from tests/idris2/interactive043/expected rename to tests/idris2/interactive/interactive043/expected diff --git a/tests/idris2/interactive043/input b/tests/idris2/interactive/interactive043/input similarity index 100% rename from tests/idris2/interactive043/input rename to tests/idris2/interactive/interactive043/input diff --git a/tests/idris2/interactive043/run b/tests/idris2/interactive/interactive043/run similarity index 90% rename from tests/idris2/interactive043/run rename to tests/idris2/interactive/interactive043/run index d2631795ff..fd21ed380a 100755 --- a/tests/idris2/interactive043/run +++ b/tests/idris2/interactive/interactive043/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh rm -f ImplicitSplitsGen.idr cp ImplicitSplits.idr ImplicitSplitsGen.idr diff --git a/tests/idris2/interactive044/SplitShadow.idr b/tests/idris2/interactive/interactive044/SplitShadow.idr similarity index 100% rename from tests/idris2/interactive044/SplitShadow.idr rename to tests/idris2/interactive/interactive044/SplitShadow.idr diff --git a/tests/idris2/interactive044/expected b/tests/idris2/interactive/interactive044/expected similarity index 100% rename from tests/idris2/interactive044/expected rename to tests/idris2/interactive/interactive044/expected diff --git a/tests/idris2/interactive044/input b/tests/idris2/interactive/interactive044/input similarity index 100% rename from tests/idris2/interactive044/input rename to tests/idris2/interactive/interactive044/input diff --git a/tests/idris2/interactive044/run b/tests/idris2/interactive/interactive044/run similarity index 84% rename from tests/idris2/interactive044/run rename to tests/idris2/interactive/interactive044/run index 71d5a34bd4..fbd0ab5bc4 100755 --- a/tests/idris2/interactive044/run +++ b/tests/idris2/interactive/interactive044/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh rm -f SplitShadowGen.idr cp SplitShadow.idr SplitShadowGen.idr diff --git a/tests/idris2/interactive045/Issue1742.idr b/tests/idris2/interactive/interactive045/Issue1742.idr similarity index 100% rename from tests/idris2/interactive045/Issue1742.idr rename to tests/idris2/interactive/interactive045/Issue1742.idr diff --git a/tests/idris2/interactive045/expected b/tests/idris2/interactive/interactive045/expected similarity index 100% rename from tests/idris2/interactive045/expected rename to tests/idris2/interactive/interactive045/expected diff --git a/tests/idris2/interactive045/input b/tests/idris2/interactive/interactive045/input similarity index 100% rename from tests/idris2/interactive045/input rename to tests/idris2/interactive/interactive045/input diff --git a/tests/idris2/interactive045/run b/tests/idris2/interactive/interactive045/run similarity index 83% rename from tests/idris2/interactive045/run rename to tests/idris2/interactive/interactive045/run index 9041703d8a..3ba50ad94e 100755 --- a/tests/idris2/interactive045/run +++ b/tests/idris2/interactive/interactive045/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh rm -f Issue1742Gen.idr cp Issue1742.idr Issue1742Gen.idr diff --git a/tests/idris2/interactive046/Issue2712.idr b/tests/idris2/interactive/interactive046/Issue2712.idr similarity index 100% rename from tests/idris2/interactive046/Issue2712.idr rename to tests/idris2/interactive/interactive046/Issue2712.idr diff --git a/tests/idris2/interactive046/expected b/tests/idris2/interactive/interactive046/expected similarity index 100% rename from tests/idris2/interactive046/expected rename to tests/idris2/interactive/interactive046/expected diff --git a/tests/idris2/interactive/interactive046/run b/tests/idris2/interactive/interactive046/run new file mode 100755 index 0000000000..107bda4e31 --- /dev/null +++ b/tests/idris2/interactive/interactive046/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Issue2712.idr diff --git a/tests/idris2/interactive007/run b/tests/idris2/interactive007/run deleted file mode 100755 index ec050c77b1..0000000000 --- a/tests/idris2/interactive007/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -idris2 IEdit.idr < input diff --git a/tests/idris2/interactive008/run b/tests/idris2/interactive008/run deleted file mode 100755 index ec050c77b1..0000000000 --- a/tests/idris2/interactive008/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -idris2 IEdit.idr < input diff --git a/tests/idris2/interactive010/run b/tests/idris2/interactive010/run deleted file mode 100755 index ec050c77b1..0000000000 --- a/tests/idris2/interactive010/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -idris2 IEdit.idr < input diff --git a/tests/idris2/interactive011/run b/tests/idris2/interactive011/run deleted file mode 100755 index ec050c77b1..0000000000 --- a/tests/idris2/interactive011/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -idris2 IEdit.idr < input diff --git a/tests/idris2/interactive015/run b/tests/idris2/interactive015/run deleted file mode 100755 index ec050c77b1..0000000000 --- a/tests/idris2/interactive015/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -idris2 IEdit.idr < input diff --git a/tests/idris2/interactive028/run b/tests/idris2/interactive028/run deleted file mode 100644 index fecfb4b107..0000000000 --- a/tests/idris2/interactive028/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -idris2 < input diff --git a/tests/idris2/interactive030/run b/tests/idris2/interactive030/run deleted file mode 100644 index fecfb4b107..0000000000 --- a/tests/idris2/interactive030/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -idris2 < input diff --git a/tests/idris2/interactive038/run b/tests/idris2/interactive038/run deleted file mode 100755 index ec050c77b1..0000000000 --- a/tests/idris2/interactive038/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -idris2 IEdit.idr < input diff --git a/tests/idris2/interactive040/run b/tests/idris2/interactive040/run deleted file mode 100755 index fecfb4b107..0000000000 --- a/tests/idris2/interactive040/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -idris2 < input diff --git a/tests/idris2/interactive046/run b/tests/idris2/interactive046/run deleted file mode 100755 index dd6bd50c7f..0000000000 --- a/tests/idris2/interactive046/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check Issue2712.idr diff --git a/tests/idris2/interface001/IFace.idr b/tests/idris2/interface/interface001/IFace.idr similarity index 100% rename from tests/idris2/interface001/IFace.idr rename to tests/idris2/interface/interface001/IFace.idr diff --git a/tests/idris2/interface001/IFace1.idr b/tests/idris2/interface/interface001/IFace1.idr similarity index 100% rename from tests/idris2/interface001/IFace1.idr rename to tests/idris2/interface/interface001/IFace1.idr diff --git a/tests/idris2/interface001/Stuff.idr b/tests/idris2/interface/interface001/Stuff.idr similarity index 100% rename from tests/idris2/interface001/Stuff.idr rename to tests/idris2/interface/interface001/Stuff.idr diff --git a/tests/idris2/interface001/expected b/tests/idris2/interface/interface001/expected similarity index 100% rename from tests/idris2/interface001/expected rename to tests/idris2/interface/interface001/expected diff --git a/tests/idris2/interface001/input b/tests/idris2/interface/interface001/input similarity index 100% rename from tests/idris2/interface001/input rename to tests/idris2/interface/interface001/input diff --git a/tests/idris2/interface001/input1 b/tests/idris2/interface/interface001/input1 similarity index 100% rename from tests/idris2/interface001/input1 rename to tests/idris2/interface/interface001/input1 diff --git a/tests/idris2/interface001/run b/tests/idris2/interface/interface001/run similarity index 76% rename from tests/idris2/interface001/run rename to tests/idris2/interface/interface001/run index cca7bbd7b7..8ecbe60a7f 100755 --- a/tests/idris2/interface001/run +++ b/tests/idris2/interface/interface001/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 --no-prelude IFace.idr < input idris2 --no-prelude IFace1.idr < input1 diff --git a/tests/idris2/interface002/Functor.idr b/tests/idris2/interface/interface002/Functor.idr similarity index 100% rename from tests/idris2/interface002/Functor.idr rename to tests/idris2/interface/interface002/Functor.idr diff --git a/tests/idris2/interface002/Stuff.idr b/tests/idris2/interface/interface002/Stuff.idr similarity index 100% rename from tests/idris2/interface002/Stuff.idr rename to tests/idris2/interface/interface002/Stuff.idr diff --git a/tests/idris2/interface002/expected b/tests/idris2/interface/interface002/expected similarity index 100% rename from tests/idris2/interface002/expected rename to tests/idris2/interface/interface002/expected diff --git a/tests/idris2/interface002/input b/tests/idris2/interface/interface002/input similarity index 100% rename from tests/idris2/interface002/input rename to tests/idris2/interface/interface002/input diff --git a/tests/idris2/interface002/run b/tests/idris2/interface/interface002/run similarity index 63% rename from tests/idris2/interface002/run rename to tests/idris2/interface/interface002/run index 19265959fa..031503983f 100755 --- a/tests/idris2/interface002/run +++ b/tests/idris2/interface/interface002/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 --no-prelude Functor.idr < input diff --git a/tests/idris2/interface003/Do.idr b/tests/idris2/interface/interface003/Do.idr similarity index 100% rename from tests/idris2/interface003/Do.idr rename to tests/idris2/interface/interface003/Do.idr diff --git a/tests/idris2/interface003/expected b/tests/idris2/interface/interface003/expected similarity index 100% rename from tests/idris2/interface003/expected rename to tests/idris2/interface/interface003/expected diff --git a/tests/idris2/interface003/input b/tests/idris2/interface/interface003/input similarity index 100% rename from tests/idris2/interface003/input rename to tests/idris2/interface/interface003/input diff --git a/tests/idris2/interface/interface003/run b/tests/idris2/interface/interface003/run new file mode 100755 index 0000000000..190b4da72f --- /dev/null +++ b/tests/idris2/interface/interface003/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +idris2 Do.idr < input diff --git a/tests/idris2/interface004/Do.idr b/tests/idris2/interface/interface004/Do.idr similarity index 100% rename from tests/idris2/interface004/Do.idr rename to tests/idris2/interface/interface004/Do.idr diff --git a/tests/idris2/interface004/expected b/tests/idris2/interface/interface004/expected similarity index 100% rename from tests/idris2/interface004/expected rename to tests/idris2/interface/interface004/expected diff --git a/tests/idris2/interface004/input b/tests/idris2/interface/interface004/input similarity index 100% rename from tests/idris2/interface004/input rename to tests/idris2/interface/interface004/input diff --git a/tests/idris2/interface/interface004/run b/tests/idris2/interface/interface004/run new file mode 100755 index 0000000000..190b4da72f --- /dev/null +++ b/tests/idris2/interface/interface004/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +idris2 Do.idr < input diff --git a/tests/idris2/interface005/Deps.idr b/tests/idris2/interface/interface005/Deps.idr similarity index 100% rename from tests/idris2/interface005/Deps.idr rename to tests/idris2/interface/interface005/Deps.idr diff --git a/tests/idris2/interface005/expected b/tests/idris2/interface/interface005/expected similarity index 100% rename from tests/idris2/interface005/expected rename to tests/idris2/interface/interface005/expected diff --git a/tests/idris2/interface/interface005/run b/tests/idris2/interface/interface005/run new file mode 100755 index 0000000000..3c4242494a --- /dev/null +++ b/tests/idris2/interface/interface005/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Deps.idr diff --git a/tests/idris2/interface006/Apply.idr b/tests/idris2/interface/interface006/Apply.idr similarity index 100% rename from tests/idris2/interface006/Apply.idr rename to tests/idris2/interface/interface006/Apply.idr diff --git a/tests/idris2/interface006/Biapplicative.idr b/tests/idris2/interface/interface006/Biapplicative.idr similarity index 100% rename from tests/idris2/interface006/Biapplicative.idr rename to tests/idris2/interface/interface006/Biapplicative.idr diff --git a/tests/idris2/interface006/Bimonad.idr b/tests/idris2/interface/interface006/Bimonad.idr similarity index 100% rename from tests/idris2/interface006/Bimonad.idr rename to tests/idris2/interface/interface006/Bimonad.idr diff --git a/tests/idris2/interface006/expected b/tests/idris2/interface/interface006/expected similarity index 100% rename from tests/idris2/interface006/expected rename to tests/idris2/interface/interface006/expected diff --git a/tests/idris2/interface/interface006/run b/tests/idris2/interface/interface006/run new file mode 100755 index 0000000000..31d6d97908 --- /dev/null +++ b/tests/idris2/interface/interface006/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Bimonad.idr diff --git a/tests/idris2/interface007/A.idr b/tests/idris2/interface/interface007/A.idr similarity index 100% rename from tests/idris2/interface007/A.idr rename to tests/idris2/interface/interface007/A.idr diff --git a/tests/idris2/interface007/expected b/tests/idris2/interface/interface007/expected similarity index 100% rename from tests/idris2/interface007/expected rename to tests/idris2/interface/interface007/expected diff --git a/tests/idris2/interface007/run b/tests/idris2/interface/interface007/run similarity index 57% rename from tests/idris2/interface007/run rename to tests/idris2/interface/interface007/run index 24086a6c87..b85de6bf00 100755 --- a/tests/idris2/interface007/run +++ b/tests/idris2/interface/interface007/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh check --debug-elab-check A.idr diff --git a/tests/idris2/interface008/Deps.idr b/tests/idris2/interface/interface008/Deps.idr similarity index 100% rename from tests/idris2/interface008/Deps.idr rename to tests/idris2/interface/interface008/Deps.idr diff --git a/tests/idris2/interface008/expected b/tests/idris2/interface/interface008/expected similarity index 100% rename from tests/idris2/interface008/expected rename to tests/idris2/interface/interface008/expected diff --git a/tests/idris2/interface/interface008/run b/tests/idris2/interface/interface008/run new file mode 100755 index 0000000000..3c4242494a --- /dev/null +++ b/tests/idris2/interface/interface008/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Deps.idr diff --git a/tests/idris2/interface009/Odd.idr b/tests/idris2/interface/interface009/Odd.idr similarity index 100% rename from tests/idris2/interface009/Odd.idr rename to tests/idris2/interface/interface009/Odd.idr diff --git a/tests/idris2/interface009/expected b/tests/idris2/interface/interface009/expected similarity index 100% rename from tests/idris2/interface009/expected rename to tests/idris2/interface/interface009/expected diff --git a/tests/idris2/interface009/input b/tests/idris2/interface/interface009/input similarity index 100% rename from tests/idris2/interface009/input rename to tests/idris2/interface/interface009/input diff --git a/tests/idris2/interface009/run b/tests/idris2/interface/interface009/run similarity index 50% rename from tests/idris2/interface009/run rename to tests/idris2/interface/interface009/run index 1118c65bcc..1869a78cde 100755 --- a/tests/idris2/interface009/run +++ b/tests/idris2/interface/interface009/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 Odd.idr < input diff --git a/tests/idris2/interface010/Dep.idr b/tests/idris2/interface/interface010/Dep.idr similarity index 100% rename from tests/idris2/interface010/Dep.idr rename to tests/idris2/interface/interface010/Dep.idr diff --git a/tests/idris2/interface010/expected b/tests/idris2/interface/interface010/expected similarity index 100% rename from tests/idris2/interface010/expected rename to tests/idris2/interface/interface010/expected diff --git a/tests/idris2/interface/interface010/run b/tests/idris2/interface/interface010/run new file mode 100755 index 0000000000..e1df128915 --- /dev/null +++ b/tests/idris2/interface/interface010/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Dep.idr diff --git a/tests/idris2/interface011/FuncImpl.idr b/tests/idris2/interface/interface011/FuncImpl.idr similarity index 100% rename from tests/idris2/interface011/FuncImpl.idr rename to tests/idris2/interface/interface011/FuncImpl.idr diff --git a/tests/idris2/interface011/expected b/tests/idris2/interface/interface011/expected similarity index 100% rename from tests/idris2/interface011/expected rename to tests/idris2/interface/interface011/expected diff --git a/tests/idris2/interface/interface011/run b/tests/idris2/interface/interface011/run new file mode 100755 index 0000000000..0b544f166f --- /dev/null +++ b/tests/idris2/interface/interface011/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check FuncImpl.idr diff --git a/tests/idris2/interface012/Defmeth.idr b/tests/idris2/interface/interface012/Defmeth.idr similarity index 100% rename from tests/idris2/interface012/Defmeth.idr rename to tests/idris2/interface/interface012/Defmeth.idr diff --git a/tests/idris2/interface012/expected b/tests/idris2/interface/interface012/expected similarity index 100% rename from tests/idris2/interface012/expected rename to tests/idris2/interface/interface012/expected diff --git a/tests/idris2/interface/interface012/run b/tests/idris2/interface/interface012/run new file mode 100755 index 0000000000..b74ac6cffa --- /dev/null +++ b/tests/idris2/interface/interface012/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Defmeth.idr diff --git a/tests/idris2/interface013/TypeInt.idr b/tests/idris2/interface/interface013/TypeInt.idr similarity index 100% rename from tests/idris2/interface013/TypeInt.idr rename to tests/idris2/interface/interface013/TypeInt.idr diff --git a/tests/idris2/interface013/expected b/tests/idris2/interface/interface013/expected similarity index 100% rename from tests/idris2/interface013/expected rename to tests/idris2/interface/interface013/expected diff --git a/tests/idris2/interface/interface013/run b/tests/idris2/interface/interface013/run new file mode 100755 index 0000000000..4913feed1f --- /dev/null +++ b/tests/idris2/interface/interface013/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check TypeInt.idr diff --git a/tests/idris2/interface014/DepInt.idr b/tests/idris2/interface/interface014/DepInt.idr similarity index 100% rename from tests/idris2/interface014/DepInt.idr rename to tests/idris2/interface/interface014/DepInt.idr diff --git a/tests/idris2/interface014/expected b/tests/idris2/interface/interface014/expected similarity index 100% rename from tests/idris2/interface014/expected rename to tests/idris2/interface/interface014/expected diff --git a/tests/idris2/interface/interface014/run b/tests/idris2/interface/interface014/run new file mode 100755 index 0000000000..4a3e1484d6 --- /dev/null +++ b/tests/idris2/interface/interface014/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check DepInt.idr diff --git a/tests/idris2/interface015/expected b/tests/idris2/interface/interface015/expected similarity index 100% rename from tests/idris2/interface015/expected rename to tests/idris2/interface/interface015/expected diff --git a/tests/idris2/interface015/gnu.idr b/tests/idris2/interface/interface015/gnu.idr similarity index 100% rename from tests/idris2/interface015/gnu.idr rename to tests/idris2/interface/interface015/gnu.idr diff --git a/tests/idris2/interface/interface015/run b/tests/idris2/interface/interface015/run new file mode 100755 index 0000000000..0b3d6002bb --- /dev/null +++ b/tests/idris2/interface/interface015/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check gnu.idr diff --git a/tests/idris2/interface016/TwoNum.idr b/tests/idris2/interface/interface016/TwoNum.idr similarity index 100% rename from tests/idris2/interface016/TwoNum.idr rename to tests/idris2/interface/interface016/TwoNum.idr diff --git a/tests/idris2/interface016/expected b/tests/idris2/interface/interface016/expected similarity index 100% rename from tests/idris2/interface016/expected rename to tests/idris2/interface/interface016/expected diff --git a/tests/idris2/interface/interface016/run b/tests/idris2/interface/interface016/run new file mode 100755 index 0000000000..f0dca8250c --- /dev/null +++ b/tests/idris2/interface/interface016/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check TwoNum.idr diff --git a/tests/idris2/interface017/Tricho.idr b/tests/idris2/interface/interface017/Tricho.idr similarity index 100% rename from tests/idris2/interface017/Tricho.idr rename to tests/idris2/interface/interface017/Tricho.idr diff --git a/tests/idris2/interface017/expected b/tests/idris2/interface/interface017/expected similarity index 100% rename from tests/idris2/interface017/expected rename to tests/idris2/interface/interface017/expected diff --git a/tests/idris2/interface/interface017/run b/tests/idris2/interface/interface017/run new file mode 100755 index 0000000000..bf37ec45e8 --- /dev/null +++ b/tests/idris2/interface/interface017/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Tricho.idr diff --git a/tests/idris2/interface018/Sized.idr b/tests/idris2/interface/interface018/Sized.idr similarity index 100% rename from tests/idris2/interface018/Sized.idr rename to tests/idris2/interface/interface018/Sized.idr diff --git a/tests/idris2/interface018/Sized2.idr b/tests/idris2/interface/interface018/Sized2.idr similarity index 100% rename from tests/idris2/interface018/Sized2.idr rename to tests/idris2/interface/interface018/Sized2.idr diff --git a/tests/idris2/interface018/Sized3.idr b/tests/idris2/interface/interface018/Sized3.idr similarity index 100% rename from tests/idris2/interface018/Sized3.idr rename to tests/idris2/interface/interface018/Sized3.idr diff --git a/tests/idris2/interface018/expected b/tests/idris2/interface/interface018/expected similarity index 100% rename from tests/idris2/interface018/expected rename to tests/idris2/interface/interface018/expected diff --git a/tests/idris2/interface018/input b/tests/idris2/interface/interface018/input similarity index 100% rename from tests/idris2/interface018/input rename to tests/idris2/interface/interface018/input diff --git a/tests/idris2/interface018/run b/tests/idris2/interface/interface018/run similarity index 71% rename from tests/idris2/interface018/run rename to tests/idris2/interface/interface018/run index f29f8f6faf..0018393bc9 100755 --- a/tests/idris2/interface018/run +++ b/tests/idris2/interface/interface018/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 Sized.idr < input check Sized2.idr diff --git a/tests/idris2/interface019/LocalHints.idr b/tests/idris2/interface/interface019/LocalHints.idr similarity index 100% rename from tests/idris2/interface019/LocalHints.idr rename to tests/idris2/interface/interface019/LocalHints.idr diff --git a/tests/idris2/interface019/expected b/tests/idris2/interface/interface019/expected similarity index 100% rename from tests/idris2/interface019/expected rename to tests/idris2/interface/interface019/expected diff --git a/tests/idris2/interface/interface019/run b/tests/idris2/interface/interface019/run new file mode 100755 index 0000000000..57e1ed8dc8 --- /dev/null +++ b/tests/idris2/interface/interface019/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check LocalHints.idr diff --git a/tests/idris2/interface020/LocalInterface.idr b/tests/idris2/interface/interface020/LocalInterface.idr similarity index 100% rename from tests/idris2/interface020/LocalInterface.idr rename to tests/idris2/interface/interface020/LocalInterface.idr diff --git a/tests/idris2/interface020/expected b/tests/idris2/interface/interface020/expected similarity index 100% rename from tests/idris2/interface020/expected rename to tests/idris2/interface/interface020/expected diff --git a/tests/idris2/interface020/run b/tests/idris2/interface/interface020/run similarity index 52% rename from tests/idris2/interface020/run rename to tests/idris2/interface/interface020/run index b6cf755f88..2aac792ce7 100755 --- a/tests/idris2/interface020/run +++ b/tests/idris2/interface/interface020/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh check LocalInterface.idr diff --git a/tests/idris2/interface021/LocalHint.idr b/tests/idris2/interface/interface021/LocalHint.idr similarity index 100% rename from tests/idris2/interface021/LocalHint.idr rename to tests/idris2/interface/interface021/LocalHint.idr diff --git a/tests/idris2/interface021/expected b/tests/idris2/interface/interface021/expected similarity index 100% rename from tests/idris2/interface021/expected rename to tests/idris2/interface/interface021/expected diff --git a/tests/idris2/interface/interface021/run b/tests/idris2/interface/interface021/run new file mode 100755 index 0000000000..41f415a1ec --- /dev/null +++ b/tests/idris2/interface/interface021/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check LocalHint.idr diff --git a/tests/idris2/interface022/DotMethod.idr b/tests/idris2/interface/interface022/DotMethod.idr similarity index 100% rename from tests/idris2/interface022/DotMethod.idr rename to tests/idris2/interface/interface022/DotMethod.idr diff --git a/tests/idris2/interface022/expected b/tests/idris2/interface/interface022/expected similarity index 100% rename from tests/idris2/interface022/expected rename to tests/idris2/interface/interface022/expected diff --git a/tests/idris2/interface022/input b/tests/idris2/interface/interface022/input similarity index 100% rename from tests/idris2/interface022/input rename to tests/idris2/interface/interface022/input diff --git a/tests/idris2/interface022/run b/tests/idris2/interface/interface022/run similarity index 55% rename from tests/idris2/interface022/run rename to tests/idris2/interface/interface022/run index 273dd0c211..80a4cd334f 100755 --- a/tests/idris2/interface022/run +++ b/tests/idris2/interface/interface022/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 DotMethod.idr < input diff --git a/tests/idris2/interface023/AppComp.idr b/tests/idris2/interface/interface023/AppComp.idr similarity index 100% rename from tests/idris2/interface023/AppComp.idr rename to tests/idris2/interface/interface023/AppComp.idr diff --git a/tests/idris2/interface023/expected b/tests/idris2/interface/interface023/expected similarity index 100% rename from tests/idris2/interface023/expected rename to tests/idris2/interface/interface023/expected diff --git a/tests/idris2/interface023/input b/tests/idris2/interface/interface023/input similarity index 100% rename from tests/idris2/interface023/input rename to tests/idris2/interface/interface023/input diff --git a/tests/idris2/interface023/run b/tests/idris2/interface/interface023/run similarity index 53% rename from tests/idris2/interface023/run rename to tests/idris2/interface/interface023/run index b8b0938790..f01ef6a0e1 100644 --- a/tests/idris2/interface023/run +++ b/tests/idris2/interface/interface023/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 AppComp.idr < input diff --git a/tests/idris2/interface024/EH.idr b/tests/idris2/interface/interface024/EH.idr similarity index 100% rename from tests/idris2/interface024/EH.idr rename to tests/idris2/interface/interface024/EH.idr diff --git a/tests/idris2/interface024/expected b/tests/idris2/interface/interface024/expected similarity index 100% rename from tests/idris2/interface024/expected rename to tests/idris2/interface/interface024/expected diff --git a/tests/idris2/interface024/run b/tests/idris2/interface/interface024/run similarity index 51% rename from tests/idris2/interface024/run rename to tests/idris2/interface/interface024/run index 4db6d1aee3..e02dda05ac 100644 --- a/tests/idris2/interface024/run +++ b/tests/idris2/interface/interface024/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh check EH.idr -p contrib diff --git a/tests/idris2/interface025/AutoSearchHide1.idr b/tests/idris2/interface/interface025/AutoSearchHide1.idr similarity index 100% rename from tests/idris2/interface025/AutoSearchHide1.idr rename to tests/idris2/interface/interface025/AutoSearchHide1.idr diff --git a/tests/idris2/interface025/AutoSearchHide2.idr b/tests/idris2/interface/interface025/AutoSearchHide2.idr similarity index 100% rename from tests/idris2/interface025/AutoSearchHide2.idr rename to tests/idris2/interface/interface025/AutoSearchHide2.idr diff --git a/tests/idris2/interface025/expected b/tests/idris2/interface/interface025/expected similarity index 100% rename from tests/idris2/interface025/expected rename to tests/idris2/interface/interface025/expected diff --git a/tests/idris2/interface026/input b/tests/idris2/interface/interface025/input similarity index 100% rename from tests/idris2/interface026/input rename to tests/idris2/interface/interface025/input diff --git a/tests/idris2/interface025/run b/tests/idris2/interface/interface025/run similarity index 60% rename from tests/idris2/interface025/run rename to tests/idris2/interface/interface025/run index 4493d93c10..7550da9290 100644 --- a/tests/idris2/interface025/run +++ b/tests/idris2/interface/interface025/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 AutoSearchHide2.idr < input diff --git a/tests/idris2/interface026/UninhabitedRec.idr b/tests/idris2/interface/interface026/UninhabitedRec.idr similarity index 100% rename from tests/idris2/interface026/UninhabitedRec.idr rename to tests/idris2/interface/interface026/UninhabitedRec.idr diff --git a/tests/idris2/interface026/expected b/tests/idris2/interface/interface026/expected similarity index 100% rename from tests/idris2/interface026/expected rename to tests/idris2/interface/interface026/expected diff --git a/tests/idris2/record007/input b/tests/idris2/interface/interface026/input similarity index 100% rename from tests/idris2/record007/input rename to tests/idris2/interface/interface026/input diff --git a/tests/idris2/interface026/run b/tests/idris2/interface/interface026/run similarity index 59% rename from tests/idris2/interface026/run rename to tests/idris2/interface/interface026/run index c3d09e3f6c..c301964182 100644 --- a/tests/idris2/interface026/run +++ b/tests/idris2/interface/interface026/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 UninhabitedRec.idr < input diff --git a/tests/idris2/interface027/expected b/tests/idris2/interface/interface027/expected similarity index 100% rename from tests/idris2/interface027/expected rename to tests/idris2/interface/interface027/expected diff --git a/tests/idris2/interface027/input b/tests/idris2/interface/interface027/input similarity index 100% rename from tests/idris2/interface027/input rename to tests/idris2/interface/interface027/input diff --git a/tests/idris2/interface027/params.idr b/tests/idris2/interface/interface027/params.idr similarity index 100% rename from tests/idris2/interface027/params.idr rename to tests/idris2/interface/interface027/params.idr diff --git a/tests/idris2/interface027/run b/tests/idris2/interface/interface027/run similarity index 52% rename from tests/idris2/interface027/run rename to tests/idris2/interface/interface027/run index adb525f312..86b61c53e0 100644 --- a/tests/idris2/interface027/run +++ b/tests/idris2/interface/interface027/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 params.idr < input diff --git a/tests/idris2/interface028/InterfaceArgs.idr b/tests/idris2/interface/interface028/InterfaceArgs.idr similarity index 100% rename from tests/idris2/interface028/InterfaceArgs.idr rename to tests/idris2/interface/interface028/InterfaceArgs.idr diff --git a/tests/idris2/interface028/expected b/tests/idris2/interface/interface028/expected similarity index 100% rename from tests/idris2/interface028/expected rename to tests/idris2/interface/interface028/expected diff --git a/tests/idris2/interface028/run b/tests/idris2/interface/interface028/run similarity index 51% rename from tests/idris2/interface028/run rename to tests/idris2/interface/interface028/run index cec8cd1b4b..eae471ca27 100644 --- a/tests/idris2/interface028/run +++ b/tests/idris2/interface/interface028/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh check InterfaceArgs.idr diff --git a/tests/idris2/interface029/ForwardImpl1.idr b/tests/idris2/interface/interface029/ForwardImpl1.idr similarity index 100% rename from tests/idris2/interface029/ForwardImpl1.idr rename to tests/idris2/interface/interface029/ForwardImpl1.idr diff --git a/tests/idris2/interface029/ForwardImpl2.idr b/tests/idris2/interface/interface029/ForwardImpl2.idr similarity index 100% rename from tests/idris2/interface029/ForwardImpl2.idr rename to tests/idris2/interface/interface029/ForwardImpl2.idr diff --git a/tests/idris2/interface029/ForwardImpl3.idr b/tests/idris2/interface/interface029/ForwardImpl3.idr similarity index 100% rename from tests/idris2/interface029/ForwardImpl3.idr rename to tests/idris2/interface/interface029/ForwardImpl3.idr diff --git a/tests/idris2/interface029/expected b/tests/idris2/interface/interface029/expected similarity index 100% rename from tests/idris2/interface029/expected rename to tests/idris2/interface/interface029/expected diff --git a/tests/idris2/interface029/run b/tests/idris2/interface/interface029/run similarity index 73% rename from tests/idris2/interface029/run rename to tests/idris2/interface/interface029/run index 9c15937d55..bb7c8ce7b2 100644 --- a/tests/idris2/interface029/run +++ b/tests/idris2/interface/interface029/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh run ForwardImpl1.idr check ForwardImpl2.idr diff --git a/tests/idris2/interface003/run b/tests/idris2/interface003/run deleted file mode 100755 index 070cba614b..0000000000 --- a/tests/idris2/interface003/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -idris2 Do.idr < input diff --git a/tests/idris2/interface004/run b/tests/idris2/interface004/run deleted file mode 100755 index 070cba614b..0000000000 --- a/tests/idris2/interface004/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -idris2 Do.idr < input diff --git a/tests/idris2/interface005/run b/tests/idris2/interface005/run deleted file mode 100755 index 5b4317229e..0000000000 --- a/tests/idris2/interface005/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check Deps.idr diff --git a/tests/idris2/interface006/run b/tests/idris2/interface006/run deleted file mode 100755 index fe3c8cd87c..0000000000 --- a/tests/idris2/interface006/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check Bimonad.idr diff --git a/tests/idris2/interface008/run b/tests/idris2/interface008/run deleted file mode 100755 index 5b4317229e..0000000000 --- a/tests/idris2/interface008/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check Deps.idr diff --git a/tests/idris2/interface010/run b/tests/idris2/interface010/run deleted file mode 100755 index c7c671d744..0000000000 --- a/tests/idris2/interface010/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check Dep.idr diff --git a/tests/idris2/interface011/run b/tests/idris2/interface011/run deleted file mode 100755 index e4c61ef462..0000000000 --- a/tests/idris2/interface011/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check FuncImpl.idr diff --git a/tests/idris2/interface012/run b/tests/idris2/interface012/run deleted file mode 100755 index 82ae339542..0000000000 --- a/tests/idris2/interface012/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check Defmeth.idr diff --git a/tests/idris2/interface013/run b/tests/idris2/interface013/run deleted file mode 100755 index 9bb60668eb..0000000000 --- a/tests/idris2/interface013/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check TypeInt.idr diff --git a/tests/idris2/interface014/run b/tests/idris2/interface014/run deleted file mode 100755 index d12a274b1c..0000000000 --- a/tests/idris2/interface014/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check DepInt.idr diff --git a/tests/idris2/interface015/run b/tests/idris2/interface015/run deleted file mode 100755 index d1a25333b9..0000000000 --- a/tests/idris2/interface015/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check gnu.idr diff --git a/tests/idris2/interface016/run b/tests/idris2/interface016/run deleted file mode 100755 index 25ba0b8b70..0000000000 --- a/tests/idris2/interface016/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check TwoNum.idr diff --git a/tests/idris2/interface017/run b/tests/idris2/interface017/run deleted file mode 100755 index fdc050c0bf..0000000000 --- a/tests/idris2/interface017/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check Tricho.idr diff --git a/tests/idris2/interface019/run b/tests/idris2/interface019/run deleted file mode 100755 index 6d8e1e5605..0000000000 --- a/tests/idris2/interface019/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check LocalHints.idr diff --git a/tests/idris2/interface021/run b/tests/idris2/interface021/run deleted file mode 100755 index a0611e6d32..0000000000 --- a/tests/idris2/interface021/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check LocalHint.idr diff --git a/tests/idris2/interpolation003/run b/tests/idris2/interpolation003/run deleted file mode 100755 index 441b64ba30..0000000000 --- a/tests/idris2/interpolation003/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -idris2 Test.idr < input diff --git a/tests/idris2/interpreter001/run b/tests/idris2/interpreter001/run deleted file mode 100755 index fecfb4b107..0000000000 --- a/tests/idris2/interpreter001/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -idris2 < input diff --git a/tests/idris2/interpreter003/run b/tests/idris2/interpreter003/run deleted file mode 100755 index fecfb4b107..0000000000 --- a/tests/idris2/interpreter003/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -idris2 < input diff --git a/tests/idris2/interpreter004/run b/tests/idris2/interpreter004/run deleted file mode 100755 index fecfb4b107..0000000000 --- a/tests/idris2/interpreter004/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -idris2 < input diff --git a/tests/idris2/interpreter006/run b/tests/idris2/interpreter006/run deleted file mode 100755 index fecfb4b107..0000000000 --- a/tests/idris2/interpreter006/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -idris2 < input diff --git a/tests/idris2/interpreter007/run b/tests/idris2/interpreter007/run deleted file mode 100755 index fecfb4b107..0000000000 --- a/tests/idris2/interpreter007/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -idris2 < input diff --git a/tests/idris2/linear001/Door.idr b/tests/idris2/linear/linear001/Door.idr similarity index 100% rename from tests/idris2/linear001/Door.idr rename to tests/idris2/linear/linear001/Door.idr diff --git a/tests/idris2/linear001/Stuff.idr b/tests/idris2/linear/linear001/Stuff.idr similarity index 100% rename from tests/idris2/linear001/Stuff.idr rename to tests/idris2/linear/linear001/Stuff.idr diff --git a/tests/idris2/linear001/expected b/tests/idris2/linear/linear001/expected similarity index 100% rename from tests/idris2/linear001/expected rename to tests/idris2/linear/linear001/expected diff --git a/tests/idris2/linear001/run b/tests/idris2/linear/linear001/run similarity index 63% rename from tests/idris2/linear001/run rename to tests/idris2/linear/linear001/run index d2e5d9d81f..c8a3a17c68 100755 --- a/tests/idris2/linear001/run +++ b/tests/idris2/linear/linear001/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh echo ':q' | idris2 --no-prelude Door.idr diff --git a/tests/idris2/linear002/Door.idr b/tests/idris2/linear/linear002/Door.idr similarity index 100% rename from tests/idris2/linear002/Door.idr rename to tests/idris2/linear/linear002/Door.idr diff --git a/tests/idris2/linear002/Stuff.idr b/tests/idris2/linear/linear002/Stuff.idr similarity index 100% rename from tests/idris2/linear002/Stuff.idr rename to tests/idris2/linear/linear002/Stuff.idr diff --git a/tests/idris2/linear002/expected b/tests/idris2/linear/linear002/expected similarity index 100% rename from tests/idris2/linear002/expected rename to tests/idris2/linear/linear002/expected diff --git a/tests/idris2/linear002/input b/tests/idris2/linear/linear002/input similarity index 100% rename from tests/idris2/linear002/input rename to tests/idris2/linear/linear002/input diff --git a/tests/idris2/linear002/run b/tests/idris2/linear/linear002/run similarity index 61% rename from tests/idris2/linear002/run rename to tests/idris2/linear/linear002/run index 2708c0d900..5c41f583f1 100755 --- a/tests/idris2/linear002/run +++ b/tests/idris2/linear/linear002/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 --no-prelude Door.idr < input diff --git a/tests/idris2/linear003/Linear.idr b/tests/idris2/linear/linear003/Linear.idr similarity index 100% rename from tests/idris2/linear003/Linear.idr rename to tests/idris2/linear/linear003/Linear.idr diff --git a/tests/idris2/linear003/expected b/tests/idris2/linear/linear003/expected similarity index 100% rename from tests/idris2/linear003/expected rename to tests/idris2/linear/linear003/expected diff --git a/tests/idris2/linear003/input b/tests/idris2/linear/linear003/input similarity index 100% rename from tests/idris2/linear003/input rename to tests/idris2/linear/linear003/input diff --git a/tests/idris2/linear003/run b/tests/idris2/linear/linear003/run similarity index 62% rename from tests/idris2/linear003/run rename to tests/idris2/linear/linear003/run index f67775b529..928ea58024 100755 --- a/tests/idris2/linear003/run +++ b/tests/idris2/linear/linear003/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 --no-prelude Linear.idr < input diff --git a/tests/idris2/linear005/Door.idr b/tests/idris2/linear/linear005/Door.idr similarity index 100% rename from tests/idris2/linear005/Door.idr rename to tests/idris2/linear/linear005/Door.idr diff --git a/tests/idris2/linear005/Linear.idr b/tests/idris2/linear/linear005/Linear.idr similarity index 100% rename from tests/idris2/linear005/Linear.idr rename to tests/idris2/linear/linear005/Linear.idr diff --git a/tests/idris2/linear005/expected b/tests/idris2/linear/linear005/expected similarity index 100% rename from tests/idris2/linear005/expected rename to tests/idris2/linear/linear005/expected diff --git a/tests/idris2/linear005/input b/tests/idris2/linear/linear005/input similarity index 100% rename from tests/idris2/linear005/input rename to tests/idris2/linear/linear005/input diff --git a/tests/idris2/linear005/run b/tests/idris2/linear/linear005/run similarity index 51% rename from tests/idris2/linear005/run rename to tests/idris2/linear/linear005/run index e6c63e9986..402b977549 100755 --- a/tests/idris2/linear005/run +++ b/tests/idris2/linear/linear005/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 Door.idr < input diff --git a/tests/idris2/linear006/ZFun.idr b/tests/idris2/linear/linear006/ZFun.idr similarity index 100% rename from tests/idris2/linear006/ZFun.idr rename to tests/idris2/linear/linear006/ZFun.idr diff --git a/tests/idris2/linear006/expected b/tests/idris2/linear/linear006/expected similarity index 100% rename from tests/idris2/linear006/expected rename to tests/idris2/linear/linear006/expected diff --git a/tests/idris2/linear006/input b/tests/idris2/linear/linear006/input similarity index 100% rename from tests/idris2/linear006/input rename to tests/idris2/linear/linear006/input diff --git a/tests/idris2/linear006/run b/tests/idris2/linear/linear006/run similarity index 51% rename from tests/idris2/linear006/run rename to tests/idris2/linear/linear006/run index 29a7dcd48e..995fe8d251 100755 --- a/tests/idris2/linear006/run +++ b/tests/idris2/linear/linear006/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 ZFun.idr < input diff --git a/tests/idris2/linear007/LCase.idr b/tests/idris2/linear/linear007/LCase.idr similarity index 100% rename from tests/idris2/linear007/LCase.idr rename to tests/idris2/linear/linear007/LCase.idr diff --git a/tests/idris2/linear007/expected b/tests/idris2/linear/linear007/expected similarity index 100% rename from tests/idris2/linear007/expected rename to tests/idris2/linear/linear007/expected diff --git a/tests/idris2/linear/linear007/run b/tests/idris2/linear/linear007/run new file mode 100755 index 0000000000..6d421ca947 --- /dev/null +++ b/tests/idris2/linear/linear007/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check LCase.idr diff --git a/tests/idris2/linear008/Door.idr b/tests/idris2/linear/linear008/Door.idr similarity index 100% rename from tests/idris2/linear008/Door.idr rename to tests/idris2/linear/linear008/Door.idr diff --git a/tests/idris2/linear008/expected b/tests/idris2/linear/linear008/expected similarity index 100% rename from tests/idris2/linear008/expected rename to tests/idris2/linear/linear008/expected diff --git a/tests/idris2/linear/linear008/run b/tests/idris2/linear/linear008/run new file mode 100644 index 0000000000..ab10707c53 --- /dev/null +++ b/tests/idris2/linear/linear008/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Door.idr diff --git a/tests/idris2/linear009/expected b/tests/idris2/linear/linear009/expected similarity index 100% rename from tests/idris2/linear009/expected rename to tests/idris2/linear/linear009/expected diff --git a/tests/idris2/linear009/input b/tests/idris2/linear/linear009/input similarity index 100% rename from tests/idris2/linear009/input rename to tests/idris2/linear/linear009/input diff --git a/tests/idris2/linear009/qtt.idr b/tests/idris2/linear/linear009/qtt.idr similarity index 100% rename from tests/idris2/linear009/qtt.idr rename to tests/idris2/linear/linear009/qtt.idr diff --git a/tests/idris2/linear009/run b/tests/idris2/linear/linear009/run similarity index 50% rename from tests/idris2/linear009/run rename to tests/idris2/linear/linear009/run index 68e408f2ff..dcabf5b95c 100644 --- a/tests/idris2/linear009/run +++ b/tests/idris2/linear/linear009/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 qtt.idr < input diff --git a/tests/idris2/linear010/Door.idr b/tests/idris2/linear/linear010/Door.idr similarity index 100% rename from tests/idris2/linear010/Door.idr rename to tests/idris2/linear/linear010/Door.idr diff --git a/tests/idris2/linear010/expected b/tests/idris2/linear/linear010/expected similarity index 100% rename from tests/idris2/linear010/expected rename to tests/idris2/linear/linear010/expected diff --git a/tests/idris2/linear010/run b/tests/idris2/linear/linear010/run similarity index 52% rename from tests/idris2/linear010/run rename to tests/idris2/linear/linear010/run index 683ee35b90..83b87f2103 100644 --- a/tests/idris2/linear010/run +++ b/tests/idris2/linear/linear010/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh check Door.idr -p linear diff --git a/tests/idris2/linear011/Network.idr b/tests/idris2/linear/linear011/Network.idr similarity index 100% rename from tests/idris2/linear011/Network.idr rename to tests/idris2/linear/linear011/Network.idr diff --git a/tests/idris2/linear011/expected b/tests/idris2/linear/linear011/expected similarity index 100% rename from tests/idris2/linear011/expected rename to tests/idris2/linear/linear011/expected diff --git a/tests/idris2/linear011/input b/tests/idris2/linear/linear011/input similarity index 100% rename from tests/idris2/linear011/input rename to tests/idris2/linear/linear011/input diff --git a/tests/idris2/linear011/run b/tests/idris2/linear/linear011/run similarity index 67% rename from tests/idris2/linear011/run rename to tests/idris2/linear/linear011/run index c2f4ab208a..82073dd594 100644 --- a/tests/idris2/linear011/run +++ b/tests/idris2/linear/linear011/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 Network.idr -p linear -p network < input diff --git a/tests/idris2/linear012/expected b/tests/idris2/linear/linear012/expected similarity index 100% rename from tests/idris2/linear012/expected rename to tests/idris2/linear/linear012/expected diff --git a/tests/idris2/linear012/input b/tests/idris2/linear/linear012/input similarity index 100% rename from tests/idris2/linear012/input rename to tests/idris2/linear/linear012/input diff --git a/tests/idris2/linear012/linholes.idr b/tests/idris2/linear/linear012/linholes.idr similarity index 100% rename from tests/idris2/linear012/linholes.idr rename to tests/idris2/linear/linear012/linholes.idr diff --git a/tests/idris2/linear012/run b/tests/idris2/linear/linear012/run similarity index 54% rename from tests/idris2/linear012/run rename to tests/idris2/linear/linear012/run index a4b65dd45f..56842615b5 100644 --- a/tests/idris2/linear012/run +++ b/tests/idris2/linear/linear012/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 linholes.idr < input diff --git a/tests/idris2/linear013/Issue758.idr b/tests/idris2/linear/linear013/Issue758.idr similarity index 100% rename from tests/idris2/linear013/Issue758.idr rename to tests/idris2/linear/linear013/Issue758.idr diff --git a/tests/idris2/linear013/expected b/tests/idris2/linear/linear013/expected similarity index 100% rename from tests/idris2/linear013/expected rename to tests/idris2/linear/linear013/expected diff --git a/tests/idris2/linear013/input b/tests/idris2/linear/linear013/input similarity index 100% rename from tests/idris2/linear013/input rename to tests/idris2/linear/linear013/input diff --git a/tests/idris2/linear013/run b/tests/idris2/linear/linear013/run similarity index 54% rename from tests/idris2/linear013/run rename to tests/idris2/linear/linear013/run index b5082d31ac..43c34aaf50 100644 --- a/tests/idris2/linear013/run +++ b/tests/idris2/linear/linear013/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 Issue758.idr < input diff --git a/tests/idris2/linear014/Issue55.idr b/tests/idris2/linear/linear014/Issue55.idr similarity index 100% rename from tests/idris2/linear014/Issue55.idr rename to tests/idris2/linear/linear014/Issue55.idr diff --git a/tests/idris2/linear014/expected b/tests/idris2/linear/linear014/expected similarity index 100% rename from tests/idris2/linear014/expected rename to tests/idris2/linear/linear014/expected diff --git a/tests/idris2/linear/linear014/run b/tests/idris2/linear/linear014/run new file mode 100644 index 0000000000..0c18d67556 --- /dev/null +++ b/tests/idris2/linear/linear014/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Issue55.idr diff --git a/tests/idris2/linear015/Issue1861.idr b/tests/idris2/linear/linear015/Issue1861.idr similarity index 100% rename from tests/idris2/linear015/Issue1861.idr rename to tests/idris2/linear/linear015/Issue1861.idr diff --git a/tests/idris2/linear015/expected b/tests/idris2/linear/linear015/expected similarity index 100% rename from tests/idris2/linear015/expected rename to tests/idris2/linear/linear015/expected diff --git a/tests/idris2/linear/linear015/run b/tests/idris2/linear/linear015/run new file mode 100644 index 0000000000..fa6aeacdbb --- /dev/null +++ b/tests/idris2/linear/linear015/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Issue1861.idr diff --git a/tests/idris2/linear016/Issue2895.idr b/tests/idris2/linear/linear016/Issue2895.idr similarity index 100% rename from tests/idris2/linear016/Issue2895.idr rename to tests/idris2/linear/linear016/Issue2895.idr diff --git a/tests/idris2/linear016/Issue2895_2.idr b/tests/idris2/linear/linear016/Issue2895_2.idr similarity index 100% rename from tests/idris2/linear016/Issue2895_2.idr rename to tests/idris2/linear/linear016/Issue2895_2.idr diff --git a/tests/idris2/linear016/expected b/tests/idris2/linear/linear016/expected similarity index 100% rename from tests/idris2/linear016/expected rename to tests/idris2/linear/linear016/expected diff --git a/tests/idris2/linear016/run b/tests/idris2/linear/linear016/run similarity index 64% rename from tests/idris2/linear016/run rename to tests/idris2/linear/linear016/run index 2bf190b2d0..b3438181ae 100644 --- a/tests/idris2/linear016/run +++ b/tests/idris2/linear/linear016/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh check Issue2895.idr check Issue2895_2.idr diff --git a/tests/idris2/linear017/As.idr b/tests/idris2/linear/linear017/As.idr similarity index 100% rename from tests/idris2/linear017/As.idr rename to tests/idris2/linear/linear017/As.idr diff --git a/tests/idris2/linear017/expected b/tests/idris2/linear/linear017/expected similarity index 100% rename from tests/idris2/linear017/expected rename to tests/idris2/linear/linear017/expected diff --git a/tests/idris2/linear/linear017/run b/tests/idris2/linear/linear017/run new file mode 100644 index 0000000000..4ffffa89da --- /dev/null +++ b/tests/idris2/linear/linear017/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check As.idr diff --git a/tests/idris2/linear004/Erase.idr b/tests/idris2/linear004/Erase.idr deleted file mode 100644 index 97045c6583..0000000000 --- a/tests/idris2/linear004/Erase.idr +++ /dev/null @@ -1,13 +0,0 @@ -import Stuff - -efn : ((0 x : Nat) -> Nat -> Nat) -> Nat -efn f = f (S Z) (S Z) - -okfn : ((x : Nat) -> Nat -> Nat) -> Nat -okfn f = f (S Z) (S Z) - -ignore : (0 x : Nat) -> Nat -> Nat -ignore x y = y - -lin : (1 x : Nat) -> Nat -> Nat -lin x y = S x diff --git a/tests/idris2/linear004/Stuff.idr b/tests/idris2/linear004/Stuff.idr deleted file mode 100644 index abb61fa56a..0000000000 --- a/tests/idris2/linear004/Stuff.idr +++ /dev/null @@ -1,68 +0,0 @@ --- a mini prelude - -module Stuff - -public export -data Bool = True | False - -public export -not : Bool -> Bool -not True = False -not False = True - -public export -data Maybe : Type -> Type where - Nothing : Maybe a - Just : (1 x : a) -> Maybe a - -public export -data Either : Type -> Type -> Type where - Left : (1 x : a) -> Either a b - Right : (1 y : b) -> Either a b - -public export -intToBool : Int -> Bool -intToBool 0 = False -intToBool x = True - -public export -ifThenElse : Bool -> Lazy a -> Lazy a -> a -ifThenElse True t e = t -ifThenElse False t e = e - -public export -data Nat : Type where - Z : Nat - S : (1 _ : Nat) -> Nat - -public export -fromInteger : Integer -> Nat -fromInteger x = ifThenElse (intToBool (prim__eq_Integer x 0)) - Z (S (fromInteger (prim__sub_Integer x 1))) - -public export -plus : Nat -> Nat -> Nat -plus Z y = y -plus (S k) y = S (plus k y) - -infixr 5 :: - -public export -data List : Type -> Type where - Nil : List a - (::) : (1 _ : a) -> (1 _ : List a) -> List a - -public export -data Eq : a -> b -> Type where - Refl : (x : a) -> Eq x x - -public export -data Unit = MkUnit - -public export -data Pair : Type -> Type -> Type where - MkPair : {a, b : Type} -> a -> b -> Pair a b - -public export -the : (a : Type) -> a -> a -the _ x = x diff --git a/tests/idris2/linear004/expected b/tests/idris2/linear004/expected deleted file mode 100644 index 4df8edfb55..0000000000 --- a/tests/idris2/linear004/expected +++ /dev/null @@ -1,46 +0,0 @@ -1/2: Building Stuff (Stuff.idr) -2/2: Building Erase (Erase.idr) -Main> S Z -Main> S (S Z) -Main> S Z -Main> Error: x is not accessible in this context. - -(interactive):1:15--1:16 - | - 1 | efn (\x, y => x) -- Bad - | ^ - -Main> Error: When unifying Nat -> Nat -> Nat and (0 _ : Nat) -> Nat -> Nat. -Mismatch between: Nat -> Nat -> Nat and (0 _ : Nat) -> Nat -> Nat. - -(interactive):1:5--1:9 - | - 1 | efn plus -- Bad - | ^^^^ - -Main> Error: When unifying (1 _ : Nat) -> Nat -> Nat and (0 _ : Nat) -> Nat -> Nat. -Mismatch between: (1 _ : Nat) -> Nat -> Nat and (0 _ : Nat) -> Nat -> Nat. - -(interactive):1:5--1:8 - | - 1 | efn lin -- Bad - | ^^^ - -Main> Error: x is not accessible in this context. - -(interactive):1:20--1:21 - | - 1 | efn (\x, y => plus x y) -- Bad - | ^ - -Main> S (S Z) -Main> S (S Z) -Main> Error: When unifying (0 _ : Nat) -> Nat -> Nat and Nat -> Nat -> Nat. -Mismatch between: (0 _ : Nat) -> Nat -> Nat and Nat -> Nat -> Nat. - -(interactive):1:6--1:12 - | - 1 | okfn ignore -- Bad - | ^^^^^^ - -Main> Bye for now! diff --git a/tests/idris2/linear004/input b/tests/idris2/linear004/input deleted file mode 100644 index 50aa69d66f..0000000000 --- a/tests/idris2/linear004/input +++ /dev/null @@ -1,11 +0,0 @@ -efn (\x, y => y) -- Good -efn (\x, y => plus y y) -- Good -efn ignore -- Good -efn (\x, y => x) -- Bad -efn plus -- Bad -efn lin -- Bad -efn (\x, y => plus x y) -- Bad -okfn plus -- Good -okfn lin -- Good -okfn ignore -- Bad -:q diff --git a/tests/idris2/linear004/run b/tests/idris2/linear004/run deleted file mode 100755 index 4f4a5a77bc..0000000000 --- a/tests/idris2/linear004/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -idris2 --no-prelude Erase.idr < input diff --git a/tests/idris2/linear007/run b/tests/idris2/linear007/run deleted file mode 100755 index a74a408b45..0000000000 --- a/tests/idris2/linear007/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check LCase.idr diff --git a/tests/idris2/linear008/run b/tests/idris2/linear008/run deleted file mode 100644 index 8f9c6d68cf..0000000000 --- a/tests/idris2/linear008/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check Door.idr diff --git a/tests/idris2/linear014/run b/tests/idris2/linear014/run deleted file mode 100644 index 18b94f9894..0000000000 --- a/tests/idris2/linear014/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check Issue55.idr diff --git a/tests/idris2/linear015/run b/tests/idris2/linear015/run deleted file mode 100644 index b27c98a40c..0000000000 --- a/tests/idris2/linear015/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check Issue1861.idr diff --git a/tests/idris2/linear017/run b/tests/idris2/linear017/run deleted file mode 100644 index e5c3405335..0000000000 --- a/tests/idris2/linear017/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check As.idr diff --git a/tests/idris2/literate001/IEdit.lidr b/tests/idris2/literate/literate001/IEdit.lidr similarity index 100% rename from tests/idris2/literate001/IEdit.lidr rename to tests/idris2/literate/literate001/IEdit.lidr diff --git a/tests/idris2/literate001/expected b/tests/idris2/literate/literate001/expected similarity index 100% rename from tests/idris2/literate001/expected rename to tests/idris2/literate/literate001/expected diff --git a/tests/idris2/literate001/input b/tests/idris2/literate/literate001/input similarity index 100% rename from tests/idris2/literate001/input rename to tests/idris2/literate/literate001/input diff --git a/tests/idris2/literate001/run b/tests/idris2/literate/literate001/run similarity index 52% rename from tests/idris2/literate001/run rename to tests/idris2/literate/literate001/run index b2447bcbe5..cb65d877e0 100755 --- a/tests/idris2/literate001/run +++ b/tests/idris2/literate/literate001/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 IEdit.lidr < input diff --git a/tests/idris2/literate002/IEdit.lidr b/tests/idris2/literate/literate002/IEdit.lidr similarity index 100% rename from tests/idris2/literate002/IEdit.lidr rename to tests/idris2/literate/literate002/IEdit.lidr diff --git a/tests/idris2/literate002/IEdit2.lidr b/tests/idris2/literate/literate002/IEdit2.lidr similarity index 100% rename from tests/idris2/literate002/IEdit2.lidr rename to tests/idris2/literate/literate002/IEdit2.lidr diff --git a/tests/idris2/literate002/expected b/tests/idris2/literate/literate002/expected similarity index 100% rename from tests/idris2/literate002/expected rename to tests/idris2/literate/literate002/expected diff --git a/tests/idris2/literate002/input b/tests/idris2/literate/literate002/input similarity index 100% rename from tests/idris2/literate002/input rename to tests/idris2/literate/literate002/input diff --git a/tests/idris2/literate002/input2 b/tests/idris2/literate/literate002/input2 similarity index 100% rename from tests/idris2/literate002/input2 rename to tests/idris2/literate/literate002/input2 diff --git a/tests/idris2/literate002/run b/tests/idris2/literate/literate002/run similarity index 69% rename from tests/idris2/literate002/run rename to tests/idris2/literate/literate002/run index 8da12871cc..bce0ba01e2 100755 --- a/tests/idris2/literate002/run +++ b/tests/idris2/literate/literate002/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 IEdit.lidr < input idris2 IEdit2.lidr < input2 diff --git a/tests/idris2/literate003/IEdit.lidr b/tests/idris2/literate/literate003/IEdit.lidr similarity index 100% rename from tests/idris2/literate003/IEdit.lidr rename to tests/idris2/literate/literate003/IEdit.lidr diff --git a/tests/idris2/literate003/expected b/tests/idris2/literate/literate003/expected similarity index 100% rename from tests/idris2/literate003/expected rename to tests/idris2/literate/literate003/expected diff --git a/tests/idris2/literate003/input b/tests/idris2/literate/literate003/input similarity index 100% rename from tests/idris2/literate003/input rename to tests/idris2/literate/literate003/input diff --git a/tests/idris2/literate005/run b/tests/idris2/literate/literate003/run similarity index 52% rename from tests/idris2/literate005/run rename to tests/idris2/literate/literate003/run index b2447bcbe5..cb65d877e0 100755 --- a/tests/idris2/literate005/run +++ b/tests/idris2/literate/literate003/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 IEdit.lidr < input diff --git a/tests/idris2/literate004/IEdit.lidr b/tests/idris2/literate/literate004/IEdit.lidr similarity index 100% rename from tests/idris2/literate004/IEdit.lidr rename to tests/idris2/literate/literate004/IEdit.lidr diff --git a/tests/idris2/literate004/expected b/tests/idris2/literate/literate004/expected similarity index 100% rename from tests/idris2/literate004/expected rename to tests/idris2/literate/literate004/expected diff --git a/tests/idris2/literate004/input b/tests/idris2/literate/literate004/input similarity index 100% rename from tests/idris2/literate004/input rename to tests/idris2/literate/literate004/input diff --git a/tests/idris2/literate003/run b/tests/idris2/literate/literate004/run similarity index 52% rename from tests/idris2/literate003/run rename to tests/idris2/literate/literate004/run index b2447bcbe5..cb65d877e0 100755 --- a/tests/idris2/literate003/run +++ b/tests/idris2/literate/literate004/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 IEdit.lidr < input diff --git a/tests/idris2/literate005/IEdit.lidr b/tests/idris2/literate/literate005/IEdit.lidr similarity index 100% rename from tests/idris2/literate005/IEdit.lidr rename to tests/idris2/literate/literate005/IEdit.lidr diff --git a/tests/idris2/literate005/expected b/tests/idris2/literate/literate005/expected similarity index 100% rename from tests/idris2/literate005/expected rename to tests/idris2/literate/literate005/expected diff --git a/tests/idris2/literate005/input b/tests/idris2/literate/literate005/input similarity index 100% rename from tests/idris2/literate005/input rename to tests/idris2/literate/literate005/input diff --git a/tests/idris2/literate004/run b/tests/idris2/literate/literate005/run similarity index 52% rename from tests/idris2/literate004/run rename to tests/idris2/literate/literate005/run index b2447bcbe5..cb65d877e0 100755 --- a/tests/idris2/literate004/run +++ b/tests/idris2/literate/literate005/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 IEdit.lidr < input diff --git a/tests/idris2/literate006/Door.lidr b/tests/idris2/literate/literate006/Door.lidr similarity index 100% rename from tests/idris2/literate006/Door.lidr rename to tests/idris2/literate/literate006/Door.lidr diff --git a/tests/idris2/literate006/expected b/tests/idris2/literate/literate006/expected similarity index 100% rename from tests/idris2/literate006/expected rename to tests/idris2/literate/literate006/expected diff --git a/tests/idris2/literate006/input b/tests/idris2/literate/literate006/input similarity index 100% rename from tests/idris2/literate006/input rename to tests/idris2/literate/literate006/input diff --git a/tests/idris2/literate006/run b/tests/idris2/literate/literate006/run similarity index 52% rename from tests/idris2/literate006/run rename to tests/idris2/literate/literate006/run index 25014df805..106f531c50 100755 --- a/tests/idris2/literate006/run +++ b/tests/idris2/literate/literate006/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 Door.lidr < input diff --git a/tests/idris2/literate007/IEdit.lidr b/tests/idris2/literate/literate007/IEdit.lidr similarity index 100% rename from tests/idris2/literate007/IEdit.lidr rename to tests/idris2/literate/literate007/IEdit.lidr diff --git a/tests/idris2/literate007/IEdit.org b/tests/idris2/literate/literate007/IEdit.org similarity index 100% rename from tests/idris2/literate007/IEdit.org rename to tests/idris2/literate/literate007/IEdit.org diff --git a/tests/idris2/literate007/IEditOrg.org b/tests/idris2/literate/literate007/IEditOrg.org similarity index 100% rename from tests/idris2/literate007/IEditOrg.org rename to tests/idris2/literate/literate007/IEditOrg.org diff --git a/tests/idris2/literate007/expected b/tests/idris2/literate/literate007/expected similarity index 100% rename from tests/idris2/literate007/expected rename to tests/idris2/literate/literate007/expected diff --git a/tests/idris2/literate007/input b/tests/idris2/literate/literate007/input similarity index 100% rename from tests/idris2/literate007/input rename to tests/idris2/literate/literate007/input diff --git a/tests/idris2/literate007/input2 b/tests/idris2/literate/literate007/input2 similarity index 100% rename from tests/idris2/literate007/input2 rename to tests/idris2/literate/literate007/input2 diff --git a/tests/idris2/literate007/run b/tests/idris2/literate/literate007/run similarity index 70% rename from tests/idris2/literate007/run rename to tests/idris2/literate/literate007/run index 728d112866..2c088bec07 100755 --- a/tests/idris2/literate007/run +++ b/tests/idris2/literate/literate007/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 IEdit.lidr < input idris2 IEditOrg.org < input2 diff --git a/tests/idris2/literate008/IEdit.lidr b/tests/idris2/literate/literate008/IEdit.lidr similarity index 100% rename from tests/idris2/literate008/IEdit.lidr rename to tests/idris2/literate/literate008/IEdit.lidr diff --git a/tests/idris2/literate008/expected b/tests/idris2/literate/literate008/expected similarity index 100% rename from tests/idris2/literate008/expected rename to tests/idris2/literate/literate008/expected diff --git a/tests/idris2/literate008/input b/tests/idris2/literate/literate008/input similarity index 100% rename from tests/idris2/literate008/input rename to tests/idris2/literate/literate008/input diff --git a/tests/idris2/literate/literate008/run b/tests/idris2/literate/literate008/run new file mode 100755 index 0000000000..cb65d877e0 --- /dev/null +++ b/tests/idris2/literate/literate008/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +idris2 IEdit.lidr < input diff --git a/tests/idris2/literate009/WithLift.lidr b/tests/idris2/literate/literate009/WithLift.lidr similarity index 100% rename from tests/idris2/literate009/WithLift.lidr rename to tests/idris2/literate/literate009/WithLift.lidr diff --git a/tests/idris2/literate009/expected b/tests/idris2/literate/literate009/expected similarity index 100% rename from tests/idris2/literate009/expected rename to tests/idris2/literate/literate009/expected diff --git a/tests/idris2/literate009/input b/tests/idris2/literate/literate009/input similarity index 100% rename from tests/idris2/literate009/input rename to tests/idris2/literate/literate009/input diff --git a/tests/idris2/literate009/run b/tests/idris2/literate/literate009/run similarity index 55% rename from tests/idris2/literate009/run rename to tests/idris2/literate/literate009/run index 988f2bef06..89a3d292de 100755 --- a/tests/idris2/literate009/run +++ b/tests/idris2/literate/literate009/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 WithLift.lidr < input diff --git a/tests/idris2/literate010/MyFirstIdrisProgram.org b/tests/idris2/literate/literate010/MyFirstIdrisProgram.org similarity index 100% rename from tests/idris2/literate010/MyFirstIdrisProgram.org rename to tests/idris2/literate/literate010/MyFirstIdrisProgram.org diff --git a/tests/idris2/literate010/expected b/tests/idris2/literate/literate010/expected similarity index 100% rename from tests/idris2/literate010/expected rename to tests/idris2/literate/literate010/expected diff --git a/tests/idris2/literate010/run b/tests/idris2/literate/literate010/run similarity index 56% rename from tests/idris2/literate010/run rename to tests/idris2/literate/literate010/run index ef39e14ed1..de8d81ad93 100755 --- a/tests/idris2/literate010/run +++ b/tests/idris2/literate/literate010/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh check MyFirstIdrisProgram.org diff --git a/tests/idris2/literate011/IEdit.md b/tests/idris2/literate/literate011/IEdit.md similarity index 100% rename from tests/idris2/literate011/IEdit.md rename to tests/idris2/literate/literate011/IEdit.md diff --git a/tests/idris2/literate011/expected b/tests/idris2/literate/literate011/expected similarity index 100% rename from tests/idris2/literate011/expected rename to tests/idris2/literate/literate011/expected diff --git a/tests/idris2/literate011/input b/tests/idris2/literate/literate011/input similarity index 100% rename from tests/idris2/literate011/input rename to tests/idris2/literate/literate011/input diff --git a/tests/idris2/literate011/run b/tests/idris2/literate/literate011/run similarity index 51% rename from tests/idris2/literate011/run rename to tests/idris2/literate/literate011/run index 4c62db0e31..6194ccf3a4 100755 --- a/tests/idris2/literate011/run +++ b/tests/idris2/literate/literate011/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 IEdit.md < input diff --git a/tests/idris2/literate012/IEdit.org b/tests/idris2/literate/literate012/IEdit.org similarity index 100% rename from tests/idris2/literate012/IEdit.org rename to tests/idris2/literate/literate012/IEdit.org diff --git a/tests/idris2/literate012/IEdit2.org b/tests/idris2/literate/literate012/IEdit2.org similarity index 100% rename from tests/idris2/literate012/IEdit2.org rename to tests/idris2/literate/literate012/IEdit2.org diff --git a/tests/idris2/literate012/expected b/tests/idris2/literate/literate012/expected similarity index 100% rename from tests/idris2/literate012/expected rename to tests/idris2/literate/literate012/expected diff --git a/tests/idris2/literate012/input b/tests/idris2/literate/literate012/input similarity index 100% rename from tests/idris2/literate012/input rename to tests/idris2/literate/literate012/input diff --git a/tests/idris2/literate012/run b/tests/idris2/literate/literate012/run similarity index 52% rename from tests/idris2/literate012/run rename to tests/idris2/literate/literate012/run index 5a24d321f9..51b18fb049 100755 --- a/tests/idris2/literate012/run +++ b/tests/idris2/literate/literate012/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 IEdit.org < input diff --git a/tests/idris2/literate013/Lit.lidr b/tests/idris2/literate/literate013/Lit.lidr similarity index 100% rename from tests/idris2/literate013/Lit.lidr rename to tests/idris2/literate/literate013/Lit.lidr diff --git a/tests/idris2/literate013/LitTeX.tex b/tests/idris2/literate/literate013/LitTeX.tex similarity index 100% rename from tests/idris2/literate013/LitTeX.tex rename to tests/idris2/literate/literate013/LitTeX.tex diff --git a/tests/idris2/literate013/expected b/tests/idris2/literate/literate013/expected similarity index 100% rename from tests/idris2/literate013/expected rename to tests/idris2/literate/literate013/expected diff --git a/tests/idris2/literate013/run b/tests/idris2/literate/literate013/run similarity index 57% rename from tests/idris2/literate013/run rename to tests/idris2/literate/literate013/run index 7106e0f5d7..cb70331b0c 100755 --- a/tests/idris2/literate013/run +++ b/tests/idris2/literate/literate013/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh check Lit.lidr check LitTeX.tex diff --git a/tests/idris2/literate014/expected b/tests/idris2/literate/literate014/expected similarity index 100% rename from tests/idris2/literate014/expected rename to tests/idris2/literate/literate014/expected diff --git a/tests/idris2/literate014/input b/tests/idris2/literate/literate014/input similarity index 100% rename from tests/idris2/literate014/input rename to tests/idris2/literate/literate014/input diff --git a/tests/idris2/literate014/run b/tests/idris2/literate/literate014/run similarity index 52% rename from tests/idris2/literate014/run rename to tests/idris2/literate/literate014/run index 685423ab3f..de7a5fd4a5 100755 --- a/tests/idris2/literate014/run +++ b/tests/idris2/literate/literate014/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 with.lidr < input diff --git a/tests/idris2/literate014/with.lidr b/tests/idris2/literate/literate014/with.lidr similarity index 100% rename from tests/idris2/literate014/with.lidr rename to tests/idris2/literate/literate014/with.lidr diff --git a/tests/idris2/literate015/case.lidr b/tests/idris2/literate/literate015/case.lidr similarity index 100% rename from tests/idris2/literate015/case.lidr rename to tests/idris2/literate/literate015/case.lidr diff --git a/tests/idris2/literate015/expected b/tests/idris2/literate/literate015/expected similarity index 100% rename from tests/idris2/literate015/expected rename to tests/idris2/literate/literate015/expected diff --git a/tests/idris2/literate015/input b/tests/idris2/literate/literate015/input similarity index 100% rename from tests/idris2/literate015/input rename to tests/idris2/literate/literate015/input diff --git a/tests/idris2/literate015/run b/tests/idris2/literate/literate015/run similarity index 52% rename from tests/idris2/literate015/run rename to tests/idris2/literate/literate015/run index 55a453de6d..684fbd9374 100755 --- a/tests/idris2/literate015/run +++ b/tests/idris2/literate/literate015/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 case.lidr < input diff --git a/tests/idris2/literate016/IEdit.org b/tests/idris2/literate/literate016/IEdit.org similarity index 100% rename from tests/idris2/literate016/IEdit.org rename to tests/idris2/literate/literate016/IEdit.org diff --git a/tests/idris2/literate016/IEdit2.org b/tests/idris2/literate/literate016/IEdit2.org similarity index 100% rename from tests/idris2/literate016/IEdit2.org rename to tests/idris2/literate/literate016/IEdit2.org diff --git a/tests/idris2/literate016/expected b/tests/idris2/literate/literate016/expected similarity index 100% rename from tests/idris2/literate016/expected rename to tests/idris2/literate/literate016/expected diff --git a/tests/idris2/literate016/input b/tests/idris2/literate/literate016/input similarity index 100% rename from tests/idris2/literate016/input rename to tests/idris2/literate/literate016/input diff --git a/tests/idris2/literate016/run b/tests/idris2/literate/literate016/run similarity index 52% rename from tests/idris2/literate016/run rename to tests/idris2/literate/literate016/run index 5a24d321f9..51b18fb049 100755 --- a/tests/idris2/literate016/run +++ b/tests/idris2/literate/literate016/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 IEdit.org < input diff --git a/tests/idris2/literate017/.gitignore b/tests/idris2/literate/literate017/.gitignore similarity index 100% rename from tests/idris2/literate017/.gitignore rename to tests/idris2/literate/literate017/.gitignore diff --git a/tests/idris2/literate017/expected b/tests/idris2/literate/literate017/expected similarity index 100% rename from tests/idris2/literate017/expected rename to tests/idris2/literate/literate017/expected diff --git a/tests/idris2/literate017/input b/tests/idris2/literate/literate017/input similarity index 100% rename from tests/idris2/literate017/input rename to tests/idris2/literate/literate017/input diff --git a/tests/idris2/literate017/project-expected.ipkg b/tests/idris2/literate/literate017/project-expected.ipkg similarity index 100% rename from tests/idris2/literate017/project-expected.ipkg rename to tests/idris2/literate/literate017/project-expected.ipkg diff --git a/tests/idris2/literate017/run b/tests/idris2/literate/literate017/run similarity index 82% rename from tests/idris2/literate017/run rename to tests/idris2/literate/literate017/run index 55708b1e07..a6da6d40c7 100755 --- a/tests/idris2/literate017/run +++ b/tests/idris2/literate/literate017/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh rm project-output.ipkg diff --git a/tests/idris2/literate017/src/A/A.org b/tests/idris2/literate/literate017/src/A/A.org similarity index 100% rename from tests/idris2/literate017/src/A/A.org rename to tests/idris2/literate/literate017/src/A/A.org diff --git a/tests/idris2/literate017/src/A/AB.lidr b/tests/idris2/literate/literate017/src/A/AB.lidr similarity index 100% rename from tests/idris2/literate017/src/A/AB.lidr rename to tests/idris2/literate/literate017/src/A/AB.lidr diff --git a/tests/idris2/literate017/src/B/B.md b/tests/idris2/literate/literate017/src/B/B.md similarity index 100% rename from tests/idris2/literate017/src/B/B.md rename to tests/idris2/literate/literate017/src/B/B.md diff --git a/tests/idris2/literate017/src/B/BA.markdown b/tests/idris2/literate/literate017/src/B/BA.markdown similarity index 100% rename from tests/idris2/literate017/src/B/BA.markdown rename to tests/idris2/literate/literate017/src/B/BA.markdown diff --git a/tests/idris2/literate017/src/Main.idr b/tests/idris2/literate/literate017/src/Main.idr similarity index 100% rename from tests/idris2/literate017/src/Main.idr rename to tests/idris2/literate/literate017/src/Main.idr diff --git a/tests/idris2/literate008/run b/tests/idris2/literate008/run deleted file mode 100755 index b2447bcbe5..0000000000 --- a/tests/idris2/literate008/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -idris2 IEdit.lidr < input diff --git a/tests/idris2/docs001/expected b/tests/idris2/misc/docs001/expected similarity index 100% rename from tests/idris2/docs001/expected rename to tests/idris2/misc/docs001/expected diff --git a/tests/idris2/docs001/input b/tests/idris2/misc/docs001/input similarity index 100% rename from tests/idris2/docs001/input rename to tests/idris2/misc/docs001/input diff --git a/tests/idris2/misc/docs001/run b/tests/idris2/misc/docs001/run new file mode 100755 index 0000000000..404d7d47c6 --- /dev/null +++ b/tests/idris2/misc/docs001/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +idris2 < input diff --git a/tests/idris2/docs002/Doc.idr b/tests/idris2/misc/docs002/Doc.idr similarity index 100% rename from tests/idris2/docs002/Doc.idr rename to tests/idris2/misc/docs002/Doc.idr diff --git a/tests/idris2/docs002/expected b/tests/idris2/misc/docs002/expected similarity index 100% rename from tests/idris2/docs002/expected rename to tests/idris2/misc/docs002/expected diff --git a/tests/idris2/docs002/input b/tests/idris2/misc/docs002/input similarity index 100% rename from tests/idris2/docs002/input rename to tests/idris2/misc/docs002/input diff --git a/tests/idris2/docs002/run b/tests/idris2/misc/docs002/run similarity index 50% rename from tests/idris2/docs002/run rename to tests/idris2/misc/docs002/run index 75f6c27108..657e133ca0 100755 --- a/tests/idris2/docs002/run +++ b/tests/idris2/misc/docs002/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 Doc.idr < input diff --git a/tests/idris2/docs003/RecordDoc.idr b/tests/idris2/misc/docs003/RecordDoc.idr similarity index 100% rename from tests/idris2/docs003/RecordDoc.idr rename to tests/idris2/misc/docs003/RecordDoc.idr diff --git a/tests/idris2/docs003/expected b/tests/idris2/misc/docs003/expected similarity index 100% rename from tests/idris2/docs003/expected rename to tests/idris2/misc/docs003/expected diff --git a/tests/idris2/docs003/input b/tests/idris2/misc/docs003/input similarity index 100% rename from tests/idris2/docs003/input rename to tests/idris2/misc/docs003/input diff --git a/tests/idris2/docs003/run b/tests/idris2/misc/docs003/run similarity index 55% rename from tests/idris2/docs003/run rename to tests/idris2/misc/docs003/run index 0dfbac62c4..6233d25ca8 100755 --- a/tests/idris2/docs003/run +++ b/tests/idris2/misc/docs003/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 RecordDoc.idr < input diff --git a/tests/idris2/docs004/DocImpl.idr b/tests/idris2/misc/docs004/DocImpl.idr similarity index 100% rename from tests/idris2/docs004/DocImpl.idr rename to tests/idris2/misc/docs004/DocImpl.idr diff --git a/tests/idris2/docs004/List.idr b/tests/idris2/misc/docs004/List.idr similarity index 100% rename from tests/idris2/docs004/List.idr rename to tests/idris2/misc/docs004/List.idr diff --git a/tests/idris2/docs004/expected b/tests/idris2/misc/docs004/expected similarity index 100% rename from tests/idris2/docs004/expected rename to tests/idris2/misc/docs004/expected diff --git a/tests/idris2/docs004/input b/tests/idris2/misc/docs004/input similarity index 100% rename from tests/idris2/docs004/input rename to tests/idris2/misc/docs004/input diff --git a/tests/idris2/docs004/input2 b/tests/idris2/misc/docs004/input2 similarity index 100% rename from tests/idris2/docs004/input2 rename to tests/idris2/misc/docs004/input2 diff --git a/tests/idris2/docs004/run b/tests/idris2/misc/docs004/run similarity index 73% rename from tests/idris2/docs004/run rename to tests/idris2/misc/docs004/run index 5a4328a2f6..de06ed166a 100755 --- a/tests/idris2/docs004/run +++ b/tests/idris2/misc/docs004/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 DocImpl.idr < input idris2 --no-prelude List.idr < input2 diff --git a/tests/idris2/docs005/expected b/tests/idris2/misc/docs005/expected similarity index 100% rename from tests/idris2/docs005/expected rename to tests/idris2/misc/docs005/expected diff --git a/tests/idris2/docs005/input b/tests/idris2/misc/docs005/input similarity index 100% rename from tests/idris2/docs005/input rename to tests/idris2/misc/docs005/input diff --git a/tests/idris2/misc/docs005/run b/tests/idris2/misc/docs005/run new file mode 100755 index 0000000000..404d7d47c6 --- /dev/null +++ b/tests/idris2/misc/docs005/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +idris2 < input diff --git a/tests/idris2/eta001/Issue1370.idr b/tests/idris2/misc/eta001/Issue1370.idr similarity index 100% rename from tests/idris2/eta001/Issue1370.idr rename to tests/idris2/misc/eta001/Issue1370.idr diff --git a/tests/idris2/eta001/expected b/tests/idris2/misc/eta001/expected similarity index 100% rename from tests/idris2/eta001/expected rename to tests/idris2/misc/eta001/expected diff --git a/tests/idris2/misc/eta001/run b/tests/idris2/misc/eta001/run new file mode 100755 index 0000000000..c34a860916 --- /dev/null +++ b/tests/idris2/misc/eta001/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Issue1370.idr diff --git a/tests/idris2/golden001/000-hello/expected b/tests/idris2/misc/golden001/000-hello/expected similarity index 100% rename from tests/idris2/golden001/000-hello/expected rename to tests/idris2/misc/golden001/000-hello/expected diff --git a/tests/idris2/golden001/000-hello/run b/tests/idris2/misc/golden001/000-hello/run old mode 100644 new mode 100755 similarity index 100% rename from tests/idris2/golden001/000-hello/run rename to tests/idris2/misc/golden001/000-hello/run diff --git a/tests/idris2/golden001/Main.idr b/tests/idris2/misc/golden001/Main.idr similarity index 100% rename from tests/idris2/golden001/Main.idr rename to tests/idris2/misc/golden001/Main.idr diff --git a/tests/idris2/golden001/Test.idr b/tests/idris2/misc/golden001/Test.idr similarity index 100% rename from tests/idris2/golden001/Test.idr rename to tests/idris2/misc/golden001/Test.idr diff --git a/tests/idris2/golden001/expected b/tests/idris2/misc/golden001/expected similarity index 100% rename from tests/idris2/golden001/expected rename to tests/idris2/misc/golden001/expected diff --git a/tests/idris2/golden001/hello.ipkg b/tests/idris2/misc/golden001/hello.ipkg similarity index 100% rename from tests/idris2/golden001/hello.ipkg rename to tests/idris2/misc/golden001/hello.ipkg diff --git a/tests/idris2/golden001/run b/tests/idris2/misc/golden001/run similarity index 81% rename from tests/idris2/golden001/run rename to tests/idris2/misc/golden001/run index 0b64af4e20..d749c8d157 100755 --- a/tests/idris2/golden001/run +++ b/tests/idris2/misc/golden001/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 --build hello.ipkg idris2 --build test.ipkg diff --git a/tests/idris2/golden001/test.ipkg b/tests/idris2/misc/golden001/test.ipkg similarity index 100% rename from tests/idris2/golden001/test.ipkg rename to tests/idris2/misc/golden001/test.ipkg diff --git a/tests/idris2/import001/Mult.idr b/tests/idris2/misc/import001/Mult.idr similarity index 100% rename from tests/idris2/import001/Mult.idr rename to tests/idris2/misc/import001/Mult.idr diff --git a/tests/idris2/import001/Nat.idr b/tests/idris2/misc/import001/Nat.idr similarity index 100% rename from tests/idris2/import001/Nat.idr rename to tests/idris2/misc/import001/Nat.idr diff --git a/tests/idris2/import001/Test.idr b/tests/idris2/misc/import001/Test.idr similarity index 100% rename from tests/idris2/import001/Test.idr rename to tests/idris2/misc/import001/Test.idr diff --git a/tests/idris2/import001/expected b/tests/idris2/misc/import001/expected similarity index 100% rename from tests/idris2/import001/expected rename to tests/idris2/misc/import001/expected diff --git a/tests/idris2/import001/input b/tests/idris2/misc/import001/input similarity index 100% rename from tests/idris2/import001/input rename to tests/idris2/misc/import001/input diff --git a/tests/idris2/import001/run b/tests/idris2/misc/import001/run similarity index 90% rename from tests/idris2/import001/run rename to tests/idris2/misc/import001/run index 5a383000ed..4a2c68bf8c 100755 --- a/tests/idris2/import001/run +++ b/tests/idris2/misc/import001/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 --no-prelude Test.idr < input sleep 1 diff --git a/tests/idris2/import002/Mult.idr b/tests/idris2/misc/import002/Mult.idr similarity index 100% rename from tests/idris2/import002/Mult.idr rename to tests/idris2/misc/import002/Mult.idr diff --git a/tests/idris2/import002/Nat.idr b/tests/idris2/misc/import002/Nat.idr similarity index 100% rename from tests/idris2/import002/Nat.idr rename to tests/idris2/misc/import002/Nat.idr diff --git a/tests/idris2/import002/Test.idr b/tests/idris2/misc/import002/Test.idr similarity index 100% rename from tests/idris2/import002/Test.idr rename to tests/idris2/misc/import002/Test.idr diff --git a/tests/idris2/import002/expected b/tests/idris2/misc/import002/expected similarity index 100% rename from tests/idris2/import002/expected rename to tests/idris2/misc/import002/expected diff --git a/tests/idris2/import002/run b/tests/idris2/misc/import002/run similarity index 54% rename from tests/idris2/import002/run rename to tests/idris2/misc/import002/run index 7ce7a33b8f..2790f478da 100755 --- a/tests/idris2/import002/run +++ b/tests/idris2/misc/import002/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh check --no-prelude Test.idr diff --git a/tests/idris2/import003/A.idr b/tests/idris2/misc/import003/A.idr similarity index 100% rename from tests/idris2/import003/A.idr rename to tests/idris2/misc/import003/A.idr diff --git a/tests/idris2/import003/B.idr b/tests/idris2/misc/import003/B.idr similarity index 100% rename from tests/idris2/import003/B.idr rename to tests/idris2/misc/import003/B.idr diff --git a/tests/idris2/import003/C.idr b/tests/idris2/misc/import003/C.idr similarity index 100% rename from tests/idris2/import003/C.idr rename to tests/idris2/misc/import003/C.idr diff --git a/tests/idris2/import003/expected b/tests/idris2/misc/import003/expected similarity index 100% rename from tests/idris2/import003/expected rename to tests/idris2/misc/import003/expected diff --git a/tests/idris2/import003/input b/tests/idris2/misc/import003/input similarity index 100% rename from tests/idris2/import003/input rename to tests/idris2/misc/import003/input diff --git a/tests/idris2/import003/run b/tests/idris2/misc/import003/run similarity index 64% rename from tests/idris2/import003/run rename to tests/idris2/misc/import003/run index dcb7fa69b6..3da20812a1 100755 --- a/tests/idris2/import003/run +++ b/tests/idris2/misc/import003/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 B.idr < input idris2 C.idr < input diff --git a/tests/idris2/import004/Cycle1.idr b/tests/idris2/misc/import004/Cycle1.idr similarity index 100% rename from tests/idris2/import004/Cycle1.idr rename to tests/idris2/misc/import004/Cycle1.idr diff --git a/tests/idris2/import004/Cycle2.idr b/tests/idris2/misc/import004/Cycle2.idr similarity index 100% rename from tests/idris2/import004/Cycle2.idr rename to tests/idris2/misc/import004/Cycle2.idr diff --git a/tests/idris2/import004/Loop.idr b/tests/idris2/misc/import004/Loop.idr similarity index 100% rename from tests/idris2/import004/Loop.idr rename to tests/idris2/misc/import004/Loop.idr diff --git a/tests/idris2/import004/expected b/tests/idris2/misc/import004/expected similarity index 100% rename from tests/idris2/import004/expected rename to tests/idris2/misc/import004/expected diff --git a/tests/idris2/import004/run b/tests/idris2/misc/import004/run similarity index 59% rename from tests/idris2/import004/run rename to tests/idris2/misc/import004/run index 6cfa986dca..1b8f346b55 100755 --- a/tests/idris2/import004/run +++ b/tests/idris2/misc/import004/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 Cycle1.idr idris2 Loop.idr diff --git a/tests/idris2/import005/As.idr b/tests/idris2/misc/import005/As.idr similarity index 100% rename from tests/idris2/import005/As.idr rename to tests/idris2/misc/import005/As.idr diff --git a/tests/idris2/import005/Test.idr b/tests/idris2/misc/import005/Test.idr similarity index 100% rename from tests/idris2/import005/Test.idr rename to tests/idris2/misc/import005/Test.idr diff --git a/tests/idris2/import005/expected b/tests/idris2/misc/import005/expected similarity index 100% rename from tests/idris2/import005/expected rename to tests/idris2/misc/import005/expected diff --git a/tests/idris2/import005/input b/tests/idris2/misc/import005/input similarity index 100% rename from tests/idris2/import005/input rename to tests/idris2/misc/import005/input diff --git a/tests/idris2/misc/import005/run b/tests/idris2/misc/import005/run new file mode 100644 index 0000000000..0ede21378e --- /dev/null +++ b/tests/idris2/misc/import005/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +idris2 As.idr < input diff --git a/tests/idris2/import006/A/B.idr b/tests/idris2/misc/import006/A/B.idr similarity index 100% rename from tests/idris2/import006/A/B.idr rename to tests/idris2/misc/import006/A/B.idr diff --git a/tests/idris2/import006/A/C.idr b/tests/idris2/misc/import006/A/C.idr similarity index 100% rename from tests/idris2/import006/A/C.idr rename to tests/idris2/misc/import006/A/C.idr diff --git a/tests/idris2/import006/cyclic.ipkg b/tests/idris2/misc/import006/cyclic.ipkg similarity index 100% rename from tests/idris2/import006/cyclic.ipkg rename to tests/idris2/misc/import006/cyclic.ipkg diff --git a/tests/idris2/import006/expected b/tests/idris2/misc/import006/expected similarity index 100% rename from tests/idris2/import006/expected rename to tests/idris2/misc/import006/expected diff --git a/tests/idris2/import006/run b/tests/idris2/misc/import006/run similarity index 53% rename from tests/idris2/import006/run rename to tests/idris2/misc/import006/run index b79def1a3d..e00127a91a 100644 --- a/tests/idris2/import006/run +++ b/tests/idris2/misc/import006/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 --build cyclic.ipkg diff --git a/tests/idris2/import007/Mod.idr b/tests/idris2/misc/import007/Mod.idr similarity index 100% rename from tests/idris2/import007/Mod.idr rename to tests/idris2/misc/import007/Mod.idr diff --git a/tests/idris2/import007/Mod1.idr b/tests/idris2/misc/import007/Mod1.idr similarity index 100% rename from tests/idris2/import007/Mod1.idr rename to tests/idris2/misc/import007/Mod1.idr diff --git a/tests/idris2/import007/Mod2.idr b/tests/idris2/misc/import007/Mod2.idr similarity index 100% rename from tests/idris2/import007/Mod2.idr rename to tests/idris2/misc/import007/Mod2.idr diff --git a/tests/idris2/import007/expected b/tests/idris2/misc/import007/expected similarity index 100% rename from tests/idris2/import007/expected rename to tests/idris2/misc/import007/expected diff --git a/tests/idris2/misc/import007/run b/tests/idris2/misc/import007/run new file mode 100644 index 0000000000..9dfa487e45 --- /dev/null +++ b/tests/idris2/misc/import007/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Mod.idr diff --git a/tests/idris2/import008/Exe/Mod.idr b/tests/idris2/misc/import008/Exe/Mod.idr similarity index 100% rename from tests/idris2/import008/Exe/Mod.idr rename to tests/idris2/misc/import008/Exe/Mod.idr diff --git a/tests/idris2/import008/Exe/depends/lib-0 b/tests/idris2/misc/import008/Exe/depends/lib-0 similarity index 100% rename from tests/idris2/import008/Exe/depends/lib-0 rename to tests/idris2/misc/import008/Exe/depends/lib-0 diff --git a/tests/idris2/import008/Exe/exe.ipkg b/tests/idris2/misc/import008/Exe/exe.ipkg similarity index 100% rename from tests/idris2/import008/Exe/exe.ipkg rename to tests/idris2/misc/import008/Exe/exe.ipkg diff --git a/tests/idris2/import008/Lib/Conf.idr b/tests/idris2/misc/import008/Lib/Conf.idr similarity index 100% rename from tests/idris2/import008/Lib/Conf.idr rename to tests/idris2/misc/import008/Lib/Conf.idr diff --git a/tests/idris2/import008/Lib/lib.ipkg b/tests/idris2/misc/import008/Lib/lib.ipkg similarity index 100% rename from tests/idris2/import008/Lib/lib.ipkg rename to tests/idris2/misc/import008/Lib/lib.ipkg diff --git a/tests/idris2/import008/expected b/tests/idris2/misc/import008/expected similarity index 100% rename from tests/idris2/import008/expected rename to tests/idris2/misc/import008/expected diff --git a/tests/idris2/import008/run b/tests/idris2/misc/import008/run similarity index 83% rename from tests/idris2/import008/run rename to tests/idris2/misc/import008/run index cd798e25c4..ef960640bb 100644 --- a/tests/idris2/import008/run +++ b/tests/idris2/misc/import008/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh rm -rf Exe/build rm -rf Lib/build diff --git a/tests/idris2/import009/Import.idr b/tests/idris2/misc/import009/Import.idr similarity index 100% rename from tests/idris2/import009/Import.idr rename to tests/idris2/misc/import009/Import.idr diff --git a/tests/idris2/import009/Infix.idr b/tests/idris2/misc/import009/Infix.idr similarity index 100% rename from tests/idris2/import009/Infix.idr rename to tests/idris2/misc/import009/Infix.idr diff --git a/tests/idris2/import009/Prefix.idr b/tests/idris2/misc/import009/Prefix.idr similarity index 100% rename from tests/idris2/import009/Prefix.idr rename to tests/idris2/misc/import009/Prefix.idr diff --git a/tests/idris2/import009/Resugar.idr b/tests/idris2/misc/import009/Resugar.idr similarity index 100% rename from tests/idris2/import009/Resugar.idr rename to tests/idris2/misc/import009/Resugar.idr diff --git a/tests/idris2/import009/Test.idr b/tests/idris2/misc/import009/Test.idr similarity index 100% rename from tests/idris2/import009/Test.idr rename to tests/idris2/misc/import009/Test.idr diff --git a/tests/idris2/import009/expected b/tests/idris2/misc/import009/expected similarity index 100% rename from tests/idris2/import009/expected rename to tests/idris2/misc/import009/expected diff --git a/tests/idris2/import009/run b/tests/idris2/misc/import009/run similarity index 68% rename from tests/idris2/import009/run rename to tests/idris2/misc/import009/run index c3e8173aff..7beee22b90 100755 --- a/tests/idris2/import009/run +++ b/tests/idris2/misc/import009/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh check Test.idr check Import.idr diff --git a/tests/idris2/inlining001/Inlining.idr b/tests/idris2/misc/inlining001/Inlining.idr similarity index 100% rename from tests/idris2/inlining001/Inlining.idr rename to tests/idris2/misc/inlining001/Inlining.idr diff --git a/tests/idris2/inlining001/expected b/tests/idris2/misc/inlining001/expected similarity index 100% rename from tests/idris2/inlining001/expected rename to tests/idris2/misc/inlining001/expected diff --git a/tests/idris2/inlining001/input b/tests/idris2/misc/inlining001/input similarity index 100% rename from tests/idris2/inlining001/input rename to tests/idris2/misc/inlining001/input diff --git a/tests/idris2/inlining001/run b/tests/idris2/misc/inlining001/run similarity index 54% rename from tests/idris2/inlining001/run rename to tests/idris2/misc/inlining001/run index 7cef0d3038..cfbfc4a5e9 100644 --- a/tests/idris2/inlining001/run +++ b/tests/idris2/misc/inlining001/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 Inlining.idr < input diff --git a/tests/idris2/lazy001/Lazy.idr b/tests/idris2/misc/lazy001/Lazy.idr similarity index 100% rename from tests/idris2/lazy001/Lazy.idr rename to tests/idris2/misc/lazy001/Lazy.idr diff --git a/tests/idris2/lazy001/expected b/tests/idris2/misc/lazy001/expected similarity index 100% rename from tests/idris2/lazy001/expected rename to tests/idris2/misc/lazy001/expected diff --git a/tests/idris2/lazy001/input b/tests/idris2/misc/lazy001/input similarity index 100% rename from tests/idris2/lazy001/input rename to tests/idris2/misc/lazy001/input diff --git a/tests/idris2/lazy001/run b/tests/idris2/misc/lazy001/run similarity index 61% rename from tests/idris2/lazy001/run rename to tests/idris2/misc/lazy001/run index 781b1d3526..9ab94207b8 100755 --- a/tests/idris2/lazy001/run +++ b/tests/idris2/misc/lazy001/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 --no-prelude Lazy.idr < input diff --git a/tests/idris2/lazy002/LazyFoldlM.idr b/tests/idris2/misc/lazy002/LazyFoldlM.idr similarity index 100% rename from tests/idris2/lazy002/LazyFoldlM.idr rename to tests/idris2/misc/lazy002/LazyFoldlM.idr diff --git a/tests/idris2/lazy002/expected b/tests/idris2/misc/lazy002/expected similarity index 100% rename from tests/idris2/lazy002/expected rename to tests/idris2/misc/lazy002/expected diff --git a/tests/idris2/lazy002/input b/tests/idris2/misc/lazy002/input similarity index 100% rename from tests/idris2/lazy002/input rename to tests/idris2/misc/lazy002/input diff --git a/tests/idris2/lazy002/run b/tests/idris2/misc/lazy002/run similarity index 63% rename from tests/idris2/lazy002/run rename to tests/idris2/misc/lazy002/run index 96bad64a01..e2fe2ed2a3 100755 --- a/tests/idris2/lazy002/run +++ b/tests/idris2/misc/lazy002/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 -p contrib LazyFoldlM.idr < input diff --git a/tests/idris2/lazy003/DelayLam.idr b/tests/idris2/misc/lazy003/DelayLam.idr similarity index 100% rename from tests/idris2/lazy003/DelayLam.idr rename to tests/idris2/misc/lazy003/DelayLam.idr diff --git a/tests/idris2/lazy003/expected b/tests/idris2/misc/lazy003/expected similarity index 100% rename from tests/idris2/lazy003/expected rename to tests/idris2/misc/lazy003/expected diff --git a/tests/idris2/lazy003/input b/tests/idris2/misc/lazy003/input similarity index 100% rename from tests/idris2/lazy003/input rename to tests/idris2/misc/lazy003/input diff --git a/tests/idris2/lazy003/run b/tests/idris2/misc/lazy003/run similarity index 54% rename from tests/idris2/lazy003/run rename to tests/idris2/misc/lazy003/run index c33f96d18b..1f938e6744 100644 --- a/tests/idris2/lazy003/run +++ b/tests/idris2/misc/lazy003/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 DelayLam.idr < input diff --git a/tests/idris2/namespace001/Dup.idr b/tests/idris2/misc/namespace001/Dup.idr similarity index 100% rename from tests/idris2/namespace001/Dup.idr rename to tests/idris2/misc/namespace001/Dup.idr diff --git a/tests/idris2/namespace001/Scope.idr b/tests/idris2/misc/namespace001/Scope.idr similarity index 100% rename from tests/idris2/namespace001/Scope.idr rename to tests/idris2/misc/namespace001/Scope.idr diff --git a/tests/idris2/namespace001/expected b/tests/idris2/misc/namespace001/expected similarity index 100% rename from tests/idris2/namespace001/expected rename to tests/idris2/misc/namespace001/expected diff --git a/tests/idris2/namespace001/run b/tests/idris2/misc/namespace001/run similarity index 72% rename from tests/idris2/namespace001/run rename to tests/idris2/misc/namespace001/run index e6d0e3e284..023234e827 100644 --- a/tests/idris2/namespace001/run +++ b/tests/idris2/misc/namespace001/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh echo ':q' | idris2 Dup.idr echo ':t Test' | idris2 Scope.idr diff --git a/tests/idris2/namespace002/Issue1313.idr b/tests/idris2/misc/namespace002/Issue1313.idr similarity index 100% rename from tests/idris2/namespace002/Issue1313.idr rename to tests/idris2/misc/namespace002/Issue1313.idr diff --git a/tests/idris2/namespace002/expected b/tests/idris2/misc/namespace002/expected similarity index 100% rename from tests/idris2/namespace002/expected rename to tests/idris2/misc/namespace002/expected diff --git a/tests/idris2/namespace002/run b/tests/idris2/misc/namespace002/run similarity index 54% rename from tests/idris2/namespace002/run rename to tests/idris2/misc/namespace002/run index fdffd3d555..ee6bbcfc5f 100644 --- a/tests/idris2/namespace002/run +++ b/tests/idris2/misc/namespace002/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh check Issue1313.idr || true diff --git a/tests/idris2/namespace003/Test.idr b/tests/idris2/misc/namespace003/Test.idr similarity index 100% rename from tests/idris2/namespace003/Test.idr rename to tests/idris2/misc/namespace003/Test.idr diff --git a/tests/idris2/namespace003/expected b/tests/idris2/misc/namespace003/expected similarity index 100% rename from tests/idris2/namespace003/expected rename to tests/idris2/misc/namespace003/expected diff --git a/tests/idris2/namespace003/run b/tests/idris2/misc/namespace003/run similarity index 50% rename from tests/idris2/namespace003/run rename to tests/idris2/misc/namespace003/run index 248924d6b2..09eacddf5d 100644 --- a/tests/idris2/namespace003/run +++ b/tests/idris2/misc/namespace003/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh check Test.idr || true diff --git a/tests/idris2/namespace004/Export.idr b/tests/idris2/misc/namespace004/Export.idr similarity index 100% rename from tests/idris2/namespace004/Export.idr rename to tests/idris2/misc/namespace004/Export.idr diff --git a/tests/idris2/namespace004/expected b/tests/idris2/misc/namespace004/expected similarity index 100% rename from tests/idris2/namespace004/expected rename to tests/idris2/misc/namespace004/expected diff --git a/tests/idris2/misc/namespace004/run b/tests/idris2/misc/namespace004/run new file mode 100644 index 0000000000..9da8cf2d7a --- /dev/null +++ b/tests/idris2/misc/namespace004/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Export.idr diff --git a/tests/idris2/namespace005/Lib1.idr b/tests/idris2/misc/namespace005/Lib1.idr similarity index 100% rename from tests/idris2/namespace005/Lib1.idr rename to tests/idris2/misc/namespace005/Lib1.idr diff --git a/tests/idris2/namespace005/Lib2.idr b/tests/idris2/misc/namespace005/Lib2.idr similarity index 100% rename from tests/idris2/namespace005/Lib2.idr rename to tests/idris2/misc/namespace005/Lib2.idr diff --git a/tests/idris2/namespace005/LibPre1.idr b/tests/idris2/misc/namespace005/LibPre1.idr similarity index 100% rename from tests/idris2/namespace005/LibPre1.idr rename to tests/idris2/misc/namespace005/LibPre1.idr diff --git a/tests/idris2/namespace005/LibPre2.idr b/tests/idris2/misc/namespace005/LibPre2.idr similarity index 100% rename from tests/idris2/namespace005/LibPre2.idr rename to tests/idris2/misc/namespace005/LibPre2.idr diff --git a/tests/idris2/namespace005/Main0.idr b/tests/idris2/misc/namespace005/Main0.idr similarity index 100% rename from tests/idris2/namespace005/Main0.idr rename to tests/idris2/misc/namespace005/Main0.idr diff --git a/tests/idris2/namespace005/Main1.idr b/tests/idris2/misc/namespace005/Main1.idr similarity index 100% rename from tests/idris2/namespace005/Main1.idr rename to tests/idris2/misc/namespace005/Main1.idr diff --git a/tests/idris2/namespace005/Main3.idr b/tests/idris2/misc/namespace005/Main3.idr similarity index 100% rename from tests/idris2/namespace005/Main3.idr rename to tests/idris2/misc/namespace005/Main3.idr diff --git a/tests/idris2/namespace005/MainConflict.idr b/tests/idris2/misc/namespace005/MainConflict.idr similarity index 100% rename from tests/idris2/namespace005/MainConflict.idr rename to tests/idris2/misc/namespace005/MainConflict.idr diff --git a/tests/idris2/namespace005/MainFail.idr b/tests/idris2/misc/namespace005/MainFail.idr similarity index 100% rename from tests/idris2/namespace005/MainFail.idr rename to tests/idris2/misc/namespace005/MainFail.idr diff --git a/tests/idris2/namespace005/MainPre0.idr b/tests/idris2/misc/namespace005/MainPre0.idr similarity index 100% rename from tests/idris2/namespace005/MainPre0.idr rename to tests/idris2/misc/namespace005/MainPre0.idr diff --git a/tests/idris2/namespace005/MainPre1.idr b/tests/idris2/misc/namespace005/MainPre1.idr similarity index 100% rename from tests/idris2/namespace005/MainPre1.idr rename to tests/idris2/misc/namespace005/MainPre1.idr diff --git a/tests/idris2/namespace005/NonConflict1.idr b/tests/idris2/misc/namespace005/NonConflict1.idr similarity index 100% rename from tests/idris2/namespace005/NonConflict1.idr rename to tests/idris2/misc/namespace005/NonConflict1.idr diff --git a/tests/idris2/namespace005/NonConflict2.idr b/tests/idris2/misc/namespace005/NonConflict2.idr similarity index 100% rename from tests/idris2/namespace005/NonConflict2.idr rename to tests/idris2/misc/namespace005/NonConflict2.idr diff --git a/tests/idris2/namespace005/expected b/tests/idris2/misc/namespace005/expected similarity index 100% rename from tests/idris2/namespace005/expected rename to tests/idris2/misc/namespace005/expected diff --git a/tests/idris2/namespace005/run b/tests/idris2/misc/namespace005/run similarity index 90% rename from tests/idris2/namespace005/run rename to tests/idris2/misc/namespace005/run index 5f4086519a..3186181b12 100755 --- a/tests/idris2/namespace005/run +++ b/tests/idris2/misc/namespace005/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh check Main0.idr check Main1.idr diff --git a/tests/idris2/params001/expected b/tests/idris2/misc/params001/expected similarity index 100% rename from tests/idris2/params001/expected rename to tests/idris2/misc/params001/expected diff --git a/tests/idris2/params001/param.idr b/tests/idris2/misc/params001/param.idr similarity index 100% rename from tests/idris2/params001/param.idr rename to tests/idris2/misc/params001/param.idr diff --git a/tests/idris2/params001/parambad.idr b/tests/idris2/misc/params001/parambad.idr similarity index 100% rename from tests/idris2/params001/parambad.idr rename to tests/idris2/misc/params001/parambad.idr diff --git a/tests/idris2/params001/run b/tests/idris2/misc/params001/run similarity index 60% rename from tests/idris2/params001/run rename to tests/idris2/misc/params001/run index 15a9a0a8db..58bbe0bd7a 100755 --- a/tests/idris2/params001/run +++ b/tests/idris2/misc/params001/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh check param.idr check parambad.idr diff --git a/tests/idris2/params002/ParamsPrint.idr b/tests/idris2/misc/params002/ParamsPrint.idr similarity index 100% rename from tests/idris2/params002/ParamsPrint.idr rename to tests/idris2/misc/params002/ParamsPrint.idr diff --git a/tests/idris2/params002/expected b/tests/idris2/misc/params002/expected similarity index 100% rename from tests/idris2/params002/expected rename to tests/idris2/misc/params002/expected diff --git a/tests/idris2/params002/input b/tests/idris2/misc/params002/input similarity index 100% rename from tests/idris2/params002/input rename to tests/idris2/misc/params002/input diff --git a/tests/idris2/params002/run b/tests/idris2/misc/params002/run similarity index 57% rename from tests/idris2/params002/run rename to tests/idris2/misc/params002/run index 1984e14548..c740a56f56 100755 --- a/tests/idris2/params002/run +++ b/tests/idris2/misc/params002/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 ParamsPrint.idr < input diff --git a/tests/idris2/params003/casesplit.idr b/tests/idris2/misc/params003/casesplit.idr similarity index 100% rename from tests/idris2/params003/casesplit.idr rename to tests/idris2/misc/params003/casesplit.idr diff --git a/tests/idris2/params003/expected b/tests/idris2/misc/params003/expected similarity index 100% rename from tests/idris2/params003/expected rename to tests/idris2/misc/params003/expected diff --git a/tests/idris2/params003/input b/tests/idris2/misc/params003/input similarity index 100% rename from tests/idris2/params003/input rename to tests/idris2/misc/params003/input diff --git a/tests/idris2/params003/run b/tests/idris2/misc/params003/run similarity index 55% rename from tests/idris2/params003/run rename to tests/idris2/misc/params003/run index 8d17d8e6ce..a306665384 100755 --- a/tests/idris2/params003/run +++ b/tests/idris2/misc/params003/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 casesplit.idr < input diff --git a/tests/idris2/pretty001/Issue1328A.idr b/tests/idris2/misc/pretty001/Issue1328A.idr similarity index 100% rename from tests/idris2/pretty001/Issue1328A.idr rename to tests/idris2/misc/pretty001/Issue1328A.idr diff --git a/tests/idris2/pretty001/expected b/tests/idris2/misc/pretty001/expected similarity index 100% rename from tests/idris2/pretty001/expected rename to tests/idris2/misc/pretty001/expected diff --git a/tests/idris2/pretty001/input b/tests/idris2/misc/pretty001/input similarity index 100% rename from tests/idris2/pretty001/input rename to tests/idris2/misc/pretty001/input diff --git a/tests/idris2/pretty001/run b/tests/idris2/misc/pretty001/run similarity index 56% rename from tests/idris2/pretty001/run rename to tests/idris2/misc/pretty001/run index 7b69792ef2..87e5eb6623 100644 --- a/tests/idris2/pretty001/run +++ b/tests/idris2/misc/pretty001/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 Issue1328A.idr < input diff --git a/tests/idris2/pretty002/expected b/tests/idris2/misc/pretty002/expected similarity index 100% rename from tests/idris2/pretty002/expected rename to tests/idris2/misc/pretty002/expected diff --git a/tests/idris2/pretty002/input b/tests/idris2/misc/pretty002/input similarity index 100% rename from tests/idris2/pretty002/input rename to tests/idris2/misc/pretty002/input diff --git a/tests/idris2/misc/pretty002/run b/tests/idris2/misc/pretty002/run new file mode 100644 index 0000000000..404d7d47c6 --- /dev/null +++ b/tests/idris2/misc/pretty002/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +idris2 < input diff --git a/tests/idris2/primloop/PrimLoop.idr b/tests/idris2/misc/primloop/PrimLoop.idr similarity index 100% rename from tests/idris2/primloop/PrimLoop.idr rename to tests/idris2/misc/primloop/PrimLoop.idr diff --git a/tests/idris2/primloop/expected b/tests/idris2/misc/primloop/expected similarity index 100% rename from tests/idris2/primloop/expected rename to tests/idris2/misc/primloop/expected diff --git a/tests/idris2/misc/primloop/run b/tests/idris2/misc/primloop/run new file mode 100644 index 0000000000..00360d00f4 --- /dev/null +++ b/tests/idris2/misc/primloop/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +run PrimLoop.idr diff --git a/tests/idris2/quantifiers001/TestQuantifiers.idr b/tests/idris2/misc/quantifiers001/TestQuantifiers.idr similarity index 100% rename from tests/idris2/quantifiers001/TestQuantifiers.idr rename to tests/idris2/misc/quantifiers001/TestQuantifiers.idr diff --git a/tests/idris2/quantifiers001/expected b/tests/idris2/misc/quantifiers001/expected similarity index 100% rename from tests/idris2/quantifiers001/expected rename to tests/idris2/misc/quantifiers001/expected diff --git a/tests/idris2/quantifiers001/input b/tests/idris2/misc/quantifiers001/input similarity index 100% rename from tests/idris2/quantifiers001/input rename to tests/idris2/misc/quantifiers001/input diff --git a/tests/idris2/quantifiers001/run b/tests/idris2/misc/quantifiers001/run similarity index 60% rename from tests/idris2/quantifiers001/run rename to tests/idris2/misc/quantifiers001/run index 5ea4e26bd1..fd9be5d12a 100755 --- a/tests/idris2/quantifiers001/run +++ b/tests/idris2/misc/quantifiers001/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 TestQuantifiers.idr < input diff --git a/tests/idris2/real001/Channel.idr b/tests/idris2/misc/real001/Channel.idr similarity index 100% rename from tests/idris2/real001/Channel.idr rename to tests/idris2/misc/real001/Channel.idr diff --git a/tests/idris2/real001/Linear.idr b/tests/idris2/misc/real001/Linear.idr similarity index 100% rename from tests/idris2/real001/Linear.idr rename to tests/idris2/misc/real001/Linear.idr diff --git a/tests/idris2/real001/MakeChans.idr b/tests/idris2/misc/real001/MakeChans.idr similarity index 100% rename from tests/idris2/real001/MakeChans.idr rename to tests/idris2/misc/real001/MakeChans.idr diff --git a/tests/idris2/real001/TestProto.idr b/tests/idris2/misc/real001/TestProto.idr similarity index 100% rename from tests/idris2/real001/TestProto.idr rename to tests/idris2/misc/real001/TestProto.idr diff --git a/tests/idris2/real001/expected b/tests/idris2/misc/real001/expected similarity index 100% rename from tests/idris2/real001/expected rename to tests/idris2/misc/real001/expected diff --git a/tests/idris2/real001/run b/tests/idris2/misc/real001/run similarity index 63% rename from tests/idris2/real001/run rename to tests/idris2/misc/real001/run index dbedfe8f05..40da6e87b9 100644 --- a/tests/idris2/real001/run +++ b/tests/idris2/misc/real001/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh check TestProto.idr #check MakeChans.idr diff --git a/tests/idris2/real002/Store.idr b/tests/idris2/misc/real002/Store.idr similarity index 100% rename from tests/idris2/real002/Store.idr rename to tests/idris2/misc/real002/Store.idr diff --git a/tests/idris2/real002/StoreL.idr b/tests/idris2/misc/real002/StoreL.idr similarity index 100% rename from tests/idris2/real002/StoreL.idr rename to tests/idris2/misc/real002/StoreL.idr diff --git a/tests/idris2/real002/expected b/tests/idris2/misc/real002/expected similarity index 100% rename from tests/idris2/real002/expected rename to tests/idris2/misc/real002/expected diff --git a/tests/idris2/real002/run b/tests/idris2/misc/real002/run similarity index 58% rename from tests/idris2/real002/run rename to tests/idris2/misc/real002/run index 3bb3bdd4b3..50b0a41f8c 100644 --- a/tests/idris2/real002/run +++ b/tests/idris2/misc/real002/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh check Store.idr check StoreL.idr diff --git a/tests/idris2/unification001/Issue647.idr b/tests/idris2/misc/unification001/Issue647.idr similarity index 100% rename from tests/idris2/unification001/Issue647.idr rename to tests/idris2/misc/unification001/Issue647.idr diff --git a/tests/idris2/unification001/expected b/tests/idris2/misc/unification001/expected similarity index 100% rename from tests/idris2/unification001/expected rename to tests/idris2/misc/unification001/expected diff --git a/tests/idris2/misc/unification001/run b/tests/idris2/misc/unification001/run new file mode 100755 index 0000000000..b5bc5b4ebe --- /dev/null +++ b/tests/idris2/misc/unification001/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Issue647.idr diff --git a/tests/idris2/with003/Main.idr b/tests/idris2/misc/with003/Main.idr similarity index 100% rename from tests/idris2/with003/Main.idr rename to tests/idris2/misc/with003/Main.idr diff --git a/tests/idris2/with003/expected b/tests/idris2/misc/with003/expected similarity index 100% rename from tests/idris2/with003/expected rename to tests/idris2/misc/with003/expected diff --git a/tests/idris2/with003/input b/tests/idris2/misc/with003/input similarity index 100% rename from tests/idris2/with003/input rename to tests/idris2/misc/with003/input diff --git a/tests/idris2/with003/run b/tests/idris2/misc/with003/run similarity index 51% rename from tests/idris2/with003/run rename to tests/idris2/misc/with003/run index 37beba4789..7935482b8c 100755 --- a/tests/idris2/with003/run +++ b/tests/idris2/misc/with003/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 Main.idr < input diff --git a/tests/idris2/namespace004/run b/tests/idris2/namespace004/run deleted file mode 100644 index 3e447e08b5..0000000000 --- a/tests/idris2/namespace004/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check Export.idr diff --git a/tests/idris2/perf001/Big.idr b/tests/idris2/perf/perf001/Big.idr similarity index 100% rename from tests/idris2/perf001/Big.idr rename to tests/idris2/perf/perf001/Big.idr diff --git a/tests/idris2/perf001/expected b/tests/idris2/perf/perf001/expected similarity index 100% rename from tests/idris2/perf001/expected rename to tests/idris2/perf/perf001/expected diff --git a/tests/idris2/perf/perf001/run b/tests/idris2/perf/perf001/run new file mode 100755 index 0000000000..644d9e38d4 --- /dev/null +++ b/tests/idris2/perf/perf001/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Big.idr diff --git a/tests/idris2/perf002/Big.idr b/tests/idris2/perf/perf002/Big.idr similarity index 100% rename from tests/idris2/perf002/Big.idr rename to tests/idris2/perf/perf002/Big.idr diff --git a/tests/idris2/perf002/expected b/tests/idris2/perf/perf002/expected similarity index 100% rename from tests/idris2/perf002/expected rename to tests/idris2/perf/perf002/expected diff --git a/tests/idris2/perf/perf002/run b/tests/idris2/perf/perf002/run new file mode 100755 index 0000000000..644d9e38d4 --- /dev/null +++ b/tests/idris2/perf/perf002/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Big.idr diff --git a/tests/idris2/perf003/Auto.idr b/tests/idris2/perf/perf003/Auto.idr similarity index 100% rename from tests/idris2/perf003/Auto.idr rename to tests/idris2/perf/perf003/Auto.idr diff --git a/tests/idris2/perf003/expected b/tests/idris2/perf/perf003/expected similarity index 100% rename from tests/idris2/perf003/expected rename to tests/idris2/perf/perf003/expected diff --git a/tests/idris2/perf/perf003/run b/tests/idris2/perf/perf003/run new file mode 100755 index 0000000000..6030d7c1f5 --- /dev/null +++ b/tests/idris2/perf/perf003/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Auto.idr diff --git a/tests/idris2/perf004/bigdpair.idr b/tests/idris2/perf/perf004/bigdpair.idr similarity index 100% rename from tests/idris2/perf004/bigdpair.idr rename to tests/idris2/perf/perf004/bigdpair.idr diff --git a/tests/idris2/perf004/expected b/tests/idris2/perf/perf004/expected similarity index 100% rename from tests/idris2/perf004/expected rename to tests/idris2/perf/perf004/expected diff --git a/tests/idris2/perf/perf004/run b/tests/idris2/perf/perf004/run new file mode 100755 index 0000000000..27641ed8c4 --- /dev/null +++ b/tests/idris2/perf/perf004/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check bigdpair.idr diff --git a/tests/idris2/perf005/Bad1.idr b/tests/idris2/perf/perf005/Bad1.idr similarity index 100% rename from tests/idris2/perf005/Bad1.idr rename to tests/idris2/perf/perf005/Bad1.idr diff --git a/tests/idris2/perf005/Bad2.idr b/tests/idris2/perf/perf005/Bad2.idr similarity index 100% rename from tests/idris2/perf005/Bad2.idr rename to tests/idris2/perf/perf005/Bad2.idr diff --git a/tests/idris2/perf005/Bad3.idr b/tests/idris2/perf/perf005/Bad3.idr similarity index 100% rename from tests/idris2/perf005/Bad3.idr rename to tests/idris2/perf/perf005/Bad3.idr diff --git a/tests/idris2/perf005/Lambda.idr b/tests/idris2/perf/perf005/Lambda.idr similarity index 100% rename from tests/idris2/perf005/Lambda.idr rename to tests/idris2/perf/perf005/Lambda.idr diff --git a/tests/idris2/perf005/NoRegression.idr b/tests/idris2/perf/perf005/NoRegression.idr similarity index 100% rename from tests/idris2/perf005/NoRegression.idr rename to tests/idris2/perf/perf005/NoRegression.idr diff --git a/tests/idris2/perf005/expected b/tests/idris2/perf/perf005/expected similarity index 100% rename from tests/idris2/perf005/expected rename to tests/idris2/perf/perf005/expected diff --git a/tests/idris2/perf005/run b/tests/idris2/perf/perf005/run similarity index 73% rename from tests/idris2/perf005/run rename to tests/idris2/perf/perf005/run index 1bc7122df1..2dd852811c 100755 --- a/tests/idris2/perf005/run +++ b/tests/idris2/perf/perf005/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh check Lambda.idr diff --git a/tests/idris2/perf007/Slooow.idr b/tests/idris2/perf/perf007/Slooow.idr similarity index 100% rename from tests/idris2/perf007/Slooow.idr rename to tests/idris2/perf/perf007/Slooow.idr diff --git a/tests/idris2/perf007/expected b/tests/idris2/perf/perf007/expected similarity index 100% rename from tests/idris2/perf007/expected rename to tests/idris2/perf/perf007/expected diff --git a/tests/idris2/perf/perf007/run b/tests/idris2/perf/perf007/run new file mode 100644 index 0000000000..59a9a11b56 --- /dev/null +++ b/tests/idris2/perf/perf007/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Slooow.idr diff --git a/tests/idris2/perf008/FinPerf.idr b/tests/idris2/perf/perf008/FinPerf.idr similarity index 100% rename from tests/idris2/perf008/FinPerf.idr rename to tests/idris2/perf/perf008/FinPerf.idr diff --git a/tests/idris2/perf008/expected b/tests/idris2/perf/perf008/expected similarity index 100% rename from tests/idris2/perf008/expected rename to tests/idris2/perf/perf008/expected diff --git a/tests/idris2/perf/perf008/run b/tests/idris2/perf/perf008/run new file mode 100644 index 0000000000..f5d0a36aad --- /dev/null +++ b/tests/idris2/perf/perf008/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +run FinPerf.idr diff --git a/tests/idris2/perf009/A.idr b/tests/idris2/perf/perf009/A.idr similarity index 100% rename from tests/idris2/perf009/A.idr rename to tests/idris2/perf/perf009/A.idr diff --git a/tests/idris2/perf009/B.idr b/tests/idris2/perf/perf009/B.idr similarity index 100% rename from tests/idris2/perf009/B.idr rename to tests/idris2/perf/perf009/B.idr diff --git a/tests/idris2/perf009/C.idr b/tests/idris2/perf/perf009/C.idr similarity index 100% rename from tests/idris2/perf009/C.idr rename to tests/idris2/perf/perf009/C.idr diff --git a/tests/idris2/perf009/expected b/tests/idris2/perf/perf009/expected similarity index 100% rename from tests/idris2/perf009/expected rename to tests/idris2/perf/perf009/expected diff --git a/tests/idris2/perf/perf009/run b/tests/idris2/perf/perf009/run new file mode 100644 index 0000000000..30b175cd4b --- /dev/null +++ b/tests/idris2/perf/perf009/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check C.idr diff --git a/tests/idris2/perf010/Printf.idr b/tests/idris2/perf/perf010/Printf.idr similarity index 100% rename from tests/idris2/perf010/Printf.idr rename to tests/idris2/perf/perf010/Printf.idr diff --git a/tests/idris2/perf010/expected b/tests/idris2/perf/perf010/expected similarity index 100% rename from tests/idris2/perf010/expected rename to tests/idris2/perf/perf010/expected diff --git a/tests/idris2/perf/perf010/run b/tests/idris2/perf/perf010/run new file mode 100644 index 0000000000..3fb43e19ca --- /dev/null +++ b/tests/idris2/perf/perf010/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Printf.idr diff --git a/tests/idris2/perf011/A.idr b/tests/idris2/perf/perf011/A.idr similarity index 100% rename from tests/idris2/perf011/A.idr rename to tests/idris2/perf/perf011/A.idr diff --git a/tests/idris2/perf011/B.idr b/tests/idris2/perf/perf011/B.idr similarity index 100% rename from tests/idris2/perf011/B.idr rename to tests/idris2/perf/perf011/B.idr diff --git a/tests/idris2/perf011/C.idr b/tests/idris2/perf/perf011/C.idr similarity index 100% rename from tests/idris2/perf011/C.idr rename to tests/idris2/perf/perf011/C.idr diff --git a/tests/idris2/perf011/expected b/tests/idris2/perf/perf011/expected similarity index 100% rename from tests/idris2/perf011/expected rename to tests/idris2/perf/perf011/expected diff --git a/tests/idris2/perf/perf011/run b/tests/idris2/perf/perf011/run new file mode 100644 index 0000000000..30b175cd4b --- /dev/null +++ b/tests/idris2/perf/perf011/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check C.idr diff --git a/tests/idris2/perf012/Main.idr b/tests/idris2/perf/perf012/Main.idr similarity index 100% rename from tests/idris2/perf012/Main.idr rename to tests/idris2/perf/perf012/Main.idr diff --git a/tests/idris2/perf012/expected b/tests/idris2/perf/perf012/expected similarity index 100% rename from tests/idris2/perf012/expected rename to tests/idris2/perf/perf012/expected diff --git a/tests/idris2/perf012/run b/tests/idris2/perf/perf012/run similarity index 51% rename from tests/idris2/perf012/run rename to tests/idris2/perf/perf012/run index 6fb7ad593f..39564dccfb 100644 --- a/tests/idris2/perf012/run +++ b/tests/idris2/perf/perf012/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 Main.idr -o main diff --git a/tests/idris2/perf2202/Church.idr b/tests/idris2/perf/perf2202/Church.idr similarity index 100% rename from tests/idris2/perf2202/Church.idr rename to tests/idris2/perf/perf2202/Church.idr diff --git a/tests/idris2/perf2202/expected b/tests/idris2/perf/perf2202/expected similarity index 100% rename from tests/idris2/perf2202/expected rename to tests/idris2/perf/perf2202/expected diff --git a/tests/idris2/perf2202/many10000.idr b/tests/idris2/perf/perf2202/many10000.idr similarity index 100% rename from tests/idris2/perf2202/many10000.idr rename to tests/idris2/perf/perf2202/many10000.idr diff --git a/tests/idris2/perf2202/run b/tests/idris2/perf/perf2202/run similarity index 61% rename from tests/idris2/perf2202/run rename to tests/idris2/perf/perf2202/run index f0e81801b6..0b987998a8 100755 --- a/tests/idris2/perf2202/run +++ b/tests/idris2/perf/perf2202/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh check many10000.idr check Church.idr diff --git a/tests/idris2/perf001/run b/tests/idris2/perf001/run deleted file mode 100755 index f9651336dd..0000000000 --- a/tests/idris2/perf001/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check Big.idr diff --git a/tests/idris2/perf002/run b/tests/idris2/perf002/run deleted file mode 100755 index f9651336dd..0000000000 --- a/tests/idris2/perf002/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check Big.idr diff --git a/tests/idris2/perf003/run b/tests/idris2/perf003/run deleted file mode 100755 index 2398c31034..0000000000 --- a/tests/idris2/perf003/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check Auto.idr diff --git a/tests/idris2/perf004/run b/tests/idris2/perf004/run deleted file mode 100755 index 2af4edcc4e..0000000000 --- a/tests/idris2/perf004/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check bigdpair.idr diff --git a/tests/idris2/perf007/run b/tests/idris2/perf007/run deleted file mode 100644 index 78ac337031..0000000000 --- a/tests/idris2/perf007/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check Slooow.idr diff --git a/tests/idris2/perf008/run b/tests/idris2/perf008/run deleted file mode 100644 index 3f89ed77ae..0000000000 --- a/tests/idris2/perf008/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -run FinPerf.idr diff --git a/tests/idris2/perf009/run b/tests/idris2/perf009/run deleted file mode 100644 index 4a6e58ce7f..0000000000 --- a/tests/idris2/perf009/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check C.idr diff --git a/tests/idris2/perf010/run b/tests/idris2/perf010/run deleted file mode 100644 index 8d7dc963a0..0000000000 --- a/tests/idris2/perf010/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check Printf.idr diff --git a/tests/idris2/perf011/run b/tests/idris2/perf011/run deleted file mode 100644 index 4a6e58ce7f..0000000000 --- a/tests/idris2/perf011/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check C.idr diff --git a/tests/idris2/perror001/run b/tests/idris2/perror001/run deleted file mode 100755 index 9af9ec0f8e..0000000000 --- a/tests/idris2/perror001/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check PError.idr diff --git a/tests/idris2/perror002/run b/tests/idris2/perror002/run deleted file mode 100755 index 9af9ec0f8e..0000000000 --- a/tests/idris2/perror002/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check PError.idr diff --git a/tests/idris2/perror004/run b/tests/idris2/perror004/run deleted file mode 100755 index 9af9ec0f8e..0000000000 --- a/tests/idris2/perror004/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check PError.idr diff --git a/tests/idris2/perror005/run b/tests/idris2/perror005/run deleted file mode 100755 index 9af9ec0f8e..0000000000 --- a/tests/idris2/perror005/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check PError.idr diff --git a/tests/idris2/perror006/run b/tests/idris2/perror006/run deleted file mode 100755 index 9af9ec0f8e..0000000000 --- a/tests/idris2/perror006/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check PError.idr diff --git a/tests/idris2/perror019/run b/tests/idris2/perror019/run deleted file mode 100644 index e91b494540..0000000000 --- a/tests/idris2/perror019/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check ImplError.idr diff --git a/tests/idris2/perror021/run b/tests/idris2/perror021/run deleted file mode 100644 index 096de90fd3..0000000000 --- a/tests/idris2/perror021/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check Implicit.idr diff --git a/tests/idris2/perror022/run b/tests/idris2/perror022/run deleted file mode 100644 index 2e9711032e..0000000000 --- a/tests/idris2/perror022/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check Indent.idr diff --git a/tests/idris2/perror023/run b/tests/idris2/perror023/run deleted file mode 100755 index a5fb5a99c7..0000000000 --- a/tests/idris2/perror023/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check ParseError.idr diff --git a/tests/idris2/perror024/run b/tests/idris2/perror024/run deleted file mode 100755 index a5fb5a99c7..0000000000 --- a/tests/idris2/perror024/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check ParseError.idr diff --git a/tests/idris2/perror025/run b/tests/idris2/perror025/run deleted file mode 100755 index 6c3df8ceed..0000000000 --- a/tests/idris2/perror025/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check DataWhere.idr diff --git a/tests/idris2/perror026/run b/tests/idris2/perror026/run deleted file mode 100755 index ea25311502..0000000000 --- a/tests/idris2/perror026/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check Micro.idr diff --git a/tests/idris2/perror027/run b/tests/idris2/perror027/run deleted file mode 100755 index 422d2c27c5..0000000000 --- a/tests/idris2/perror027/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check Outdent.idr diff --git a/tests/idris2/perror028/run b/tests/idris2/perror028/run deleted file mode 100644 index a9617307b4..0000000000 --- a/tests/idris2/perror028/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check LetInDo.idr diff --git a/tests/idris2/perror029/run b/tests/idris2/perror029/run deleted file mode 100644 index 36ad497cca..0000000000 --- a/tests/idris2/perror029/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check DelayParse.idr diff --git a/tests/idris2/pkg001/Dummy.idr b/tests/idris2/pkg/pkg001/Dummy.idr similarity index 100% rename from tests/idris2/pkg001/Dummy.idr rename to tests/idris2/pkg/pkg001/Dummy.idr diff --git a/tests/idris2/pkg001/dummy.ipkg b/tests/idris2/pkg/pkg001/dummy.ipkg similarity index 100% rename from tests/idris2/pkg001/dummy.ipkg rename to tests/idris2/pkg/pkg001/dummy.ipkg diff --git a/tests/idris2/pkg001/expected b/tests/idris2/pkg/pkg001/expected similarity index 100% rename from tests/idris2/pkg001/expected rename to tests/idris2/pkg/pkg001/expected diff --git a/tests/idris2/pkg001/run b/tests/idris2/pkg/pkg001/run similarity index 52% rename from tests/idris2/pkg001/run rename to tests/idris2/pkg/pkg001/run index 3d99b32b7c..1fed4c1c33 100755 --- a/tests/idris2/pkg001/run +++ b/tests/idris2/pkg/pkg001/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 --build dummy.ipkg diff --git a/tests/idris2/pkg002/dummy.ipkg b/tests/idris2/pkg/pkg002/dummy.ipkg similarity index 100% rename from tests/idris2/pkg002/dummy.ipkg rename to tests/idris2/pkg/pkg002/dummy.ipkg diff --git a/tests/idris2/pkg002/expected b/tests/idris2/pkg/pkg002/expected similarity index 100% rename from tests/idris2/pkg002/expected rename to tests/idris2/pkg/pkg002/expected diff --git a/tests/idris2/pkg002/run b/tests/idris2/pkg/pkg002/run similarity index 67% rename from tests/idris2/pkg002/run rename to tests/idris2/pkg/pkg002/run index 3cb8823af2..1be3a3b02b 100755 --- a/tests/idris2/pkg002/run +++ b/tests/idris2/pkg/pkg002/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh cd src/Top check --find-ipkg Dummy.idr diff --git a/tests/idris2/pkg002/src/Top/Dummy.idr b/tests/idris2/pkg/pkg002/src/Top/Dummy.idr similarity index 100% rename from tests/idris2/pkg002/src/Top/Dummy.idr rename to tests/idris2/pkg/pkg002/src/Top/Dummy.idr diff --git a/tests/idris2/pkg003/Main.idr b/tests/idris2/pkg/pkg003/Main.idr similarity index 100% rename from tests/idris2/pkg003/Main.idr rename to tests/idris2/pkg/pkg003/Main.idr diff --git a/tests/idris2/pkg003/expected b/tests/idris2/pkg/pkg003/expected similarity index 100% rename from tests/idris2/pkg003/expected rename to tests/idris2/pkg/pkg003/expected diff --git a/tests/idris2/pkg003/run b/tests/idris2/pkg/pkg003/run similarity index 92% rename from tests/idris2/pkg003/run rename to tests/idris2/pkg/pkg003/run index 61526aceb7..0acfc870c8 100755 --- a/tests/idris2/pkg003/run +++ b/tests/idris2/pkg/pkg003/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 --build testpkg.ipkg rm -rf build/ diff --git a/tests/idris2/pkg003/testpkg.ipkg b/tests/idris2/pkg/pkg003/testpkg.ipkg similarity index 100% rename from tests/idris2/pkg003/testpkg.ipkg rename to tests/idris2/pkg/pkg003/testpkg.ipkg diff --git a/tests/idris2/pkg004/Dummy.idr b/tests/idris2/pkg/pkg004/Dummy.idr similarity index 100% rename from tests/idris2/pkg004/Dummy.idr rename to tests/idris2/pkg/pkg004/Dummy.idr diff --git a/tests/idris2/pkg004/dummy.ipkg b/tests/idris2/pkg/pkg004/dummy.ipkg similarity index 100% rename from tests/idris2/pkg004/dummy.ipkg rename to tests/idris2/pkg/pkg004/dummy.ipkg diff --git a/tests/idris2/pkg004/expected b/tests/idris2/pkg/pkg004/expected similarity index 100% rename from tests/idris2/pkg004/expected rename to tests/idris2/pkg/pkg004/expected diff --git a/tests/idris2/pkg004/input b/tests/idris2/pkg/pkg004/input similarity index 100% rename from tests/idris2/pkg004/input rename to tests/idris2/pkg/pkg004/input diff --git a/tests/idris2/pkg004/run b/tests/idris2/pkg/pkg004/run similarity index 58% rename from tests/idris2/pkg004/run rename to tests/idris2/pkg/pkg004/run index b32a1a5bc6..14237f183c 100755 --- a/tests/idris2/pkg004/run +++ b/tests/idris2/pkg/pkg004/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 --repl dummy.ipkg < input diff --git a/tests/idris2/pkg005/Foo.idr b/tests/idris2/pkg/pkg005/Foo.idr similarity index 100% rename from tests/idris2/pkg005/Foo.idr rename to tests/idris2/pkg/pkg005/Foo.idr diff --git a/tests/idris2/pkg005/expected b/tests/idris2/pkg/pkg005/expected similarity index 100% rename from tests/idris2/pkg005/expected rename to tests/idris2/pkg/pkg005/expected diff --git a/tests/idris2/pkg005/foo.ipkg b/tests/idris2/pkg/pkg005/foo.ipkg similarity index 100% rename from tests/idris2/pkg005/foo.ipkg rename to tests/idris2/pkg/pkg005/foo.ipkg diff --git a/tests/idris2/pkg005/input b/tests/idris2/pkg/pkg005/input similarity index 100% rename from tests/idris2/pkg005/input rename to tests/idris2/pkg/pkg005/input diff --git a/tests/idris2/pkg005/run b/tests/idris2/pkg/pkg005/run similarity index 78% rename from tests/idris2/pkg005/run rename to tests/idris2/pkg/pkg005/run index 1e1df3c0b9..a6ba253f8a 100755 --- a/tests/idris2/pkg005/run +++ b/tests/idris2/pkg/pkg005/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 --clean foo.ipkg idris2 --repl foo.ipkg < input diff --git a/tests/idris2/pkg006/depends/bar-1.0.1/bar.ipkg b/tests/idris2/pkg/pkg006/depends/bar-1.0.1/bar.ipkg similarity index 100% rename from tests/idris2/pkg006/depends/bar-1.0.1/bar.ipkg rename to tests/idris2/pkg/pkg006/depends/bar-1.0.1/bar.ipkg diff --git a/tests/idris2/pkg006/depends/bar-baz/bar-baz.ipkg b/tests/idris2/pkg/pkg006/depends/bar-baz/bar-baz.ipkg similarity index 100% rename from tests/idris2/pkg006/depends/bar-baz/bar-baz.ipkg rename to tests/idris2/pkg/pkg006/depends/bar-baz/bar-baz.ipkg diff --git a/tests/idris2/pkg006/depends/foo-0.5/foo.ipkg b/tests/idris2/pkg/pkg006/depends/foo-0.5/foo.ipkg similarity index 100% rename from tests/idris2/pkg006/depends/foo-0.5/foo.ipkg rename to tests/idris2/pkg/pkg006/depends/foo-0.5/foo.ipkg diff --git a/tests/idris2/pkg006/depends/foo-bar-1.3.1/foo-bar.ipkg b/tests/idris2/pkg/pkg006/depends/foo-bar-1.3.1/foo-bar.ipkg similarity index 100% rename from tests/idris2/pkg006/depends/foo-bar-1.3.1/foo-bar.ipkg rename to tests/idris2/pkg/pkg006/depends/foo-bar-1.3.1/foo-bar.ipkg diff --git a/tests/idris2/pkg006/depends/quux/quux.ipkg b/tests/idris2/pkg/pkg006/depends/quux/quux.ipkg similarity index 100% rename from tests/idris2/pkg006/depends/quux/quux.ipkg rename to tests/idris2/pkg/pkg006/depends/quux/quux.ipkg diff --git a/tests/idris2/pkg006/expected b/tests/idris2/pkg/pkg006/expected similarity index 100% rename from tests/idris2/pkg006/expected rename to tests/idris2/pkg/pkg006/expected diff --git a/tests/idris2/pkg006/run b/tests/idris2/pkg/pkg006/run similarity index 84% rename from tests/idris2/pkg006/run rename to tests/idris2/pkg/pkg006/run index 3c62534d24..8a0e3a397d 100755 --- a/tests/idris2/pkg006/run +++ b/tests/idris2/pkg/pkg006/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 --build test1.ipkg idris2 --build test2.ipkg diff --git a/tests/idris2/pkg006/test1.ipkg b/tests/idris2/pkg/pkg006/test1.ipkg similarity index 100% rename from tests/idris2/pkg006/test1.ipkg rename to tests/idris2/pkg/pkg006/test1.ipkg diff --git a/tests/idris2/pkg006/test2.ipkg b/tests/idris2/pkg/pkg006/test2.ipkg similarity index 100% rename from tests/idris2/pkg006/test2.ipkg rename to tests/idris2/pkg/pkg006/test2.ipkg diff --git a/tests/idris2/pkg006/test3.ipkg b/tests/idris2/pkg/pkg006/test3.ipkg similarity index 100% rename from tests/idris2/pkg006/test3.ipkg rename to tests/idris2/pkg/pkg006/test3.ipkg diff --git a/tests/idris2/pkg006/test4.ipkg b/tests/idris2/pkg/pkg006/test4.ipkg similarity index 100% rename from tests/idris2/pkg006/test4.ipkg rename to tests/idris2/pkg/pkg006/test4.ipkg diff --git a/tests/idris2/pkg006/test5.ipkg b/tests/idris2/pkg/pkg006/test5.ipkg similarity index 100% rename from tests/idris2/pkg006/test5.ipkg rename to tests/idris2/pkg/pkg006/test5.ipkg diff --git a/tests/idris2/pkg007/A/Path/Of/Dires/First.idr b/tests/idris2/pkg/pkg007/A/Path/Of/Dires/First.idr similarity index 100% rename from tests/idris2/pkg007/A/Path/Of/Dires/First.idr rename to tests/idris2/pkg/pkg007/A/Path/Of/Dires/First.idr diff --git a/tests/idris2/pkg007/A/Path/Of/Dires/Second.idr b/tests/idris2/pkg/pkg007/A/Path/Of/Dires/Second.idr similarity index 100% rename from tests/idris2/pkg007/A/Path/Of/Dires/Second.idr rename to tests/idris2/pkg/pkg007/A/Path/Of/Dires/Second.idr diff --git a/tests/idris2/pkg007/Another/Fourth.idr b/tests/idris2/pkg/pkg007/Another/Fourth.idr similarity index 100% rename from tests/idris2/pkg007/Another/Fourth.idr rename to tests/idris2/pkg/pkg007/Another/Fourth.idr diff --git a/tests/idris2/pkg007/Another/One/Third.idr b/tests/idris2/pkg/pkg007/Another/One/Third.idr similarity index 100% rename from tests/idris2/pkg007/Another/One/Third.idr rename to tests/idris2/pkg/pkg007/Another/One/Third.idr diff --git a/tests/idris2/pkg007/expected b/tests/idris2/pkg/pkg007/expected similarity index 100% rename from tests/idris2/pkg007/expected rename to tests/idris2/pkg/pkg007/expected diff --git a/tests/idris2/pkg007/input b/tests/idris2/pkg/pkg007/input similarity index 100% rename from tests/idris2/pkg007/input rename to tests/idris2/pkg/pkg007/input diff --git a/tests/idris2/pkg007/input2 b/tests/idris2/pkg/pkg007/input2 similarity index 100% rename from tests/idris2/pkg007/input2 rename to tests/idris2/pkg/pkg007/input2 diff --git a/tests/idris2/pkg007/run b/tests/idris2/pkg/pkg007/run similarity index 87% rename from tests/idris2/pkg007/run rename to tests/idris2/pkg/pkg007/run index 834d0f0ca7..bdca77a5ae 100755 --- a/tests/idris2/pkg007/run +++ b/tests/idris2/pkg/pkg007/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 --init < input idris2 --build cool.ipkg diff --git a/tests/idris2/pkg007/src/And/A/Proof.idr b/tests/idris2/pkg/pkg007/src/And/A/Proof.idr similarity index 100% rename from tests/idris2/pkg007/src/And/A/Proof.idr rename to tests/idris2/pkg/pkg007/src/And/A/Proof.idr diff --git a/tests/idris2/pkg007/src/Yet/Another/Path.idr b/tests/idris2/pkg/pkg007/src/Yet/Another/Path.idr similarity index 100% rename from tests/idris2/pkg007/src/Yet/Another/Path.idr rename to tests/idris2/pkg/pkg007/src/Yet/Another/Path.idr diff --git a/tests/idris2/pkg008/Bar.idr b/tests/idris2/pkg/pkg008/Bar.idr similarity index 100% rename from tests/idris2/pkg008/Bar.idr rename to tests/idris2/pkg/pkg008/Bar.idr diff --git a/tests/idris2/pkg008/Foo.idr b/tests/idris2/pkg/pkg008/Foo.idr similarity index 100% rename from tests/idris2/pkg008/Foo.idr rename to tests/idris2/pkg/pkg008/Foo.idr diff --git a/tests/idris2/pkg008/bar.ipkg b/tests/idris2/pkg/pkg008/bar.ipkg similarity index 100% rename from tests/idris2/pkg008/bar.ipkg rename to tests/idris2/pkg/pkg008/bar.ipkg diff --git a/tests/idris2/pkg008/expected b/tests/idris2/pkg/pkg008/expected similarity index 100% rename from tests/idris2/pkg008/expected rename to tests/idris2/pkg/pkg008/expected diff --git a/tests/idris2/pkg008/foo.ipkg b/tests/idris2/pkg/pkg008/foo.ipkg similarity index 100% rename from tests/idris2/pkg008/foo.ipkg rename to tests/idris2/pkg/pkg008/foo.ipkg diff --git a/tests/idris2/pkg008/run b/tests/idris2/pkg/pkg008/run similarity index 63% rename from tests/idris2/pkg008/run rename to tests/idris2/pkg/pkg008/run index 652be410e4..53a7d6590c 100755 --- a/tests/idris2/pkg008/run +++ b/tests/idris2/pkg/pkg008/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 --build foo.ipkg --build bar.ipkg diff --git a/tests/idris2/pkg009/expected b/tests/idris2/pkg/pkg009/expected similarity index 100% rename from tests/idris2/pkg009/expected rename to tests/idris2/pkg/pkg009/expected diff --git a/tests/idris2/pkg009/run b/tests/idris2/pkg/pkg009/run similarity index 71% rename from tests/idris2/pkg009/run rename to tests/idris2/pkg/pkg009/run index 83fb782eeb..3663e195f7 100644 --- a/tests/idris2/pkg009/run +++ b/tests/idris2/pkg/pkg009/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 --build testpkg/testpkg.ipkg rm -rf testpkg/build/ diff --git a/tests/idris2/pkg009/testpkg/Main.idr b/tests/idris2/pkg/pkg009/testpkg/Main.idr similarity index 100% rename from tests/idris2/pkg009/testpkg/Main.idr rename to tests/idris2/pkg/pkg009/testpkg/Main.idr diff --git a/tests/idris2/pkg009/testpkg/testpkg.ipkg b/tests/idris2/pkg/pkg009/testpkg/testpkg.ipkg similarity index 100% rename from tests/idris2/pkg009/testpkg/testpkg.ipkg rename to tests/idris2/pkg/pkg009/testpkg/testpkg.ipkg diff --git a/tests/idris2/pkg010/Main.idr b/tests/idris2/pkg/pkg010/Main.idr similarity index 100% rename from tests/idris2/pkg010/Main.idr rename to tests/idris2/pkg/pkg010/Main.idr diff --git a/tests/idris2/pkg010/expected b/tests/idris2/pkg/pkg010/expected similarity index 100% rename from tests/idris2/pkg010/expected rename to tests/idris2/pkg/pkg010/expected diff --git a/tests/idris2/pkg010/run b/tests/idris2/pkg/pkg010/run similarity index 90% rename from tests/idris2/pkg010/run rename to tests/idris2/pkg/pkg010/run index 9277c11787..1135ebaaa6 100755 --- a/tests/idris2/pkg010/run +++ b/tests/idris2/pkg/pkg010/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh export IDRIS2_PACKAGE_PATH=$IDRIS2_PREFIX/$NAME_VERSION export IDRIS2_PREFIX=$test_dir/currently/nonexistent/dir/ diff --git a/tests/idris2/pkg010/testpkg.ipkg b/tests/idris2/pkg/pkg010/testpkg.ipkg similarity index 100% rename from tests/idris2/pkg010/testpkg.ipkg rename to tests/idris2/pkg/pkg010/testpkg.ipkg diff --git a/tests/idris2/pkg011/dot-slash-dot/Main.idr b/tests/idris2/pkg/pkg011/dot-slash-dot/Main.idr similarity index 100% rename from tests/idris2/pkg011/dot-slash-dot/Main.idr rename to tests/idris2/pkg/pkg011/dot-slash-dot/Main.idr diff --git a/tests/idris2/pkg011/dot-slash-dot/testpkg.ipkg b/tests/idris2/pkg/pkg011/dot-slash-dot/testpkg.ipkg similarity index 100% rename from tests/idris2/pkg011/dot-slash-dot/testpkg.ipkg rename to tests/idris2/pkg/pkg011/dot-slash-dot/testpkg.ipkg diff --git a/tests/idris2/pkg011/dot-slash-name-slash/src/Main.idr b/tests/idris2/pkg/pkg011/dot-slash-name-slash/src/Main.idr similarity index 100% rename from tests/idris2/pkg011/dot-slash-name-slash/src/Main.idr rename to tests/idris2/pkg/pkg011/dot-slash-name-slash/src/Main.idr diff --git a/tests/idris2/pkg011/dot-slash-name-slash/testpkg.ipkg b/tests/idris2/pkg/pkg011/dot-slash-name-slash/testpkg.ipkg similarity index 100% rename from tests/idris2/pkg011/dot-slash-name-slash/testpkg.ipkg rename to tests/idris2/pkg/pkg011/dot-slash-name-slash/testpkg.ipkg diff --git a/tests/idris2/pkg011/dot-slash/Main.idr b/tests/idris2/pkg/pkg011/dot-slash/Main.idr similarity index 100% rename from tests/idris2/pkg011/dot-slash/Main.idr rename to tests/idris2/pkg/pkg011/dot-slash/Main.idr diff --git a/tests/idris2/pkg011/dot-slash/testpkg.ipkg b/tests/idris2/pkg/pkg011/dot-slash/testpkg.ipkg similarity index 100% rename from tests/idris2/pkg011/dot-slash/testpkg.ipkg rename to tests/idris2/pkg/pkg011/dot-slash/testpkg.ipkg diff --git a/tests/idris2/pkg011/dot/Main.idr b/tests/idris2/pkg/pkg011/dot/Main.idr similarity index 100% rename from tests/idris2/pkg011/dot/Main.idr rename to tests/idris2/pkg/pkg011/dot/Main.idr diff --git a/tests/idris2/pkg011/dot/testpkg.ipkg b/tests/idris2/pkg/pkg011/dot/testpkg.ipkg similarity index 100% rename from tests/idris2/pkg011/dot/testpkg.ipkg rename to tests/idris2/pkg/pkg011/dot/testpkg.ipkg diff --git a/tests/idris2/pkg011/expected b/tests/idris2/pkg/pkg011/expected similarity index 100% rename from tests/idris2/pkg011/expected rename to tests/idris2/pkg/pkg011/expected diff --git a/tests/idris2/pkg011/run b/tests/idris2/pkg/pkg011/run similarity index 94% rename from tests/idris2/pkg011/run rename to tests/idris2/pkg/pkg011/run index 78172fc980..b3a50422dc 100755 --- a/tests/idris2/pkg011/run +++ b/tests/idris2/pkg/pkg011/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh for t in dot dot-slash dot-slash-dot; do cd $t diff --git a/tests/idris2/pkg011/sibling/pkg/testpkg.ipkg b/tests/idris2/pkg/pkg011/sibling/pkg/testpkg.ipkg similarity index 100% rename from tests/idris2/pkg011/sibling/pkg/testpkg.ipkg rename to tests/idris2/pkg/pkg011/sibling/pkg/testpkg.ipkg diff --git a/tests/idris2/pkg011/sibling/src/Main.idr b/tests/idris2/pkg/pkg011/sibling/src/Main.idr similarity index 100% rename from tests/idris2/pkg011/sibling/src/Main.idr rename to tests/idris2/pkg/pkg011/sibling/src/Main.idr diff --git a/tests/idris2/pkg012/Main.idr b/tests/idris2/pkg/pkg012/Main.idr similarity index 100% rename from tests/idris2/pkg012/Main.idr rename to tests/idris2/pkg/pkg012/Main.idr diff --git a/tests/idris2/pkg012/expected b/tests/idris2/pkg/pkg012/expected similarity index 100% rename from tests/idris2/pkg012/expected rename to tests/idris2/pkg/pkg012/expected diff --git a/tests/idris2/pkg012/run b/tests/idris2/pkg/pkg012/run similarity index 54% rename from tests/idris2/pkg012/run rename to tests/idris2/pkg/pkg012/run index f5469d88ab..dd0e092b66 100755 --- a/tests/idris2/pkg012/run +++ b/tests/idris2/pkg/pkg012/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 --build testpkg.ipkg diff --git a/tests/idris2/pkg012/testpkg.ipkg b/tests/idris2/pkg/pkg012/testpkg.ipkg similarity index 100% rename from tests/idris2/pkg012/testpkg.ipkg rename to tests/idris2/pkg/pkg012/testpkg.ipkg diff --git a/tests/idris2/pkg013/depends/bar-0.7.2/bar.ipkg b/tests/idris2/pkg/pkg013/depends/bar-0.7.2/bar.ipkg similarity index 100% rename from tests/idris2/pkg013/depends/bar-0.7.2/bar.ipkg rename to tests/idris2/pkg/pkg013/depends/bar-0.7.2/bar.ipkg diff --git a/tests/idris2/pkg013/depends/foo-0.1.0/foo.ipkg b/tests/idris2/pkg/pkg013/depends/foo-0.1.0/foo.ipkg similarity index 100% rename from tests/idris2/pkg013/depends/foo-0.1.0/foo.ipkg rename to tests/idris2/pkg/pkg013/depends/foo-0.1.0/foo.ipkg diff --git a/tests/idris2/pkg013/expected b/tests/idris2/pkg/pkg013/expected similarity index 100% rename from tests/idris2/pkg013/expected rename to tests/idris2/pkg/pkg013/expected diff --git a/tests/idris2/pkg013/run b/tests/idris2/pkg/pkg013/run similarity index 74% rename from tests/idris2/pkg013/run rename to tests/idris2/pkg/pkg013/run index 4a39ba369b..cf15acfc2a 100644 --- a/tests/idris2/pkg013/run +++ b/tests/idris2/pkg/pkg013/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 --build test.ipkg --log package.depends:10 | filter_test_dir diff --git a/tests/idris2/pkg013/test.ipkg b/tests/idris2/pkg/pkg013/test.ipkg similarity index 100% rename from tests/idris2/pkg013/test.ipkg rename to tests/idris2/pkg/pkg013/test.ipkg diff --git a/tests/idris2/pkg014/depends/bar-0.1.0/bar.ipkg b/tests/idris2/pkg/pkg014/depends/bar-0.1.0/bar.ipkg similarity index 100% rename from tests/idris2/pkg014/depends/bar-0.1.0/bar.ipkg rename to tests/idris2/pkg/pkg014/depends/bar-0.1.0/bar.ipkg diff --git a/tests/idris2/pkg014/depends/bar-0.1.1/bar.ipkg b/tests/idris2/pkg/pkg014/depends/bar-0.1.1/bar.ipkg similarity index 100% rename from tests/idris2/pkg014/depends/bar-0.1.1/bar.ipkg rename to tests/idris2/pkg/pkg014/depends/bar-0.1.1/bar.ipkg diff --git a/tests/idris2/pkg014/depends/baz-0.1.0/baz.ipkg b/tests/idris2/pkg/pkg014/depends/baz-0.1.0/baz.ipkg similarity index 100% rename from tests/idris2/pkg014/depends/baz-0.1.0/baz.ipkg rename to tests/idris2/pkg/pkg014/depends/baz-0.1.0/baz.ipkg diff --git a/tests/idris2/pkg014/depends/baz-0.2.0/baz.ipkg b/tests/idris2/pkg/pkg014/depends/baz-0.2.0/baz.ipkg similarity index 100% rename from tests/idris2/pkg014/depends/baz-0.2.0/baz.ipkg rename to tests/idris2/pkg/pkg014/depends/baz-0.2.0/baz.ipkg diff --git a/tests/idris2/pkg014/depends/baz-0.3.0/baz.ipkg b/tests/idris2/pkg/pkg014/depends/baz-0.3.0/baz.ipkg similarity index 100% rename from tests/idris2/pkg014/depends/baz-0.3.0/baz.ipkg rename to tests/idris2/pkg/pkg014/depends/baz-0.3.0/baz.ipkg diff --git a/tests/idris2/pkg014/depends/foo-0.1.0/foo.ipkg b/tests/idris2/pkg/pkg014/depends/foo-0.1.0/foo.ipkg similarity index 100% rename from tests/idris2/pkg014/depends/foo-0.1.0/foo.ipkg rename to tests/idris2/pkg/pkg014/depends/foo-0.1.0/foo.ipkg diff --git a/tests/idris2/pkg014/depends/foo-0.2.0/foo.ipkg b/tests/idris2/pkg/pkg014/depends/foo-0.2.0/foo.ipkg similarity index 100% rename from tests/idris2/pkg014/depends/foo-0.2.0/foo.ipkg rename to tests/idris2/pkg/pkg014/depends/foo-0.2.0/foo.ipkg diff --git a/tests/idris2/pkg014/depends/foo-0.3.0/foo.ipkg b/tests/idris2/pkg/pkg014/depends/foo-0.3.0/foo.ipkg similarity index 100% rename from tests/idris2/pkg014/depends/foo-0.3.0/foo.ipkg rename to tests/idris2/pkg/pkg014/depends/foo-0.3.0/foo.ipkg diff --git a/tests/idris2/pkg014/expected b/tests/idris2/pkg/pkg014/expected similarity index 100% rename from tests/idris2/pkg014/expected rename to tests/idris2/pkg/pkg014/expected diff --git a/tests/idris2/pkg014/run b/tests/idris2/pkg/pkg014/run similarity index 74% rename from tests/idris2/pkg014/run rename to tests/idris2/pkg/pkg014/run index 4a39ba369b..cf15acfc2a 100644 --- a/tests/idris2/pkg014/run +++ b/tests/idris2/pkg/pkg014/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 --build test.ipkg --log package.depends:10 | filter_test_dir diff --git a/tests/idris2/pkg014/test.ipkg b/tests/idris2/pkg/pkg014/test.ipkg similarity index 100% rename from tests/idris2/pkg014/test.ipkg rename to tests/idris2/pkg/pkg014/test.ipkg diff --git a/tests/idris2/pkg015/depends/bar-0.1.0/bar.ipkg b/tests/idris2/pkg/pkg015/depends/bar-0.1.0/bar.ipkg similarity index 100% rename from tests/idris2/pkg015/depends/bar-0.1.0/bar.ipkg rename to tests/idris2/pkg/pkg015/depends/bar-0.1.0/bar.ipkg diff --git a/tests/idris2/pkg015/depends/bar-0.1.1/bar.ipkg b/tests/idris2/pkg/pkg015/depends/bar-0.1.1/bar.ipkg similarity index 100% rename from tests/idris2/pkg015/depends/bar-0.1.1/bar.ipkg rename to tests/idris2/pkg/pkg015/depends/bar-0.1.1/bar.ipkg diff --git a/tests/idris2/pkg015/depends/baz-0.1.0/baz.ipkg b/tests/idris2/pkg/pkg015/depends/baz-0.1.0/baz.ipkg similarity index 100% rename from tests/idris2/pkg015/depends/baz-0.1.0/baz.ipkg rename to tests/idris2/pkg/pkg015/depends/baz-0.1.0/baz.ipkg diff --git a/tests/idris2/pkg015/depends/baz-0.2.0/baz.ipkg b/tests/idris2/pkg/pkg015/depends/baz-0.2.0/baz.ipkg similarity index 100% rename from tests/idris2/pkg015/depends/baz-0.2.0/baz.ipkg rename to tests/idris2/pkg/pkg015/depends/baz-0.2.0/baz.ipkg diff --git a/tests/idris2/pkg015/depends/baz-0.3.0/baz.ipkg b/tests/idris2/pkg/pkg015/depends/baz-0.3.0/baz.ipkg similarity index 100% rename from tests/idris2/pkg015/depends/baz-0.3.0/baz.ipkg rename to tests/idris2/pkg/pkg015/depends/baz-0.3.0/baz.ipkg diff --git a/tests/idris2/pkg015/depends/foo-0.1.0/foo.ipkg b/tests/idris2/pkg/pkg015/depends/foo-0.1.0/foo.ipkg similarity index 100% rename from tests/idris2/pkg015/depends/foo-0.1.0/foo.ipkg rename to tests/idris2/pkg/pkg015/depends/foo-0.1.0/foo.ipkg diff --git a/tests/idris2/pkg015/depends/foo-0.2.0/foo.ipkg b/tests/idris2/pkg/pkg015/depends/foo-0.2.0/foo.ipkg similarity index 100% rename from tests/idris2/pkg015/depends/foo-0.2.0/foo.ipkg rename to tests/idris2/pkg/pkg015/depends/foo-0.2.0/foo.ipkg diff --git a/tests/idris2/pkg015/depends/foo-0.3.0/foo.ipkg b/tests/idris2/pkg/pkg015/depends/foo-0.3.0/foo.ipkg similarity index 100% rename from tests/idris2/pkg015/depends/foo-0.3.0/foo.ipkg rename to tests/idris2/pkg/pkg015/depends/foo-0.3.0/foo.ipkg diff --git a/tests/idris2/pkg015/depends/prz-0.1.0/prz.ipkg b/tests/idris2/pkg/pkg015/depends/prz-0.1.0/prz.ipkg similarity index 100% rename from tests/idris2/pkg015/depends/prz-0.1.0/prz.ipkg rename to tests/idris2/pkg/pkg015/depends/prz-0.1.0/prz.ipkg diff --git a/tests/idris2/pkg015/expected b/tests/idris2/pkg/pkg015/expected similarity index 100% rename from tests/idris2/pkg015/expected rename to tests/idris2/pkg/pkg015/expected diff --git a/tests/idris2/pkg015/run b/tests/idris2/pkg/pkg015/run similarity index 100% rename from tests/idris2/pkg015/run rename to tests/idris2/pkg/pkg015/run diff --git a/tests/idris2/pkg015/test.ipkg b/tests/idris2/pkg/pkg015/test.ipkg similarity index 100% rename from tests/idris2/pkg015/test.ipkg rename to tests/idris2/pkg/pkg015/test.ipkg diff --git a/tests/idris2/pkg016/bar.ipkg b/tests/idris2/pkg/pkg016/bar.ipkg similarity index 100% rename from tests/idris2/pkg016/bar.ipkg rename to tests/idris2/pkg/pkg016/bar.ipkg diff --git a/tests/idris2/pkg016/bar/Bar.idr b/tests/idris2/pkg/pkg016/bar/Bar.idr similarity index 100% rename from tests/idris2/pkg016/bar/Bar.idr rename to tests/idris2/pkg/pkg016/bar/Bar.idr diff --git a/tests/idris2/pkg016/baz.ipkg b/tests/idris2/pkg/pkg016/baz.ipkg similarity index 100% rename from tests/idris2/pkg016/baz.ipkg rename to tests/idris2/pkg/pkg016/baz.ipkg diff --git a/tests/idris2/pkg016/baz/Baz.idr b/tests/idris2/pkg/pkg016/baz/Baz.idr similarity index 100% rename from tests/idris2/pkg016/baz/Baz.idr rename to tests/idris2/pkg/pkg016/baz/Baz.idr diff --git a/tests/idris2/pkg016/expected b/tests/idris2/pkg/pkg016/expected similarity index 100% rename from tests/idris2/pkg016/expected rename to tests/idris2/pkg/pkg016/expected diff --git a/tests/idris2/pkg016/foo.ipkg b/tests/idris2/pkg/pkg016/foo.ipkg similarity index 100% rename from tests/idris2/pkg016/foo.ipkg rename to tests/idris2/pkg/pkg016/foo.ipkg diff --git a/tests/idris2/pkg016/foo/Foo.idr b/tests/idris2/pkg/pkg016/foo/Foo.idr similarity index 100% rename from tests/idris2/pkg016/foo/Foo.idr rename to tests/idris2/pkg/pkg016/foo/Foo.idr diff --git a/tests/idris2/pkg016/run b/tests/idris2/pkg/pkg016/run similarity index 100% rename from tests/idris2/pkg016/run rename to tests/idris2/pkg/pkg016/run diff --git a/tests/idris2/pkg016/src/Test.idr b/tests/idris2/pkg/pkg016/src/Test.idr similarity index 100% rename from tests/idris2/pkg016/src/Test.idr rename to tests/idris2/pkg/pkg016/src/Test.idr diff --git a/tests/idris2/pkg016/test.ipkg b/tests/idris2/pkg/pkg016/test.ipkg similarity index 100% rename from tests/idris2/pkg016/test.ipkg rename to tests/idris2/pkg/pkg016/test.ipkg diff --git a/tests/idris2/pkg017/a1/A.idr b/tests/idris2/pkg/pkg017/a1/A.idr similarity index 100% rename from tests/idris2/pkg017/a1/A.idr rename to tests/idris2/pkg/pkg017/a1/A.idr diff --git a/tests/idris2/pkg017/a1/a1.ipkg b/tests/idris2/pkg/pkg017/a1/a1.ipkg similarity index 100% rename from tests/idris2/pkg017/a1/a1.ipkg rename to tests/idris2/pkg/pkg017/a1/a1.ipkg diff --git a/tests/idris2/pkg017/a2/A.idr b/tests/idris2/pkg/pkg017/a2/A.idr similarity index 100% rename from tests/idris2/pkg017/a2/A.idr rename to tests/idris2/pkg/pkg017/a2/A.idr diff --git a/tests/idris2/pkg017/a2/a2.ipkg b/tests/idris2/pkg/pkg017/a2/a2.ipkg similarity index 100% rename from tests/idris2/pkg017/a2/a2.ipkg rename to tests/idris2/pkg/pkg017/a2/a2.ipkg diff --git a/tests/idris2/pkg017/b1/b1.ipkg b/tests/idris2/pkg/pkg017/b1/b1.ipkg similarity index 100% rename from tests/idris2/pkg017/b1/b1.ipkg rename to tests/idris2/pkg/pkg017/b1/b1.ipkg diff --git a/tests/idris2/pkg017/b1/src/B1.idr b/tests/idris2/pkg/pkg017/b1/src/B1.idr similarity index 100% rename from tests/idris2/pkg017/b1/src/B1.idr rename to tests/idris2/pkg/pkg017/b1/src/B1.idr diff --git a/tests/idris2/pkg017/b2/b2.ipkg b/tests/idris2/pkg/pkg017/b2/b2.ipkg similarity index 100% rename from tests/idris2/pkg017/b2/b2.ipkg rename to tests/idris2/pkg/pkg017/b2/b2.ipkg diff --git a/tests/idris2/pkg017/b2/src/B2.idr b/tests/idris2/pkg/pkg017/b2/src/B2.idr similarity index 100% rename from tests/idris2/pkg017/b2/src/B2.idr rename to tests/idris2/pkg/pkg017/b2/src/B2.idr diff --git a/tests/idris2/pkg017/expected b/tests/idris2/pkg/pkg017/expected similarity index 100% rename from tests/idris2/pkg017/expected rename to tests/idris2/pkg/pkg017/expected diff --git a/tests/idris2/pkg017/input1 b/tests/idris2/pkg/pkg017/input1 similarity index 100% rename from tests/idris2/pkg017/input1 rename to tests/idris2/pkg/pkg017/input1 diff --git a/tests/idris2/pkg017/input2 b/tests/idris2/pkg/pkg017/input2 similarity index 100% rename from tests/idris2/pkg017/input2 rename to tests/idris2/pkg/pkg017/input2 diff --git a/tests/idris2/pkg017/run b/tests/idris2/pkg/pkg017/run similarity index 95% rename from tests/idris2/pkg017/run rename to tests/idris2/pkg/pkg017/run index b99a65f8ad..f53a450eba 100644 --- a/tests/idris2/pkg017/run +++ b/tests/idris2/pkg/pkg017/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh mkdir prefix diff --git a/tests/idris2/positivity001/run b/tests/idris2/positivity001/run deleted file mode 100644 index 96ab89ccf9..0000000000 --- a/tests/idris2/positivity001/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check Issue660.idr diff --git a/tests/idris2/positivity002/run b/tests/idris2/positivity002/run deleted file mode 100644 index 96ab89ccf9..0000000000 --- a/tests/idris2/positivity002/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check Issue660.idr diff --git a/tests/idris2/positivity003/run b/tests/idris2/positivity003/run deleted file mode 100644 index 96ab89ccf9..0000000000 --- a/tests/idris2/positivity003/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check Issue660.idr diff --git a/tests/idris2/pretty002/run b/tests/idris2/pretty002/run deleted file mode 100644 index fecfb4b107..0000000000 --- a/tests/idris2/pretty002/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -idris2 < input diff --git a/tests/idris2/primloop/run b/tests/idris2/primloop/run deleted file mode 100644 index 63df8b24f7..0000000000 --- a/tests/idris2/primloop/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -run PrimLoop.idr diff --git a/tests/idris2/record003/run b/tests/idris2/record003/run deleted file mode 100755 index 80e434e591..0000000000 --- a/tests/idris2/record003/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -idris2 -c Record.idr diff --git a/tests/idris2/record010/run b/tests/idris2/record010/run deleted file mode 100755 index df93277329..0000000000 --- a/tests/idris2/record010/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check record.idr diff --git a/tests/idris2/record011/run b/tests/idris2/record011/run deleted file mode 100755 index 19ee100b0b..0000000000 --- a/tests/idris2/record011/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check Issue2095.idr diff --git a/tests/idris2/record012/run b/tests/idris2/record012/run deleted file mode 100755 index c0ea24df05..0000000000 --- a/tests/idris2/record012/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check Issue2065.idr diff --git a/tests/idris2/record014/run b/tests/idris2/record014/run deleted file mode 100755 index a5c8d5d776..0000000000 --- a/tests/idris2/record014/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check Issue1404.idr diff --git a/tests/idris2/record015/run b/tests/idris2/record015/run deleted file mode 100755 index de9479f270..0000000000 --- a/tests/idris2/record015/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -run Issue2176.idr diff --git a/tests/idris2/record018/run b/tests/idris2/record018/run deleted file mode 100755 index ce2877d08b..0000000000 --- a/tests/idris2/record018/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check mut.idr diff --git a/tests/idris2/reflection001/expected b/tests/idris2/reflection/reflection001/expected similarity index 100% rename from tests/idris2/reflection001/expected rename to tests/idris2/reflection/reflection001/expected diff --git a/tests/idris2/reflection001/input b/tests/idris2/reflection/reflection001/input similarity index 100% rename from tests/idris2/reflection001/input rename to tests/idris2/reflection/reflection001/input diff --git a/tests/idris2/reflection001/quote.idr b/tests/idris2/reflection/reflection001/quote.idr similarity index 100% rename from tests/idris2/reflection001/quote.idr rename to tests/idris2/reflection/reflection001/quote.idr diff --git a/tests/idris2/reflection001/run b/tests/idris2/reflection/reflection001/run similarity index 52% rename from tests/idris2/reflection001/run rename to tests/idris2/reflection/reflection001/run index 59012e6d4f..3a94b09377 100755 --- a/tests/idris2/reflection001/run +++ b/tests/idris2/reflection/reflection001/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 quote.idr < input diff --git a/tests/idris2/reflection002/expected b/tests/idris2/reflection/reflection002/expected similarity index 100% rename from tests/idris2/reflection002/expected rename to tests/idris2/reflection/reflection002/expected diff --git a/tests/idris2/reflection002/input b/tests/idris2/reflection/reflection002/input similarity index 100% rename from tests/idris2/reflection002/input rename to tests/idris2/reflection/reflection002/input diff --git a/tests/idris2/reflection002/power.idr b/tests/idris2/reflection/reflection002/power.idr similarity index 100% rename from tests/idris2/reflection002/power.idr rename to tests/idris2/reflection/reflection002/power.idr diff --git a/tests/idris2/reflection002/run b/tests/idris2/reflection/reflection002/run similarity index 52% rename from tests/idris2/reflection002/run rename to tests/idris2/reflection/reflection002/run index 8b5520427c..5647c63a30 100755 --- a/tests/idris2/reflection002/run +++ b/tests/idris2/reflection/reflection002/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 power.idr < input diff --git a/tests/idris2/reflection003/expected b/tests/idris2/reflection/reflection003/expected similarity index 100% rename from tests/idris2/reflection003/expected rename to tests/idris2/reflection/reflection003/expected diff --git a/tests/idris2/reflection003/refprims.idr b/tests/idris2/reflection/reflection003/refprims.idr similarity index 100% rename from tests/idris2/reflection003/refprims.idr rename to tests/idris2/reflection/reflection003/refprims.idr diff --git a/tests/idris2/reflection/reflection003/run b/tests/idris2/reflection/reflection003/run new file mode 100755 index 0000000000..85543664ae --- /dev/null +++ b/tests/idris2/reflection/reflection003/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check refprims.idr diff --git a/tests/idris2/reflection004/expected b/tests/idris2/reflection/reflection004/expected similarity index 100% rename from tests/idris2/reflection004/expected rename to tests/idris2/reflection/reflection004/expected diff --git a/tests/idris2/reflection004/input b/tests/idris2/reflection/reflection004/input similarity index 100% rename from tests/idris2/reflection004/input rename to tests/idris2/reflection/reflection004/input diff --git a/tests/idris2/reflection004/refdecl.idr b/tests/idris2/reflection/reflection004/refdecl.idr similarity index 100% rename from tests/idris2/reflection004/refdecl.idr rename to tests/idris2/reflection/reflection004/refdecl.idr diff --git a/tests/idris2/reflection004/run b/tests/idris2/reflection/reflection004/run similarity index 53% rename from tests/idris2/reflection004/run rename to tests/idris2/reflection/reflection004/run index e86f7b9b4f..11fac262bb 100755 --- a/tests/idris2/reflection004/run +++ b/tests/idris2/reflection/reflection004/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 refdecl.idr < input diff --git a/tests/idris2/reflection005/expected b/tests/idris2/reflection/reflection005/expected similarity index 100% rename from tests/idris2/reflection005/expected rename to tests/idris2/reflection/reflection005/expected diff --git a/tests/idris2/reflection005/input b/tests/idris2/reflection/reflection005/input similarity index 100% rename from tests/idris2/reflection005/input rename to tests/idris2/reflection/reflection005/input diff --git a/tests/idris2/reflection005/refdecl.idr b/tests/idris2/reflection/reflection005/refdecl.idr similarity index 100% rename from tests/idris2/reflection005/refdecl.idr rename to tests/idris2/reflection/reflection005/refdecl.idr diff --git a/tests/idris2/reflection005/run b/tests/idris2/reflection/reflection005/run similarity index 53% rename from tests/idris2/reflection005/run rename to tests/idris2/reflection/reflection005/run index e86f7b9b4f..11fac262bb 100755 --- a/tests/idris2/reflection005/run +++ b/tests/idris2/reflection/reflection005/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 refdecl.idr < input diff --git a/tests/idris2/reflection006/expected b/tests/idris2/reflection/reflection006/expected similarity index 100% rename from tests/idris2/reflection006/expected rename to tests/idris2/reflection/reflection006/expected diff --git a/tests/idris2/reflection006/input b/tests/idris2/reflection/reflection006/input similarity index 100% rename from tests/idris2/reflection006/input rename to tests/idris2/reflection/reflection006/input diff --git a/tests/idris2/reflection006/refleq.idr b/tests/idris2/reflection/reflection006/refleq.idr similarity index 100% rename from tests/idris2/reflection006/refleq.idr rename to tests/idris2/reflection/reflection006/refleq.idr diff --git a/tests/idris2/reflection/reflection006/run b/tests/idris2/reflection/reflection006/run new file mode 100755 index 0000000000..36a672918a --- /dev/null +++ b/tests/idris2/reflection/reflection006/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check refleq.idr diff --git a/tests/idris2/reflection007/NatExpr.idr b/tests/idris2/reflection/reflection007/NatExpr.idr similarity index 100% rename from tests/idris2/reflection007/NatExpr.idr rename to tests/idris2/reflection/reflection007/NatExpr.idr diff --git a/tests/idris2/reflection007/expected b/tests/idris2/reflection/reflection007/expected similarity index 100% rename from tests/idris2/reflection007/expected rename to tests/idris2/reflection/reflection007/expected diff --git a/tests/idris2/reflection007/input b/tests/idris2/reflection/reflection007/input similarity index 100% rename from tests/idris2/reflection007/input rename to tests/idris2/reflection/reflection007/input diff --git a/tests/idris2/reflection007/run b/tests/idris2/reflection/reflection007/run similarity index 53% rename from tests/idris2/reflection007/run rename to tests/idris2/reflection/reflection007/run index be06b99c47..2fec5bcc97 100755 --- a/tests/idris2/reflection007/run +++ b/tests/idris2/reflection/reflection007/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 NatExpr.idr < input diff --git a/tests/idris2/reflection008/Interp.idr b/tests/idris2/reflection/reflection008/Interp.idr similarity index 100% rename from tests/idris2/reflection008/Interp.idr rename to tests/idris2/reflection/reflection008/Interp.idr diff --git a/tests/idris2/reflection008/expected b/tests/idris2/reflection/reflection008/expected similarity index 100% rename from tests/idris2/reflection008/expected rename to tests/idris2/reflection/reflection008/expected diff --git a/tests/idris2/reflection008/input b/tests/idris2/reflection/reflection008/input similarity index 100% rename from tests/idris2/reflection008/input rename to tests/idris2/reflection/reflection008/input diff --git a/tests/idris2/reflection008/run b/tests/idris2/reflection/reflection008/run similarity index 52% rename from tests/idris2/reflection008/run rename to tests/idris2/reflection/reflection008/run index 43638a3c41..181d0dea6e 100755 --- a/tests/idris2/reflection008/run +++ b/tests/idris2/reflection/reflection008/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 Interp.idr < input diff --git a/tests/idris2/reflection009/expected b/tests/idris2/reflection/reflection009/expected similarity index 100% rename from tests/idris2/reflection009/expected rename to tests/idris2/reflection/reflection009/expected diff --git a/tests/idris2/reflection009/perf.idr b/tests/idris2/reflection/reflection009/perf.idr similarity index 100% rename from tests/idris2/reflection009/perf.idr rename to tests/idris2/reflection/reflection009/perf.idr diff --git a/tests/idris2/reflection/reflection009/run b/tests/idris2/reflection/reflection009/run new file mode 100755 index 0000000000..d9bfccbcb6 --- /dev/null +++ b/tests/idris2/reflection/reflection009/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check perf.idr diff --git a/tests/idris2/reflection010/Name.idr b/tests/idris2/reflection/reflection010/Name.idr similarity index 100% rename from tests/idris2/reflection010/Name.idr rename to tests/idris2/reflection/reflection010/Name.idr diff --git a/tests/idris2/reflection010/expected b/tests/idris2/reflection/reflection010/expected similarity index 100% rename from tests/idris2/reflection010/expected rename to tests/idris2/reflection/reflection010/expected diff --git a/tests/idris2/reflection010/run b/tests/idris2/reflection/reflection010/run similarity index 68% rename from tests/idris2/reflection010/run rename to tests/idris2/reflection/reflection010/run index bc5391405b..bf1c8e5f8c 100755 --- a/tests/idris2/reflection010/run +++ b/tests/idris2/reflection/reflection010/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh check Name.idr | sed -E "s/\.[0-9]+:[0-9]+/\.N:N/" diff --git a/tests/idris2/reflection011/expected b/tests/idris2/reflection/reflection011/expected similarity index 100% rename from tests/idris2/reflection011/expected rename to tests/idris2/reflection/reflection011/expected diff --git a/tests/idris2/reflection011/input b/tests/idris2/reflection/reflection011/input similarity index 100% rename from tests/idris2/reflection011/input rename to tests/idris2/reflection/reflection011/input diff --git a/tests/idris2/reflection011/run b/tests/idris2/reflection/reflection011/run similarity index 52% rename from tests/idris2/reflection011/run rename to tests/idris2/reflection/reflection011/run index e7143a45d0..20eec466b2 100755 --- a/tests/idris2/reflection011/run +++ b/tests/idris2/reflection/reflection011/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 tryref.idr < input diff --git a/tests/idris2/reflection011/tryref.idr b/tests/idris2/reflection/reflection011/tryref.idr similarity index 100% rename from tests/idris2/reflection011/tryref.idr rename to tests/idris2/reflection/reflection011/tryref.idr diff --git a/tests/idris2/reflection012/expected b/tests/idris2/reflection/reflection012/expected similarity index 100% rename from tests/idris2/reflection012/expected rename to tests/idris2/reflection/reflection012/expected diff --git a/tests/idris2/reflection012/input b/tests/idris2/reflection/reflection012/input similarity index 100% rename from tests/idris2/reflection012/input rename to tests/idris2/reflection/reflection012/input diff --git a/tests/idris2/reflection012/nameinfo.idr b/tests/idris2/reflection/reflection012/nameinfo.idr similarity index 100% rename from tests/idris2/reflection012/nameinfo.idr rename to tests/idris2/reflection/reflection012/nameinfo.idr diff --git a/tests/idris2/reflection012/run b/tests/idris2/reflection/reflection012/run similarity index 54% rename from tests/idris2/reflection012/run rename to tests/idris2/reflection/reflection012/run index 28c36f5f1d..b84cf2f7af 100755 --- a/tests/idris2/reflection012/run +++ b/tests/idris2/reflection/reflection012/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 nameinfo.idr < input diff --git a/tests/idris2/reflection013/WithUnambig.idr b/tests/idris2/reflection/reflection013/WithUnambig.idr similarity index 100% rename from tests/idris2/reflection013/WithUnambig.idr rename to tests/idris2/reflection/reflection013/WithUnambig.idr diff --git a/tests/idris2/reflection013/expected b/tests/idris2/reflection/reflection013/expected similarity index 100% rename from tests/idris2/reflection013/expected rename to tests/idris2/reflection/reflection013/expected diff --git a/tests/idris2/reflection/reflection013/run b/tests/idris2/reflection/reflection013/run new file mode 100755 index 0000000000..7220f45577 --- /dev/null +++ b/tests/idris2/reflection/reflection013/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check WithUnambig.idr diff --git a/tests/idris2/reflection014/expected b/tests/idris2/reflection/reflection014/expected similarity index 100% rename from tests/idris2/reflection014/expected rename to tests/idris2/reflection/reflection014/expected diff --git a/tests/idris2/reflection014/refdecl.idr b/tests/idris2/reflection/reflection014/refdecl.idr similarity index 100% rename from tests/idris2/reflection014/refdecl.idr rename to tests/idris2/reflection/reflection014/refdecl.idr diff --git a/tests/idris2/reflection/reflection014/run b/tests/idris2/reflection/reflection014/run new file mode 100755 index 0000000000..f63319fa06 --- /dev/null +++ b/tests/idris2/reflection/reflection014/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +run refdecl.idr diff --git a/tests/idris2/reflection015/MacroRetFunc.idr b/tests/idris2/reflection/reflection015/MacroRetFunc.idr similarity index 100% rename from tests/idris2/reflection015/MacroRetFunc.idr rename to tests/idris2/reflection/reflection015/MacroRetFunc.idr diff --git a/tests/idris2/reflection015/expected b/tests/idris2/reflection/reflection015/expected similarity index 100% rename from tests/idris2/reflection015/expected rename to tests/idris2/reflection/reflection015/expected diff --git a/tests/idris2/reflection015/run b/tests/idris2/reflection/reflection015/run similarity index 50% rename from tests/idris2/reflection015/run rename to tests/idris2/reflection/reflection015/run index 2e3dd06660..5adb5e1123 100755 --- a/tests/idris2/reflection015/run +++ b/tests/idris2/reflection/reflection015/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh check MacroRetFunc.idr diff --git a/tests/idris2/reflection016/BindElabScBug.idr b/tests/idris2/reflection/reflection016/BindElabScBug.idr similarity index 100% rename from tests/idris2/reflection016/BindElabScBug.idr rename to tests/idris2/reflection/reflection016/BindElabScBug.idr diff --git a/tests/idris2/reflection016/Eta.idr b/tests/idris2/reflection/reflection016/Eta.idr similarity index 100% rename from tests/idris2/reflection016/Eta.idr rename to tests/idris2/reflection/reflection016/Eta.idr diff --git a/tests/idris2/reflection016/expected b/tests/idris2/reflection/reflection016/expected similarity index 100% rename from tests/idris2/reflection016/expected rename to tests/idris2/reflection/reflection016/expected diff --git a/tests/idris2/reflection016/input b/tests/idris2/reflection/reflection016/input similarity index 100% rename from tests/idris2/reflection016/input rename to tests/idris2/reflection/reflection016/input diff --git a/tests/idris2/reflection016/run b/tests/idris2/reflection/reflection016/run similarity index 66% rename from tests/idris2/reflection016/run rename to tests/idris2/reflection/reflection016/run index c7981f9d56..f58d5b2419 100755 --- a/tests/idris2/reflection016/run +++ b/tests/idris2/reflection/reflection016/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 Eta.idr < input check BindElabScBug.idr diff --git a/tests/idris2/reflection017/CanElabType.idr b/tests/idris2/reflection/reflection017/CanElabType.idr similarity index 100% rename from tests/idris2/reflection017/CanElabType.idr rename to tests/idris2/reflection/reflection017/CanElabType.idr diff --git a/tests/idris2/reflection017/StillCantEscape.idr b/tests/idris2/reflection/reflection017/StillCantEscape.idr similarity index 100% rename from tests/idris2/reflection017/StillCantEscape.idr rename to tests/idris2/reflection/reflection017/StillCantEscape.idr diff --git a/tests/idris2/reflection017/expected b/tests/idris2/reflection/reflection017/expected similarity index 100% rename from tests/idris2/reflection017/expected rename to tests/idris2/reflection/reflection017/expected diff --git a/tests/idris2/reflection017/run b/tests/idris2/reflection/reflection017/run similarity index 67% rename from tests/idris2/reflection017/run rename to tests/idris2/reflection/reflection017/run index d2458e5ed4..199af796f0 100755 --- a/tests/idris2/reflection017/run +++ b/tests/idris2/reflection/reflection017/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh check CanElabType.idr check StillCantEscape.idr diff --git a/tests/idris2/reflection018/AtTypeLevel.idr b/tests/idris2/reflection/reflection018/AtTypeLevel.idr similarity index 100% rename from tests/idris2/reflection018/AtTypeLevel.idr rename to tests/idris2/reflection/reflection018/AtTypeLevel.idr diff --git a/tests/idris2/reflection018/expected b/tests/idris2/reflection/reflection018/expected similarity index 100% rename from tests/idris2/reflection018/expected rename to tests/idris2/reflection/reflection018/expected diff --git a/tests/idris2/reflection/reflection018/run b/tests/idris2/reflection/reflection018/run new file mode 100755 index 0000000000..46c3b1c5ba --- /dev/null +++ b/tests/idris2/reflection/reflection018/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check AtTypeLevel.idr diff --git a/tests/idris2/reflection019/ElabScriptWarning.idr b/tests/idris2/reflection/reflection019/ElabScriptWarning.idr similarity index 100% rename from tests/idris2/reflection019/ElabScriptWarning.idr rename to tests/idris2/reflection/reflection019/ElabScriptWarning.idr diff --git a/tests/idris2/reflection019/expected b/tests/idris2/reflection/reflection019/expected similarity index 100% rename from tests/idris2/reflection019/expected rename to tests/idris2/reflection/reflection019/expected diff --git a/tests/idris2/reflection019/input b/tests/idris2/reflection/reflection019/input similarity index 100% rename from tests/idris2/reflection019/input rename to tests/idris2/reflection/reflection019/input diff --git a/tests/idris2/reflection019/run b/tests/idris2/reflection/reflection019/run similarity index 61% rename from tests/idris2/reflection019/run rename to tests/idris2/reflection/reflection019/run index c779442ba2..d9ba0698c9 100755 --- a/tests/idris2/reflection019/run +++ b/tests/idris2/reflection/reflection019/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 ElabScriptWarning.idr < input diff --git a/tests/idris2/reflection019/test.ipkg b/tests/idris2/reflection/reflection019/test.ipkg similarity index 100% rename from tests/idris2/reflection019/test.ipkg rename to tests/idris2/reflection/reflection019/test.ipkg diff --git a/tests/idris2/reflection020/FromDecls.idr b/tests/idris2/reflection/reflection020/FromDecls.idr similarity index 100% rename from tests/idris2/reflection020/FromDecls.idr rename to tests/idris2/reflection/reflection020/FromDecls.idr diff --git a/tests/idris2/reflection020/FromName.idr b/tests/idris2/reflection/reflection020/FromName.idr similarity index 100% rename from tests/idris2/reflection020/FromName.idr rename to tests/idris2/reflection/reflection020/FromName.idr diff --git a/tests/idris2/reflection020/FromTTImp.idr b/tests/idris2/reflection/reflection020/FromTTImp.idr similarity index 100% rename from tests/idris2/reflection020/FromTTImp.idr rename to tests/idris2/reflection/reflection020/FromTTImp.idr diff --git a/tests/idris2/reflection020/expected b/tests/idris2/reflection/reflection020/expected similarity index 100% rename from tests/idris2/reflection020/expected rename to tests/idris2/reflection/reflection020/expected diff --git a/tests/idris2/reflection020/run b/tests/idris2/reflection/reflection020/run similarity index 94% rename from tests/idris2/reflection020/run rename to tests/idris2/reflection/reflection020/run index c82eda21a5..e80bc66483 100644 --- a/tests/idris2/reflection020/run +++ b/tests/idris2/reflection/reflection020/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 FromTTImp.idr << EOF :printdef natExprMacroTest diff --git a/tests/idris2/reflection021/QuoteSearch.idr b/tests/idris2/reflection/reflection021/QuoteSearch.idr similarity index 100% rename from tests/idris2/reflection021/QuoteSearch.idr rename to tests/idris2/reflection/reflection021/QuoteSearch.idr diff --git a/tests/idris2/reflection021/expected b/tests/idris2/reflection/reflection021/expected similarity index 100% rename from tests/idris2/reflection021/expected rename to tests/idris2/reflection/reflection021/expected diff --git a/tests/idris2/reflection/reflection021/run b/tests/idris2/reflection/reflection021/run new file mode 100755 index 0000000000..ebfe42ba08 --- /dev/null +++ b/tests/idris2/reflection/reflection021/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check QuoteSearch.idr diff --git a/tests/idris2/reflection003/run b/tests/idris2/reflection003/run deleted file mode 100755 index 5c4359d7a3..0000000000 --- a/tests/idris2/reflection003/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check refprims.idr diff --git a/tests/idris2/reflection006/run b/tests/idris2/reflection006/run deleted file mode 100755 index f70c2e6a82..0000000000 --- a/tests/idris2/reflection006/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check refleq.idr diff --git a/tests/idris2/reflection009/run b/tests/idris2/reflection009/run deleted file mode 100755 index c938b03d82..0000000000 --- a/tests/idris2/reflection009/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check perf.idr diff --git a/tests/idris2/reflection013/run b/tests/idris2/reflection013/run deleted file mode 100755 index 7a48fcc609..0000000000 --- a/tests/idris2/reflection013/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check WithUnambig.idr diff --git a/tests/idris2/reflection014/run b/tests/idris2/reflection014/run deleted file mode 100755 index 5c104e274c..0000000000 --- a/tests/idris2/reflection014/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -run refdecl.idr diff --git a/tests/idris2/reflection018/run b/tests/idris2/reflection018/run deleted file mode 100755 index 4abf199262..0000000000 --- a/tests/idris2/reflection018/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check AtTypeLevel.idr diff --git a/tests/idris2/reflection021/run b/tests/idris2/reflection021/run deleted file mode 100755 index 4ca6d240e0..0000000000 --- a/tests/idris2/reflection021/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check QuoteSearch.idr diff --git a/tests/idris2/reg001/D.idr b/tests/idris2/reg/reg001/D.idr similarity index 100% rename from tests/idris2/reg001/D.idr rename to tests/idris2/reg/reg001/D.idr diff --git a/tests/idris2/reg001/expected b/tests/idris2/reg/reg001/expected similarity index 100% rename from tests/idris2/reg001/expected rename to tests/idris2/reg/reg001/expected diff --git a/tests/idris2/reg/reg001/run b/tests/idris2/reg/reg001/run new file mode 100755 index 0000000000..a5c6c13ca1 --- /dev/null +++ b/tests/idris2/reg/reg001/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check D.idr diff --git a/tests/idris2/reg002/expected b/tests/idris2/reg/reg002/expected similarity index 100% rename from tests/idris2/reg002/expected rename to tests/idris2/reg/reg002/expected diff --git a/tests/idris2/reg002/linm.idr b/tests/idris2/reg/reg002/linm.idr similarity index 100% rename from tests/idris2/reg002/linm.idr rename to tests/idris2/reg/reg002/linm.idr diff --git a/tests/idris2/reg/reg002/run b/tests/idris2/reg/reg002/run new file mode 100755 index 0000000000..aee2831ace --- /dev/null +++ b/tests/idris2/reg/reg002/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check linm.idr diff --git a/tests/idris2/reg003/Holes.idr b/tests/idris2/reg/reg003/Holes.idr similarity index 100% rename from tests/idris2/reg003/Holes.idr rename to tests/idris2/reg/reg003/Holes.idr diff --git a/tests/idris2/reg003/expected b/tests/idris2/reg/reg003/expected similarity index 100% rename from tests/idris2/reg003/expected rename to tests/idris2/reg/reg003/expected diff --git a/tests/idris2/reg003/input b/tests/idris2/reg/reg003/input similarity index 100% rename from tests/idris2/reg003/input rename to tests/idris2/reg/reg003/input diff --git a/tests/idris2/reg003/run b/tests/idris2/reg/reg003/run similarity index 52% rename from tests/idris2/reg003/run rename to tests/idris2/reg/reg003/run index e33955b1a0..33d7384221 100755 --- a/tests/idris2/reg003/run +++ b/tests/idris2/reg/reg003/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 Holes.idr < input diff --git a/tests/idris2/reg004/ambig.idr b/tests/idris2/reg/reg004/ambig.idr similarity index 100% rename from tests/idris2/reg004/ambig.idr rename to tests/idris2/reg/reg004/ambig.idr diff --git a/tests/idris2/reg004/expected b/tests/idris2/reg/reg004/expected similarity index 100% rename from tests/idris2/reg004/expected rename to tests/idris2/reg/reg004/expected diff --git a/tests/idris2/reg/reg004/run b/tests/idris2/reg/reg004/run new file mode 100755 index 0000000000..735312c932 --- /dev/null +++ b/tests/idris2/reg/reg004/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check ambig.idr diff --git a/tests/idris2/reg005/expected b/tests/idris2/reg/reg005/expected similarity index 100% rename from tests/idris2/reg005/expected rename to tests/idris2/reg/reg005/expected diff --git a/tests/idris2/reg005/iftype.idr b/tests/idris2/reg/reg005/iftype.idr similarity index 100% rename from tests/idris2/reg005/iftype.idr rename to tests/idris2/reg/reg005/iftype.idr diff --git a/tests/idris2/reg/reg005/run b/tests/idris2/reg/reg005/run new file mode 100755 index 0000000000..1f2a4bae16 --- /dev/null +++ b/tests/idris2/reg/reg005/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check iftype.idr diff --git a/tests/idris2/reg006/Cmd.idr b/tests/idris2/reg/reg006/Cmd.idr similarity index 100% rename from tests/idris2/reg006/Cmd.idr rename to tests/idris2/reg/reg006/Cmd.idr diff --git a/tests/idris2/reg006/expected b/tests/idris2/reg/reg006/expected similarity index 100% rename from tests/idris2/reg006/expected rename to tests/idris2/reg/reg006/expected diff --git a/tests/idris2/reg/reg006/run b/tests/idris2/reg/reg006/run new file mode 100755 index 0000000000..7cce14eda5 --- /dev/null +++ b/tests/idris2/reg/reg006/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Cmd.idr diff --git a/tests/idris2/reg007/Main.idr b/tests/idris2/reg/reg007/Main.idr similarity index 100% rename from tests/idris2/reg007/Main.idr rename to tests/idris2/reg/reg007/Main.idr diff --git a/tests/idris2/reg007/expected b/tests/idris2/reg/reg007/expected similarity index 100% rename from tests/idris2/reg007/expected rename to tests/idris2/reg/reg007/expected diff --git a/tests/idris2/reg/reg007/run b/tests/idris2/reg/reg007/run new file mode 100755 index 0000000000..d49bfd87de --- /dev/null +++ b/tests/idris2/reg/reg007/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Main.idr diff --git a/tests/idris2/reg008/Vending.idr b/tests/idris2/reg/reg008/Vending.idr similarity index 100% rename from tests/idris2/reg008/Vending.idr rename to tests/idris2/reg/reg008/Vending.idr diff --git a/tests/idris2/reg008/expected b/tests/idris2/reg/reg008/expected similarity index 100% rename from tests/idris2/reg008/expected rename to tests/idris2/reg/reg008/expected diff --git a/tests/idris2/reg008/input b/tests/idris2/reg/reg008/input similarity index 100% rename from tests/idris2/reg008/input rename to tests/idris2/reg/reg008/input diff --git a/tests/idris2/reg008/run b/tests/idris2/reg/reg008/run similarity index 66% rename from tests/idris2/reg008/run rename to tests/idris2/reg/reg008/run index 7457509d1f..8ca87f86c1 100755 --- a/tests/idris2/reg008/run +++ b/tests/idris2/reg/reg008/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 Vending.idr --debug-elab-check < input diff --git a/tests/idris2/reg009/Case.idr b/tests/idris2/reg/reg009/Case.idr similarity index 100% rename from tests/idris2/reg009/Case.idr rename to tests/idris2/reg/reg009/Case.idr diff --git a/tests/idris2/reg009/expected b/tests/idris2/reg/reg009/expected similarity index 100% rename from tests/idris2/reg009/expected rename to tests/idris2/reg/reg009/expected diff --git a/tests/idris2/reg/reg009/run b/tests/idris2/reg/reg009/run new file mode 100755 index 0000000000..e289dcebfd --- /dev/null +++ b/tests/idris2/reg/reg009/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Case.idr diff --git a/tests/idris2/reg010/Recordname.idr b/tests/idris2/reg/reg010/Recordname.idr similarity index 100% rename from tests/idris2/reg010/Recordname.idr rename to tests/idris2/reg/reg010/Recordname.idr diff --git a/tests/idris2/reg010/expected b/tests/idris2/reg/reg010/expected similarity index 100% rename from tests/idris2/reg010/expected rename to tests/idris2/reg/reg010/expected diff --git a/tests/idris2/reg/reg010/run b/tests/idris2/reg/reg010/run new file mode 100755 index 0000000000..55306fab82 --- /dev/null +++ b/tests/idris2/reg/reg010/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Recordname.idr diff --git a/tests/idris2/reg011/expected b/tests/idris2/reg/reg011/expected similarity index 100% rename from tests/idris2/reg011/expected rename to tests/idris2/reg/reg011/expected diff --git a/tests/idris2/reg011/mut.idr b/tests/idris2/reg/reg011/mut.idr similarity index 100% rename from tests/idris2/reg011/mut.idr rename to tests/idris2/reg/reg011/mut.idr diff --git a/tests/idris2/reg/reg011/run b/tests/idris2/reg/reg011/run new file mode 100755 index 0000000000..ed26311e5d --- /dev/null +++ b/tests/idris2/reg/reg011/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check mut.idr diff --git a/tests/idris2/reg012/Foo.idr b/tests/idris2/reg/reg012/Foo.idr similarity index 100% rename from tests/idris2/reg012/Foo.idr rename to tests/idris2/reg/reg012/Foo.idr diff --git a/tests/idris2/reg012/expected b/tests/idris2/reg/reg012/expected similarity index 100% rename from tests/idris2/reg012/expected rename to tests/idris2/reg/reg012/expected diff --git a/tests/idris2/reg/reg012/run b/tests/idris2/reg/reg012/run new file mode 100755 index 0000000000..7bb0c6fdd4 --- /dev/null +++ b/tests/idris2/reg/reg012/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Foo.idr diff --git a/tests/idris2/reg013/UnboundImplicits.idr b/tests/idris2/reg/reg013/UnboundImplicits.idr similarity index 100% rename from tests/idris2/reg013/UnboundImplicits.idr rename to tests/idris2/reg/reg013/UnboundImplicits.idr diff --git a/tests/idris2/reg013/expected b/tests/idris2/reg/reg013/expected similarity index 100% rename from tests/idris2/reg013/expected rename to tests/idris2/reg/reg013/expected diff --git a/tests/idris2/reg013/run b/tests/idris2/reg/reg013/run similarity index 53% rename from tests/idris2/reg013/run rename to tests/idris2/reg/reg013/run index 8b47b02e13..3fcbc0fb40 100755 --- a/tests/idris2/reg013/run +++ b/tests/idris2/reg/reg013/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh check UnboundImplicits.idr diff --git a/tests/idris2/reg014/casecase.idr b/tests/idris2/reg/reg014/casecase.idr similarity index 100% rename from tests/idris2/reg014/casecase.idr rename to tests/idris2/reg/reg014/casecase.idr diff --git a/tests/idris2/reg014/expected b/tests/idris2/reg/reg014/expected similarity index 100% rename from tests/idris2/reg014/expected rename to tests/idris2/reg/reg014/expected diff --git a/tests/idris2/reg014/run b/tests/idris2/reg/reg014/run similarity index 61% rename from tests/idris2/reg014/run rename to tests/idris2/reg/reg014/run index be0ede11d9..81ca0d5bb6 100755 --- a/tests/idris2/reg014/run +++ b/tests/idris2/reg/reg014/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh check casecase.idr --debug-elab-check diff --git a/tests/idris2/reg015/anyfail.idr b/tests/idris2/reg/reg015/anyfail.idr similarity index 100% rename from tests/idris2/reg015/anyfail.idr rename to tests/idris2/reg/reg015/anyfail.idr diff --git a/tests/idris2/reg015/expected b/tests/idris2/reg/reg015/expected similarity index 100% rename from tests/idris2/reg015/expected rename to tests/idris2/reg/reg015/expected diff --git a/tests/idris2/reg/reg015/run b/tests/idris2/reg/reg015/run new file mode 100755 index 0000000000..4fe465460c --- /dev/null +++ b/tests/idris2/reg/reg015/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check anyfail.idr diff --git a/tests/idris2/reg016/Using.idr b/tests/idris2/reg/reg016/Using.idr similarity index 100% rename from tests/idris2/reg016/Using.idr rename to tests/idris2/reg/reg016/Using.idr diff --git a/tests/idris2/reg016/expected b/tests/idris2/reg/reg016/expected similarity index 100% rename from tests/idris2/reg016/expected rename to tests/idris2/reg/reg016/expected diff --git a/tests/idris2/reg/reg016/run b/tests/idris2/reg/reg016/run new file mode 100755 index 0000000000..8ca58acea0 --- /dev/null +++ b/tests/idris2/reg/reg016/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Using.idr diff --git a/tests/idris2/reg017/expected b/tests/idris2/reg/reg017/expected similarity index 100% rename from tests/idris2/reg017/expected rename to tests/idris2/reg/reg017/expected diff --git a/tests/idris2/reg017/lammult.idr b/tests/idris2/reg/reg017/lammult.idr similarity index 100% rename from tests/idris2/reg017/lammult.idr rename to tests/idris2/reg/reg017/lammult.idr diff --git a/tests/idris2/reg/reg017/run b/tests/idris2/reg/reg017/run new file mode 100755 index 0000000000..00c764f3e7 --- /dev/null +++ b/tests/idris2/reg/reg017/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check lammult.idr diff --git a/tests/idris2/reg018/cycle.idr b/tests/idris2/reg/reg018/cycle.idr similarity index 100% rename from tests/idris2/reg018/cycle.idr rename to tests/idris2/reg/reg018/cycle.idr diff --git a/tests/idris2/reg018/expected b/tests/idris2/reg/reg018/expected similarity index 100% rename from tests/idris2/reg018/expected rename to tests/idris2/reg/reg018/expected diff --git a/tests/idris2/reg/reg018/run b/tests/idris2/reg/reg018/run new file mode 100755 index 0000000000..14702a32b8 --- /dev/null +++ b/tests/idris2/reg/reg018/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check cycle.idr diff --git a/tests/idris2/reg019/expected b/tests/idris2/reg/reg019/expected similarity index 100% rename from tests/idris2/reg019/expected rename to tests/idris2/reg/reg019/expected diff --git a/tests/idris2/reg019/lazybug.idr b/tests/idris2/reg/reg019/lazybug.idr similarity index 100% rename from tests/idris2/reg019/lazybug.idr rename to tests/idris2/reg/reg019/lazybug.idr diff --git a/tests/idris2/reg/reg019/run b/tests/idris2/reg/reg019/run new file mode 100755 index 0000000000..b5a938dfea --- /dev/null +++ b/tests/idris2/reg/reg019/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check lazybug.idr diff --git a/tests/idris2/reg020/expected b/tests/idris2/reg/reg020/expected similarity index 100% rename from tests/idris2/reg020/expected rename to tests/idris2/reg/reg020/expected diff --git a/tests/idris2/reg020/input b/tests/idris2/reg/reg020/input similarity index 100% rename from tests/idris2/reg020/input rename to tests/idris2/reg/reg020/input diff --git a/tests/idris2/reg020/matchlits.idr b/tests/idris2/reg/reg020/matchlits.idr similarity index 100% rename from tests/idris2/reg020/matchlits.idr rename to tests/idris2/reg/reg020/matchlits.idr diff --git a/tests/idris2/reg020/run b/tests/idris2/reg/reg020/run similarity index 55% rename from tests/idris2/reg020/run rename to tests/idris2/reg/reg020/run index 9de98b0472..c1de738a7e 100755 --- a/tests/idris2/reg020/run +++ b/tests/idris2/reg/reg020/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 matchlits.idr < input diff --git a/tests/idris2/reg021/case.idr b/tests/idris2/reg/reg021/case.idr similarity index 100% rename from tests/idris2/reg021/case.idr rename to tests/idris2/reg/reg021/case.idr diff --git a/tests/idris2/reg021/expected b/tests/idris2/reg/reg021/expected similarity index 100% rename from tests/idris2/reg021/expected rename to tests/idris2/reg/reg021/expected diff --git a/tests/idris2/reg/reg021/run b/tests/idris2/reg/reg021/run new file mode 100755 index 0000000000..ece2f87a9b --- /dev/null +++ b/tests/idris2/reg/reg021/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check case.idr diff --git a/tests/idris2/reg022/case.idr b/tests/idris2/reg/reg022/case.idr similarity index 100% rename from tests/idris2/reg022/case.idr rename to tests/idris2/reg/reg022/case.idr diff --git a/tests/idris2/reg022/expected b/tests/idris2/reg/reg022/expected similarity index 100% rename from tests/idris2/reg022/expected rename to tests/idris2/reg/reg022/expected diff --git a/tests/idris2/reg022/input b/tests/idris2/reg/reg022/input similarity index 100% rename from tests/idris2/reg022/input rename to tests/idris2/reg/reg022/input diff --git a/tests/idris2/reg/reg022/run b/tests/idris2/reg/reg022/run new file mode 100755 index 0000000000..ece2f87a9b --- /dev/null +++ b/tests/idris2/reg/reg022/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check case.idr diff --git a/tests/idris2/reg023/boom.idr b/tests/idris2/reg/reg023/boom.idr similarity index 100% rename from tests/idris2/reg023/boom.idr rename to tests/idris2/reg/reg023/boom.idr diff --git a/tests/idris2/reg023/expected b/tests/idris2/reg/reg023/expected similarity index 100% rename from tests/idris2/reg023/expected rename to tests/idris2/reg/reg023/expected diff --git a/tests/idris2/reg/reg023/run b/tests/idris2/reg/reg023/run new file mode 100755 index 0000000000..a35ea0c594 --- /dev/null +++ b/tests/idris2/reg/reg023/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check boom.idr diff --git a/tests/idris2/reg024/expected b/tests/idris2/reg/reg024/expected similarity index 100% rename from tests/idris2/reg024/expected rename to tests/idris2/reg/reg024/expected diff --git a/tests/idris2/reg024/input b/tests/idris2/reg/reg024/input similarity index 100% rename from tests/idris2/reg024/input rename to tests/idris2/reg/reg024/input diff --git a/tests/idris2/reg024/run b/tests/idris2/reg/reg024/run similarity index 52% rename from tests/idris2/reg024/run rename to tests/idris2/reg/reg024/run index 17f6211da4..800ee3e4f2 100755 --- a/tests/idris2/reg024/run +++ b/tests/idris2/reg/reg024/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 split.idr < input diff --git a/tests/idris2/reg024/split.idr b/tests/idris2/reg/reg024/split.idr similarity index 100% rename from tests/idris2/reg024/split.idr rename to tests/idris2/reg/reg024/split.idr diff --git a/tests/idris2/reg025/expected b/tests/idris2/reg/reg025/expected similarity index 100% rename from tests/idris2/reg025/expected rename to tests/idris2/reg/reg025/expected diff --git a/tests/idris2/reg025/input b/tests/idris2/reg/reg025/input similarity index 100% rename from tests/idris2/reg025/input rename to tests/idris2/reg/reg025/input diff --git a/tests/idris2/reg025/lift.idr b/tests/idris2/reg/reg025/lift.idr similarity index 100% rename from tests/idris2/reg025/lift.idr rename to tests/idris2/reg/reg025/lift.idr diff --git a/tests/idris2/reg025/run b/tests/idris2/reg/reg025/run similarity index 51% rename from tests/idris2/reg025/run rename to tests/idris2/reg/reg025/run index 157940aa4b..d65a908ea7 100755 --- a/tests/idris2/reg025/run +++ b/tests/idris2/reg/reg025/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 lift.idr < input diff --git a/tests/idris2/reg026/Meh.idr b/tests/idris2/reg/reg026/Meh.idr similarity index 100% rename from tests/idris2/reg026/Meh.idr rename to tests/idris2/reg/reg026/Meh.idr diff --git a/tests/idris2/reg026/expected b/tests/idris2/reg/reg026/expected similarity index 100% rename from tests/idris2/reg026/expected rename to tests/idris2/reg/reg026/expected diff --git a/tests/idris2/reg/reg026/run b/tests/idris2/reg/reg026/run new file mode 100755 index 0000000000..0012d13f2a --- /dev/null +++ b/tests/idris2/reg/reg026/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Meh.idr diff --git a/tests/idris2/reg027/expected b/tests/idris2/reg/reg027/expected similarity index 100% rename from tests/idris2/reg027/expected rename to tests/idris2/reg/reg027/expected diff --git a/tests/idris2/reg027/pwhere.idr b/tests/idris2/reg/reg027/pwhere.idr similarity index 100% rename from tests/idris2/reg027/pwhere.idr rename to tests/idris2/reg/reg027/pwhere.idr diff --git a/tests/idris2/reg/reg027/run b/tests/idris2/reg/reg027/run new file mode 100755 index 0000000000..4f32d300f6 --- /dev/null +++ b/tests/idris2/reg/reg027/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check pwhere.idr diff --git a/tests/idris2/reg028/Test.idr b/tests/idris2/reg/reg028/Test.idr similarity index 100% rename from tests/idris2/reg028/Test.idr rename to tests/idris2/reg/reg028/Test.idr diff --git a/tests/idris2/reg028/expected b/tests/idris2/reg/reg028/expected similarity index 100% rename from tests/idris2/reg028/expected rename to tests/idris2/reg/reg028/expected diff --git a/tests/idris2/reg/reg028/run b/tests/idris2/reg/reg028/run new file mode 100755 index 0000000000..d35387bb39 --- /dev/null +++ b/tests/idris2/reg/reg028/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Test.idr diff --git a/tests/idris2/reg029/expected b/tests/idris2/reg/reg029/expected similarity index 100% rename from tests/idris2/reg029/expected rename to tests/idris2/reg/reg029/expected diff --git a/tests/idris2/reg029/lqueue.idr b/tests/idris2/reg/reg029/lqueue.idr similarity index 100% rename from tests/idris2/reg029/lqueue.idr rename to tests/idris2/reg/reg029/lqueue.idr diff --git a/tests/idris2/reg/reg029/run b/tests/idris2/reg/reg029/run new file mode 100755 index 0000000000..d929310cef --- /dev/null +++ b/tests/idris2/reg/reg029/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check lqueue.idr diff --git a/tests/idris2/reg030/A.idr b/tests/idris2/reg/reg030/A.idr similarity index 100% rename from tests/idris2/reg030/A.idr rename to tests/idris2/reg/reg030/A.idr diff --git a/tests/idris2/reg030/B.idr b/tests/idris2/reg/reg030/B.idr similarity index 100% rename from tests/idris2/reg030/B.idr rename to tests/idris2/reg/reg030/B.idr diff --git a/tests/idris2/reg030/C.idr b/tests/idris2/reg/reg030/C.idr similarity index 100% rename from tests/idris2/reg030/C.idr rename to tests/idris2/reg/reg030/C.idr diff --git a/tests/idris2/reg030/expected b/tests/idris2/reg/reg030/expected similarity index 100% rename from tests/idris2/reg030/expected rename to tests/idris2/reg/reg030/expected diff --git a/tests/idris2/reg/reg030/run b/tests/idris2/reg/reg030/run new file mode 100755 index 0000000000..54a33fa376 --- /dev/null +++ b/tests/idris2/reg/reg030/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check A.idr diff --git a/tests/idris2/reg031/dpair.idr b/tests/idris2/reg/reg031/dpair.idr similarity index 100% rename from tests/idris2/reg031/dpair.idr rename to tests/idris2/reg/reg031/dpair.idr diff --git a/tests/idris2/reg031/expected b/tests/idris2/reg/reg031/expected similarity index 100% rename from tests/idris2/reg031/expected rename to tests/idris2/reg/reg031/expected diff --git a/tests/idris2/reg/reg031/run b/tests/idris2/reg/reg031/run new file mode 100755 index 0000000000..5c91651b0d --- /dev/null +++ b/tests/idris2/reg/reg031/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check dpair.idr diff --git a/tests/idris2/reg032/expected b/tests/idris2/reg/reg032/expected similarity index 100% rename from tests/idris2/reg032/expected rename to tests/idris2/reg/reg032/expected diff --git a/tests/idris2/reg032/recupdate.idr b/tests/idris2/reg/reg032/recupdate.idr similarity index 100% rename from tests/idris2/reg032/recupdate.idr rename to tests/idris2/reg/reg032/recupdate.idr diff --git a/tests/idris2/reg/reg032/run b/tests/idris2/reg/reg032/run new file mode 100755 index 0000000000..d3347bad5d --- /dev/null +++ b/tests/idris2/reg/reg032/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check recupdate.idr diff --git a/tests/idris2/reg033/DerivingEq.idr b/tests/idris2/reg/reg033/DerivingEq.idr similarity index 100% rename from tests/idris2/reg033/DerivingEq.idr rename to tests/idris2/reg/reg033/DerivingEq.idr diff --git a/tests/idris2/reg033/expected b/tests/idris2/reg/reg033/expected similarity index 100% rename from tests/idris2/reg033/expected rename to tests/idris2/reg/reg033/expected diff --git a/tests/idris2/reg/reg033/run b/tests/idris2/reg/reg033/run new file mode 100755 index 0000000000..663b5eec02 --- /dev/null +++ b/tests/idris2/reg/reg033/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check test.idr diff --git a/tests/idris2/reg033/test.idr b/tests/idris2/reg/reg033/test.idr similarity index 100% rename from tests/idris2/reg033/test.idr rename to tests/idris2/reg/reg033/test.idr diff --git a/tests/idris2/reg034/expected b/tests/idris2/reg/reg034/expected similarity index 100% rename from tests/idris2/reg034/expected rename to tests/idris2/reg/reg034/expected diff --git a/tests/idris2/reg/reg034/run b/tests/idris2/reg/reg034/run new file mode 100755 index 0000000000..0e271ecf5b --- /dev/null +++ b/tests/idris2/reg/reg034/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check void.idr diff --git a/tests/idris2/reg034/void.idr b/tests/idris2/reg/reg034/void.idr similarity index 100% rename from tests/idris2/reg034/void.idr rename to tests/idris2/reg/reg034/void.idr diff --git a/tests/idris2/reg035/Implicit.idr b/tests/idris2/reg/reg035/Implicit.idr similarity index 100% rename from tests/idris2/reg035/Implicit.idr rename to tests/idris2/reg/reg035/Implicit.idr diff --git a/tests/idris2/reg035/expected b/tests/idris2/reg/reg035/expected similarity index 100% rename from tests/idris2/reg035/expected rename to tests/idris2/reg/reg035/expected diff --git a/tests/idris2/reg/reg035/run b/tests/idris2/reg/reg035/run new file mode 100755 index 0000000000..c88212c0ed --- /dev/null +++ b/tests/idris2/reg/reg035/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Implicit.idr diff --git a/tests/idris2/reg036/Test.idr b/tests/idris2/reg/reg036/Test.idr similarity index 100% rename from tests/idris2/reg036/Test.idr rename to tests/idris2/reg/reg036/Test.idr diff --git a/tests/idris2/reg036/expected b/tests/idris2/reg/reg036/expected similarity index 100% rename from tests/idris2/reg036/expected rename to tests/idris2/reg/reg036/expected diff --git a/tests/idris2/reg/reg036/run b/tests/idris2/reg/reg036/run new file mode 100755 index 0000000000..d35387bb39 --- /dev/null +++ b/tests/idris2/reg/reg036/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Test.idr diff --git a/tests/idris2/reg037/Test.idr b/tests/idris2/reg/reg037/Test.idr similarity index 100% rename from tests/idris2/reg037/Test.idr rename to tests/idris2/reg/reg037/Test.idr diff --git a/tests/idris2/reg037/expected b/tests/idris2/reg/reg037/expected similarity index 100% rename from tests/idris2/reg037/expected rename to tests/idris2/reg/reg037/expected diff --git a/tests/idris2/reg/reg037/run b/tests/idris2/reg/reg037/run new file mode 100755 index 0000000000..d35387bb39 --- /dev/null +++ b/tests/idris2/reg/reg037/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Test.idr diff --git a/tests/idris2/reg038/Test1.idr b/tests/idris2/reg/reg038/Test1.idr similarity index 100% rename from tests/idris2/reg038/Test1.idr rename to tests/idris2/reg/reg038/Test1.idr diff --git a/tests/idris2/reg038/Test2.idr b/tests/idris2/reg/reg038/Test2.idr similarity index 100% rename from tests/idris2/reg038/Test2.idr rename to tests/idris2/reg/reg038/Test2.idr diff --git a/tests/idris2/reg038/expected b/tests/idris2/reg/reg038/expected similarity index 100% rename from tests/idris2/reg038/expected rename to tests/idris2/reg/reg038/expected diff --git a/tests/idris2/reg038/run b/tests/idris2/reg/reg038/run similarity index 57% rename from tests/idris2/reg038/run rename to tests/idris2/reg/reg038/run index 50feabb225..d49c9cf383 100755 --- a/tests/idris2/reg038/run +++ b/tests/idris2/reg/reg038/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh check Test1.idr check Test2.idr diff --git a/tests/idris2/reg039/dupdup.idr b/tests/idris2/reg/reg039/dupdup.idr similarity index 100% rename from tests/idris2/reg039/dupdup.idr rename to tests/idris2/reg/reg039/dupdup.idr diff --git a/tests/idris2/reg039/expected b/tests/idris2/reg/reg039/expected similarity index 100% rename from tests/idris2/reg039/expected rename to tests/idris2/reg/reg039/expected diff --git a/tests/idris2/reg/reg039/run b/tests/idris2/reg/reg039/run new file mode 100755 index 0000000000..c7225e0972 --- /dev/null +++ b/tests/idris2/reg/reg039/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check dupdup.idr diff --git a/tests/idris2/reg040/CoverBug.idr b/tests/idris2/reg/reg040/CoverBug.idr similarity index 100% rename from tests/idris2/reg040/CoverBug.idr rename to tests/idris2/reg/reg040/CoverBug.idr diff --git a/tests/idris2/reg040/expected b/tests/idris2/reg/reg040/expected similarity index 100% rename from tests/idris2/reg040/expected rename to tests/idris2/reg/reg040/expected diff --git a/tests/idris2/reg/reg040/run b/tests/idris2/reg/reg040/run new file mode 100755 index 0000000000..159f83fc25 --- /dev/null +++ b/tests/idris2/reg/reg040/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check CoverBug.idr diff --git a/tests/idris2/reg041/expected b/tests/idris2/reg/reg041/expected similarity index 100% rename from tests/idris2/reg041/expected rename to tests/idris2/reg/reg041/expected diff --git a/tests/idris2/reg/reg041/run b/tests/idris2/reg/reg041/run new file mode 100755 index 0000000000..4f15a90464 --- /dev/null +++ b/tests/idris2/reg/reg041/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check tuple.idr diff --git a/tests/idris2/reg041/tuple.idr b/tests/idris2/reg/reg041/tuple.idr similarity index 100% rename from tests/idris2/reg041/tuple.idr rename to tests/idris2/reg/reg041/tuple.idr diff --git a/tests/idris2/reg042/NatOpts.idr b/tests/idris2/reg/reg042/NatOpts.idr similarity index 100% rename from tests/idris2/reg042/NatOpts.idr rename to tests/idris2/reg/reg042/NatOpts.idr diff --git a/tests/idris2/reg042/expected b/tests/idris2/reg/reg042/expected similarity index 100% rename from tests/idris2/reg042/expected rename to tests/idris2/reg/reg042/expected diff --git a/tests/idris2/reg042/input b/tests/idris2/reg/reg042/input similarity index 100% rename from tests/idris2/reg042/input rename to tests/idris2/reg/reg042/input diff --git a/tests/idris2/reg042/run b/tests/idris2/reg/reg042/run similarity index 53% rename from tests/idris2/reg042/run rename to tests/idris2/reg/reg042/run index 2f7509220c..508f18835f 100644 --- a/tests/idris2/reg042/run +++ b/tests/idris2/reg/reg042/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 NatOpts.idr < input diff --git a/tests/idris2/reg043/NotFake.idr b/tests/idris2/reg/reg043/NotFake.idr similarity index 100% rename from tests/idris2/reg043/NotFake.idr rename to tests/idris2/reg/reg043/NotFake.idr diff --git a/tests/idris2/reg043/TestFake.idr b/tests/idris2/reg/reg043/TestFake.idr similarity index 100% rename from tests/idris2/reg043/TestFake.idr rename to tests/idris2/reg/reg043/TestFake.idr diff --git a/tests/idris2/reg043/expected b/tests/idris2/reg/reg043/expected similarity index 100% rename from tests/idris2/reg043/expected rename to tests/idris2/reg/reg043/expected diff --git a/tests/idris2/reg043/run b/tests/idris2/reg/reg043/run similarity index 81% rename from tests/idris2/reg043/run rename to tests/idris2/reg/reg043/run index 740fd28d27..de20f836da 100755 --- a/tests/idris2/reg043/run +++ b/tests/idris2/reg/reg043/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh check NotFake.idr diff --git a/tests/idris2/reg044/Methods.idr b/tests/idris2/reg/reg044/Methods.idr similarity index 100% rename from tests/idris2/reg044/Methods.idr rename to tests/idris2/reg/reg044/Methods.idr diff --git a/tests/idris2/reg044/expected b/tests/idris2/reg/reg044/expected similarity index 100% rename from tests/idris2/reg044/expected rename to tests/idris2/reg/reg044/expected diff --git a/tests/idris2/reg/reg044/run b/tests/idris2/reg/reg044/run new file mode 100755 index 0000000000..c870602d6c --- /dev/null +++ b/tests/idris2/reg/reg044/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Methods.idr diff --git a/tests/idris2/reg045/expected b/tests/idris2/reg/reg045/expected similarity index 100% rename from tests/idris2/reg045/expected rename to tests/idris2/reg/reg045/expected diff --git a/tests/idris2/reg/reg045/run b/tests/idris2/reg/reg045/run new file mode 100755 index 0000000000..bf95787c03 --- /dev/null +++ b/tests/idris2/reg/reg045/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check withparams.idr diff --git a/tests/idris2/reg045/withparams.idr b/tests/idris2/reg/reg045/withparams.idr similarity index 100% rename from tests/idris2/reg045/withparams.idr rename to tests/idris2/reg/reg045/withparams.idr diff --git a/tests/idris2/reg046/Postpone.idr b/tests/idris2/reg/reg046/Postpone.idr similarity index 100% rename from tests/idris2/reg046/Postpone.idr rename to tests/idris2/reg/reg046/Postpone.idr diff --git a/tests/idris2/reg046/expected b/tests/idris2/reg/reg046/expected similarity index 100% rename from tests/idris2/reg046/expected rename to tests/idris2/reg/reg046/expected diff --git a/tests/idris2/reg046/input b/tests/idris2/reg/reg046/input similarity index 100% rename from tests/idris2/reg046/input rename to tests/idris2/reg/reg046/input diff --git a/tests/idris2/reg046/run b/tests/idris2/reg/reg046/run similarity index 54% rename from tests/idris2/reg046/run rename to tests/idris2/reg/reg046/run index 1b63d28c3d..606432c5d4 100755 --- a/tests/idris2/reg046/run +++ b/tests/idris2/reg/reg046/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 Postpone.idr < input diff --git a/tests/idris2/reg047/QualifiedDoBang.idr b/tests/idris2/reg/reg047/QualifiedDoBang.idr similarity index 100% rename from tests/idris2/reg047/QualifiedDoBang.idr rename to tests/idris2/reg/reg047/QualifiedDoBang.idr diff --git a/tests/idris2/reg047/expected b/tests/idris2/reg/reg047/expected similarity index 100% rename from tests/idris2/reg047/expected rename to tests/idris2/reg/reg047/expected diff --git a/tests/idris2/reg047/input b/tests/idris2/reg/reg047/input similarity index 100% rename from tests/idris2/reg047/input rename to tests/idris2/reg/reg047/input diff --git a/tests/idris2/reg047/run b/tests/idris2/reg/reg047/run similarity index 60% rename from tests/idris2/reg047/run rename to tests/idris2/reg/reg047/run index d1c2a87cbb..335b0ec338 100755 --- a/tests/idris2/reg047/run +++ b/tests/idris2/reg/reg047/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 QualifiedDoBang.idr < input diff --git a/tests/idris2/reg048/expected b/tests/idris2/reg/reg048/expected similarity index 100% rename from tests/idris2/reg048/expected rename to tests/idris2/reg/reg048/expected diff --git a/tests/idris2/reg048/inferror.idr b/tests/idris2/reg/reg048/inferror.idr similarity index 100% rename from tests/idris2/reg048/inferror.idr rename to tests/idris2/reg/reg048/inferror.idr diff --git a/tests/idris2/reg048/run b/tests/idris2/reg/reg048/run similarity index 55% rename from tests/idris2/reg048/run rename to tests/idris2/reg/reg048/run index d5473efbc4..5bf04de68a 100755 --- a/tests/idris2/reg048/run +++ b/tests/idris2/reg/reg048/run @@ -1,2 +1,2 @@ -. ../../testutils.sh +. ../../../testutils.sh check inferror.idr -p contrib diff --git a/tests/idris2/reg049/expected b/tests/idris2/reg/reg049/expected similarity index 100% rename from tests/idris2/reg049/expected rename to tests/idris2/reg/reg049/expected diff --git a/tests/idris2/reg049/lettype.idr b/tests/idris2/reg/reg049/lettype.idr similarity index 100% rename from tests/idris2/reg049/lettype.idr rename to tests/idris2/reg/reg049/lettype.idr diff --git a/tests/idris2/reg049/run b/tests/idris2/reg/reg049/run similarity index 54% rename from tests/idris2/reg049/run rename to tests/idris2/reg/reg049/run index e5dc521b16..f41a41bb10 100755 --- a/tests/idris2/reg049/run +++ b/tests/idris2/reg/reg049/run @@ -1,2 +1,2 @@ -. ../../testutils.sh +. ../../../testutils.sh check lettype.idr -p contrib diff --git a/tests/idris2/reg050/expected b/tests/idris2/reg/reg050/expected similarity index 100% rename from tests/idris2/reg050/expected rename to tests/idris2/reg/reg050/expected diff --git a/tests/idris2/reg050/loopy.idr b/tests/idris2/reg/reg050/loopy.idr similarity index 100% rename from tests/idris2/reg050/loopy.idr rename to tests/idris2/reg/reg050/loopy.idr diff --git a/tests/idris2/reg050/run b/tests/idris2/reg/reg050/run similarity index 52% rename from tests/idris2/reg050/run rename to tests/idris2/reg/reg050/run index 66707ed52f..c14600be1c 100755 --- a/tests/idris2/reg050/run +++ b/tests/idris2/reg/reg050/run @@ -1,2 +1,2 @@ -. ../../testutils.sh +. ../../../testutils.sh check loopy.idr -p contrib diff --git a/tests/idris2/reg051/BigFins.idr b/tests/idris2/reg/reg051/BigFins.idr similarity index 100% rename from tests/idris2/reg051/BigFins.idr rename to tests/idris2/reg/reg051/BigFins.idr diff --git a/tests/idris2/reg051/expected b/tests/idris2/reg/reg051/expected similarity index 100% rename from tests/idris2/reg051/expected rename to tests/idris2/reg/reg051/expected diff --git a/tests/idris2/reg/reg051/run b/tests/idris2/reg/reg051/run new file mode 100755 index 0000000000..144530cf6a --- /dev/null +++ b/tests/idris2/reg/reg051/run @@ -0,0 +1,2 @@ +. ../../../testutils.sh +check BigFins.idr diff --git a/tests/idris2/reg051/test.ipkg b/tests/idris2/reg/reg051/test.ipkg similarity index 100% rename from tests/idris2/reg051/test.ipkg rename to tests/idris2/reg/reg051/test.ipkg diff --git a/tests/idris2/reg052/DPairQuote.idr b/tests/idris2/reg/reg052/DPairQuote.idr similarity index 100% rename from tests/idris2/reg052/DPairQuote.idr rename to tests/idris2/reg/reg052/DPairQuote.idr diff --git a/tests/idris2/reg052/expected b/tests/idris2/reg/reg052/expected similarity index 100% rename from tests/idris2/reg052/expected rename to tests/idris2/reg/reg052/expected diff --git a/tests/idris2/reg/reg052/run b/tests/idris2/reg/reg052/run new file mode 100755 index 0000000000..5ed96ec7f5 --- /dev/null +++ b/tests/idris2/reg/reg052/run @@ -0,0 +1,2 @@ +. ../../../testutils.sh +check DPairQuote.idr diff --git a/tests/idris2/reg052/test.ipkg b/tests/idris2/reg/reg052/test.ipkg similarity index 100% rename from tests/idris2/reg052/test.ipkg rename to tests/idris2/reg/reg052/test.ipkg diff --git a/tests/idris2/reg001/run b/tests/idris2/reg001/run deleted file mode 100755 index e927e12c0b..0000000000 --- a/tests/idris2/reg001/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check D.idr diff --git a/tests/idris2/reg002/run b/tests/idris2/reg002/run deleted file mode 100755 index c2084f482a..0000000000 --- a/tests/idris2/reg002/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check linm.idr diff --git a/tests/idris2/reg004/run b/tests/idris2/reg004/run deleted file mode 100755 index 10632b2def..0000000000 --- a/tests/idris2/reg004/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check ambig.idr diff --git a/tests/idris2/reg005/run b/tests/idris2/reg005/run deleted file mode 100755 index 6d1af86408..0000000000 --- a/tests/idris2/reg005/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check iftype.idr diff --git a/tests/idris2/reg006/run b/tests/idris2/reg006/run deleted file mode 100755 index 05d093b40b..0000000000 --- a/tests/idris2/reg006/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check Cmd.idr diff --git a/tests/idris2/reg007/run b/tests/idris2/reg007/run deleted file mode 100755 index 64c80a5b21..0000000000 --- a/tests/idris2/reg007/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check Main.idr diff --git a/tests/idris2/reg009/run b/tests/idris2/reg009/run deleted file mode 100755 index 190bd11b91..0000000000 --- a/tests/idris2/reg009/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check Case.idr diff --git a/tests/idris2/reg010/run b/tests/idris2/reg010/run deleted file mode 100755 index 33f2d1ada4..0000000000 --- a/tests/idris2/reg010/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check Recordname.idr diff --git a/tests/idris2/reg011/run b/tests/idris2/reg011/run deleted file mode 100755 index ce2877d08b..0000000000 --- a/tests/idris2/reg011/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check mut.idr diff --git a/tests/idris2/reg012/run b/tests/idris2/reg012/run deleted file mode 100755 index 8b984ed2bc..0000000000 --- a/tests/idris2/reg012/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check Foo.idr diff --git a/tests/idris2/reg015/run b/tests/idris2/reg015/run deleted file mode 100755 index 581897bbad..0000000000 --- a/tests/idris2/reg015/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check anyfail.idr diff --git a/tests/idris2/reg016/run b/tests/idris2/reg016/run deleted file mode 100755 index 29afa74be8..0000000000 --- a/tests/idris2/reg016/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check Using.idr diff --git a/tests/idris2/reg017/run b/tests/idris2/reg017/run deleted file mode 100755 index 5465fa2dfc..0000000000 --- a/tests/idris2/reg017/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check lammult.idr diff --git a/tests/idris2/reg018/run b/tests/idris2/reg018/run deleted file mode 100755 index 508b8282e7..0000000000 --- a/tests/idris2/reg018/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check cycle.idr diff --git a/tests/idris2/reg019/run b/tests/idris2/reg019/run deleted file mode 100755 index 546e07fcd2..0000000000 --- a/tests/idris2/reg019/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check lazybug.idr diff --git a/tests/idris2/reg021/run b/tests/idris2/reg021/run deleted file mode 100755 index 6f7f117224..0000000000 --- a/tests/idris2/reg021/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check case.idr diff --git a/tests/idris2/reg022/run b/tests/idris2/reg022/run deleted file mode 100755 index 6f7f117224..0000000000 --- a/tests/idris2/reg022/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check case.idr diff --git a/tests/idris2/reg023/run b/tests/idris2/reg023/run deleted file mode 100755 index 3f140f93b2..0000000000 --- a/tests/idris2/reg023/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check boom.idr diff --git a/tests/idris2/reg026/run b/tests/idris2/reg026/run deleted file mode 100755 index f5d5dd5d96..0000000000 --- a/tests/idris2/reg026/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check Meh.idr diff --git a/tests/idris2/reg027/run b/tests/idris2/reg027/run deleted file mode 100755 index 4d6e1ff77b..0000000000 --- a/tests/idris2/reg027/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check pwhere.idr diff --git a/tests/idris2/reg028/run b/tests/idris2/reg028/run deleted file mode 100755 index 64b10e3f8e..0000000000 --- a/tests/idris2/reg028/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check Test.idr diff --git a/tests/idris2/reg029/run b/tests/idris2/reg029/run deleted file mode 100755 index fcbc83b59b..0000000000 --- a/tests/idris2/reg029/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check lqueue.idr diff --git a/tests/idris2/reg030/run b/tests/idris2/reg030/run deleted file mode 100755 index e8267b3ab2..0000000000 --- a/tests/idris2/reg030/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check A.idr diff --git a/tests/idris2/reg031/run b/tests/idris2/reg031/run deleted file mode 100755 index a08b0bef9c..0000000000 --- a/tests/idris2/reg031/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check dpair.idr diff --git a/tests/idris2/reg032/run b/tests/idris2/reg032/run deleted file mode 100755 index 72cff3a29c..0000000000 --- a/tests/idris2/reg032/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check recupdate.idr diff --git a/tests/idris2/reg033/run b/tests/idris2/reg033/run deleted file mode 100755 index d0f513a107..0000000000 --- a/tests/idris2/reg033/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check test.idr diff --git a/tests/idris2/reg034/run b/tests/idris2/reg034/run deleted file mode 100755 index 1092d1eb15..0000000000 --- a/tests/idris2/reg034/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check void.idr diff --git a/tests/idris2/reg035/run b/tests/idris2/reg035/run deleted file mode 100755 index 096de90fd3..0000000000 --- a/tests/idris2/reg035/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check Implicit.idr diff --git a/tests/idris2/reg036/run b/tests/idris2/reg036/run deleted file mode 100755 index 64b10e3f8e..0000000000 --- a/tests/idris2/reg036/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check Test.idr diff --git a/tests/idris2/reg037/run b/tests/idris2/reg037/run deleted file mode 100755 index 64b10e3f8e..0000000000 --- a/tests/idris2/reg037/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check Test.idr diff --git a/tests/idris2/reg039/run b/tests/idris2/reg039/run deleted file mode 100755 index 52897db3e3..0000000000 --- a/tests/idris2/reg039/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check dupdup.idr diff --git a/tests/idris2/reg040/run b/tests/idris2/reg040/run deleted file mode 100755 index 4c3a475e32..0000000000 --- a/tests/idris2/reg040/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check CoverBug.idr diff --git a/tests/idris2/reg041/run b/tests/idris2/reg041/run deleted file mode 100755 index 00fb744d9e..0000000000 --- a/tests/idris2/reg041/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check tuple.idr diff --git a/tests/idris2/reg044/run b/tests/idris2/reg044/run deleted file mode 100755 index d459abd54a..0000000000 --- a/tests/idris2/reg044/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check Methods.idr diff --git a/tests/idris2/reg045/run b/tests/idris2/reg045/run deleted file mode 100755 index 8426db0e77..0000000000 --- a/tests/idris2/reg045/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check withparams.idr diff --git a/tests/idris2/reg051/run b/tests/idris2/reg051/run deleted file mode 100755 index 2b332393bf..0000000000 --- a/tests/idris2/reg051/run +++ /dev/null @@ -1,2 +0,0 @@ -. ../../testutils.sh -check BigFins.idr diff --git a/tests/idris2/reg052/run b/tests/idris2/reg052/run deleted file mode 100755 index bdd7a10f4e..0000000000 --- a/tests/idris2/reg052/run +++ /dev/null @@ -1,2 +0,0 @@ -. ../../testutils.sh -check DPairQuote.idr diff --git a/tests/idris2/repl001/expected b/tests/idris2/repl/repl001/expected similarity index 100% rename from tests/idris2/repl001/expected rename to tests/idris2/repl/repl001/expected diff --git a/tests/idris2/repl001/input b/tests/idris2/repl/repl001/input similarity index 100% rename from tests/idris2/repl001/input rename to tests/idris2/repl/repl001/input diff --git a/tests/idris2/repl/repl001/run b/tests/idris2/repl/repl001/run new file mode 100755 index 0000000000..404d7d47c6 --- /dev/null +++ b/tests/idris2/repl/repl001/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +idris2 < input diff --git a/tests/idris2/repl002/expected b/tests/idris2/repl/repl002/expected similarity index 100% rename from tests/idris2/repl002/expected rename to tests/idris2/repl/repl002/expected diff --git a/tests/idris2/repl002/input b/tests/idris2/repl/repl002/input similarity index 100% rename from tests/idris2/repl002/input rename to tests/idris2/repl/repl002/input diff --git a/tests/idris2/repl/repl002/run b/tests/idris2/repl/repl002/run new file mode 100755 index 0000000000..404d7d47c6 --- /dev/null +++ b/tests/idris2/repl/repl002/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +idris2 < input diff --git a/tests/idris2/repl003/expected b/tests/idris2/repl/repl003/expected similarity index 100% rename from tests/idris2/repl003/expected rename to tests/idris2/repl/repl003/expected diff --git a/tests/idris2/repl003/input b/tests/idris2/repl/repl003/input similarity index 100% rename from tests/idris2/repl003/input rename to tests/idris2/repl/repl003/input diff --git a/tests/idris2/repl/repl003/run b/tests/idris2/repl/repl003/run new file mode 100755 index 0000000000..404d7d47c6 --- /dev/null +++ b/tests/idris2/repl/repl003/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +idris2 < input diff --git a/tests/idris2/repl004/expected b/tests/idris2/repl/repl004/expected similarity index 100% rename from tests/idris2/repl004/expected rename to tests/idris2/repl/repl004/expected diff --git a/tests/idris2/repl004/input b/tests/idris2/repl/repl004/input similarity index 100% rename from tests/idris2/repl004/input rename to tests/idris2/repl/repl004/input diff --git a/tests/idris2/repl/repl004/run b/tests/idris2/repl/repl004/run new file mode 100755 index 0000000000..404d7d47c6 --- /dev/null +++ b/tests/idris2/repl/repl004/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +idris2 < input diff --git a/tests/idris2/repl005/expected b/tests/idris2/repl/repl005/expected similarity index 100% rename from tests/idris2/repl005/expected rename to tests/idris2/repl/repl005/expected diff --git a/tests/idris2/repl005/input b/tests/idris2/repl/repl005/input similarity index 100% rename from tests/idris2/repl005/input rename to tests/idris2/repl/repl005/input diff --git a/tests/idris2/repl/repl005/run b/tests/idris2/repl/repl005/run new file mode 100755 index 0000000000..404d7d47c6 --- /dev/null +++ b/tests/idris2/repl/repl005/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +idris2 < input diff --git a/tests/idris2/repl006/expected b/tests/idris2/repl/repl006/expected similarity index 100% rename from tests/idris2/repl006/expected rename to tests/idris2/repl/repl006/expected diff --git a/tests/idris2/repl006/input b/tests/idris2/repl/repl006/input similarity index 100% rename from tests/idris2/repl006/input rename to tests/idris2/repl/repl006/input diff --git a/tests/idris2/repl/repl006/run b/tests/idris2/repl/repl006/run new file mode 100755 index 0000000000..404d7d47c6 --- /dev/null +++ b/tests/idris2/repl/repl006/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +idris2 < input diff --git a/tests/idris2/repl001/run b/tests/idris2/repl001/run deleted file mode 100755 index fecfb4b107..0000000000 --- a/tests/idris2/repl001/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -idris2 < input diff --git a/tests/idris2/repl002/run b/tests/idris2/repl002/run deleted file mode 100755 index fecfb4b107..0000000000 --- a/tests/idris2/repl002/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -idris2 < input diff --git a/tests/idris2/repl003/run b/tests/idris2/repl003/run deleted file mode 100755 index fecfb4b107..0000000000 --- a/tests/idris2/repl003/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -idris2 < input diff --git a/tests/idris2/repl004/run b/tests/idris2/repl004/run deleted file mode 100755 index fecfb4b107..0000000000 --- a/tests/idris2/repl004/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -idris2 < input diff --git a/tests/idris2/repl005/run b/tests/idris2/repl005/run deleted file mode 100755 index fecfb4b107..0000000000 --- a/tests/idris2/repl005/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -idris2 < input diff --git a/tests/idris2/repl006/run b/tests/idris2/repl006/run deleted file mode 100755 index fecfb4b107..0000000000 --- a/tests/idris2/repl006/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -idris2 < input diff --git a/tests/idris2/rewrite001/run b/tests/idris2/rewrite001/run deleted file mode 100755 index 48e0430d41..0000000000 --- a/tests/idris2/rewrite001/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check Issue2573.idr diff --git a/tests/idris2/schemeeval001/expected b/tests/idris2/schemeeval/schemeeval001/expected similarity index 100% rename from tests/idris2/schemeeval001/expected rename to tests/idris2/schemeeval/schemeeval001/expected diff --git a/tests/idris2/schemeeval001/input b/tests/idris2/schemeeval/schemeeval001/input similarity index 100% rename from tests/idris2/schemeeval001/input rename to tests/idris2/schemeeval/schemeeval001/input diff --git a/tests/idris2/schemeeval/schemeeval001/run b/tests/idris2/schemeeval/schemeeval001/run new file mode 100755 index 0000000000..404d7d47c6 --- /dev/null +++ b/tests/idris2/schemeeval/schemeeval001/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +idris2 < input diff --git a/tests/idris2/schemeeval002/expected b/tests/idris2/schemeeval/schemeeval002/expected similarity index 100% rename from tests/idris2/schemeeval002/expected rename to tests/idris2/schemeeval/schemeeval002/expected diff --git a/tests/idris2/schemeeval002/input b/tests/idris2/schemeeval/schemeeval002/input similarity index 100% rename from tests/idris2/schemeeval002/input rename to tests/idris2/schemeeval/schemeeval002/input diff --git a/tests/idris2/schemeeval/schemeeval002/run b/tests/idris2/schemeeval/schemeeval002/run new file mode 100755 index 0000000000..404d7d47c6 --- /dev/null +++ b/tests/idris2/schemeeval/schemeeval002/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +idris2 < input diff --git a/tests/idris2/schemeeval003/expected b/tests/idris2/schemeeval/schemeeval003/expected similarity index 100% rename from tests/idris2/schemeeval003/expected rename to tests/idris2/schemeeval/schemeeval003/expected diff --git a/tests/idris2/schemeeval003/input b/tests/idris2/schemeeval/schemeeval003/input similarity index 100% rename from tests/idris2/schemeeval003/input rename to tests/idris2/schemeeval/schemeeval003/input diff --git a/tests/idris2/schemeeval/schemeeval003/run b/tests/idris2/schemeeval/schemeeval003/run new file mode 100755 index 0000000000..404d7d47c6 --- /dev/null +++ b/tests/idris2/schemeeval/schemeeval003/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +idris2 < input diff --git a/tests/idris2/schemeeval004/expected b/tests/idris2/schemeeval/schemeeval004/expected similarity index 100% rename from tests/idris2/schemeeval004/expected rename to tests/idris2/schemeeval/schemeeval004/expected diff --git a/tests/idris2/schemeeval004/input b/tests/idris2/schemeeval/schemeeval004/input similarity index 100% rename from tests/idris2/schemeeval004/input rename to tests/idris2/schemeeval/schemeeval004/input diff --git a/tests/idris2/schemeeval004/list.idr b/tests/idris2/schemeeval/schemeeval004/list.idr similarity index 100% rename from tests/idris2/schemeeval004/list.idr rename to tests/idris2/schemeeval/schemeeval004/list.idr diff --git a/tests/idris2/schemeeval004/run b/tests/idris2/schemeeval/schemeeval004/run similarity index 51% rename from tests/idris2/schemeeval004/run rename to tests/idris2/schemeeval/schemeeval004/run index 6d6d987b11..11a21eb370 100755 --- a/tests/idris2/schemeeval004/run +++ b/tests/idris2/schemeeval/schemeeval004/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 list.idr < input diff --git a/tests/idris2/schemeeval005/Printf.idr b/tests/idris2/schemeeval/schemeeval005/Printf.idr similarity index 100% rename from tests/idris2/schemeeval005/Printf.idr rename to tests/idris2/schemeeval/schemeeval005/Printf.idr diff --git a/tests/idris2/schemeeval005/expected b/tests/idris2/schemeeval/schemeeval005/expected similarity index 100% rename from tests/idris2/schemeeval005/expected rename to tests/idris2/schemeeval/schemeeval005/expected diff --git a/tests/idris2/schemeeval005/input b/tests/idris2/schemeeval/schemeeval005/input similarity index 100% rename from tests/idris2/schemeeval005/input rename to tests/idris2/schemeeval/schemeeval005/input diff --git a/tests/idris2/schemeeval005/run b/tests/idris2/schemeeval/schemeeval005/run similarity index 52% rename from tests/idris2/schemeeval005/run rename to tests/idris2/schemeeval/schemeeval005/run index a854f30fd6..bc2f18c94b 100755 --- a/tests/idris2/schemeeval005/run +++ b/tests/idris2/schemeeval/schemeeval005/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 Printf.idr < input diff --git a/tests/idris2/schemeeval006/expected b/tests/idris2/schemeeval/schemeeval006/expected similarity index 100% rename from tests/idris2/schemeeval006/expected rename to tests/idris2/schemeeval/schemeeval006/expected diff --git a/tests/idris2/schemeeval006/input b/tests/idris2/schemeeval/schemeeval006/input similarity index 100% rename from tests/idris2/schemeeval006/input rename to tests/idris2/schemeeval/schemeeval006/input diff --git a/tests/idris2/schemeeval/schemeeval006/run b/tests/idris2/schemeeval/schemeeval006/run new file mode 100644 index 0000000000..404d7d47c6 --- /dev/null +++ b/tests/idris2/schemeeval/schemeeval006/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +idris2 < input diff --git a/tests/idris2/schemeeval001/run b/tests/idris2/schemeeval001/run deleted file mode 100755 index fecfb4b107..0000000000 --- a/tests/idris2/schemeeval001/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -idris2 < input diff --git a/tests/idris2/schemeeval002/run b/tests/idris2/schemeeval002/run deleted file mode 100755 index fecfb4b107..0000000000 --- a/tests/idris2/schemeeval002/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -idris2 < input diff --git a/tests/idris2/schemeeval003/run b/tests/idris2/schemeeval003/run deleted file mode 100755 index fecfb4b107..0000000000 --- a/tests/idris2/schemeeval003/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -idris2 < input diff --git a/tests/idris2/schemeeval006/run b/tests/idris2/schemeeval006/run deleted file mode 100644 index fecfb4b107..0000000000 --- a/tests/idris2/schemeeval006/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -idris2 < input diff --git a/tests/idris2/termination001/AgdaIssue6059.idr b/tests/idris2/termination/termination001/AgdaIssue6059.idr similarity index 100% rename from tests/idris2/termination001/AgdaIssue6059.idr rename to tests/idris2/termination/termination001/AgdaIssue6059.idr diff --git a/tests/idris2/termination001/expected b/tests/idris2/termination/termination001/expected similarity index 100% rename from tests/idris2/termination001/expected rename to tests/idris2/termination/termination001/expected diff --git a/tests/idris2/termination001/run b/tests/idris2/termination/termination001/run similarity index 54% rename from tests/idris2/termination001/run rename to tests/idris2/termination/termination001/run index a035eb51c1..f60e0a73a8 100755 --- a/tests/idris2/termination001/run +++ b/tests/idris2/termination/termination001/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 -c AgdaIssue6059.idr diff --git a/tests/idris2/positivity001/Issue660.idr b/tests/idris2/total/positivity001/Issue660.idr similarity index 100% rename from tests/idris2/positivity001/Issue660.idr rename to tests/idris2/total/positivity001/Issue660.idr diff --git a/tests/idris2/positivity001/expected b/tests/idris2/total/positivity001/expected similarity index 100% rename from tests/idris2/positivity001/expected rename to tests/idris2/total/positivity001/expected diff --git a/tests/idris2/total/positivity001/run b/tests/idris2/total/positivity001/run new file mode 100644 index 0000000000..a1d32c7cb8 --- /dev/null +++ b/tests/idris2/total/positivity001/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Issue660.idr diff --git a/tests/idris2/positivity002/Issue660.idr b/tests/idris2/total/positivity002/Issue660.idr similarity index 100% rename from tests/idris2/positivity002/Issue660.idr rename to tests/idris2/total/positivity002/Issue660.idr diff --git a/tests/idris2/positivity002/expected b/tests/idris2/total/positivity002/expected similarity index 100% rename from tests/idris2/positivity002/expected rename to tests/idris2/total/positivity002/expected diff --git a/tests/idris2/total/positivity002/run b/tests/idris2/total/positivity002/run new file mode 100644 index 0000000000..a1d32c7cb8 --- /dev/null +++ b/tests/idris2/total/positivity002/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Issue660.idr diff --git a/tests/idris2/positivity003/Issue660.idr b/tests/idris2/total/positivity003/Issue660.idr similarity index 100% rename from tests/idris2/positivity003/Issue660.idr rename to tests/idris2/total/positivity003/Issue660.idr diff --git a/tests/idris2/positivity003/expected b/tests/idris2/total/positivity003/expected similarity index 100% rename from tests/idris2/positivity003/expected rename to tests/idris2/total/positivity003/expected diff --git a/tests/idris2/total/positivity003/run b/tests/idris2/total/positivity003/run new file mode 100644 index 0000000000..a1d32c7cb8 --- /dev/null +++ b/tests/idris2/total/positivity003/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Issue660.idr diff --git a/tests/idris2/positivity004/Issue1771-1.idr b/tests/idris2/total/positivity004/Issue1771-1.idr similarity index 100% rename from tests/idris2/positivity004/Issue1771-1.idr rename to tests/idris2/total/positivity004/Issue1771-1.idr diff --git a/tests/idris2/positivity004/Issue1771-2.idr b/tests/idris2/total/positivity004/Issue1771-2.idr similarity index 100% rename from tests/idris2/positivity004/Issue1771-2.idr rename to tests/idris2/total/positivity004/Issue1771-2.idr diff --git a/tests/idris2/positivity004/Issue1771-3.idr b/tests/idris2/total/positivity004/Issue1771-3.idr similarity index 100% rename from tests/idris2/positivity004/Issue1771-3.idr rename to tests/idris2/total/positivity004/Issue1771-3.idr diff --git a/tests/idris2/positivity004/expected b/tests/idris2/total/positivity004/expected similarity index 100% rename from tests/idris2/positivity004/expected rename to tests/idris2/total/positivity004/expected diff --git a/tests/idris2/positivity004/run b/tests/idris2/total/positivity004/run similarity index 79% rename from tests/idris2/positivity004/run rename to tests/idris2/total/positivity004/run index 274b479254..0785bd6655 100644 --- a/tests/idris2/positivity004/run +++ b/tests/idris2/total/positivity004/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh check Issue1771-1.idr || true check Issue1771-2.idr || true diff --git a/tests/idris2/total001/Total.idr b/tests/idris2/total/total001/Total.idr similarity index 100% rename from tests/idris2/total001/Total.idr rename to tests/idris2/total/total001/Total.idr diff --git a/tests/idris2/total001/expected b/tests/idris2/total/total001/expected similarity index 100% rename from tests/idris2/total001/expected rename to tests/idris2/total/total001/expected diff --git a/tests/idris2/total001/input b/tests/idris2/total/total001/input similarity index 100% rename from tests/idris2/total001/input rename to tests/idris2/total/total001/input diff --git a/tests/idris2/total001/run b/tests/idris2/total/total001/run similarity index 52% rename from tests/idris2/total001/run rename to tests/idris2/total/total001/run index c1204f255d..e3074a9579 100755 --- a/tests/idris2/total001/run +++ b/tests/idris2/total/total001/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 Total.idr < input diff --git a/tests/idris2/total002/Total.idr b/tests/idris2/total/total002/Total.idr similarity index 100% rename from tests/idris2/total002/Total.idr rename to tests/idris2/total/total002/Total.idr diff --git a/tests/idris2/total002/expected b/tests/idris2/total/total002/expected similarity index 100% rename from tests/idris2/total002/expected rename to tests/idris2/total/total002/expected diff --git a/tests/idris2/total002/input b/tests/idris2/total/total002/input similarity index 100% rename from tests/idris2/total002/input rename to tests/idris2/total/total002/input diff --git a/tests/idris2/total002/run b/tests/idris2/total/total002/run similarity index 52% rename from tests/idris2/total002/run rename to tests/idris2/total/total002/run index c1204f255d..e3074a9579 100755 --- a/tests/idris2/total002/run +++ b/tests/idris2/total/total002/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 Total.idr < input diff --git a/tests/idris2/total003/Total.idr b/tests/idris2/total/total003/Total.idr similarity index 100% rename from tests/idris2/total003/Total.idr rename to tests/idris2/total/total003/Total.idr diff --git a/tests/idris2/total003/expected b/tests/idris2/total/total003/expected similarity index 100% rename from tests/idris2/total003/expected rename to tests/idris2/total/total003/expected diff --git a/tests/idris2/total003/input b/tests/idris2/total/total003/input similarity index 100% rename from tests/idris2/total003/input rename to tests/idris2/total/total003/input diff --git a/tests/idris2/total003/run b/tests/idris2/total/total003/run similarity index 52% rename from tests/idris2/total003/run rename to tests/idris2/total/total003/run index c1204f255d..e3074a9579 100755 --- a/tests/idris2/total003/run +++ b/tests/idris2/total/total003/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 Total.idr < input diff --git a/tests/idris2/total004/Total.idr b/tests/idris2/total/total004/Total.idr similarity index 100% rename from tests/idris2/total004/Total.idr rename to tests/idris2/total/total004/Total.idr diff --git a/tests/idris2/total004/expected b/tests/idris2/total/total004/expected similarity index 100% rename from tests/idris2/total004/expected rename to tests/idris2/total/total004/expected diff --git a/tests/idris2/total004/input b/tests/idris2/total/total004/input similarity index 100% rename from tests/idris2/total004/input rename to tests/idris2/total/total004/input diff --git a/tests/idris2/total004/run b/tests/idris2/total/total004/run similarity index 52% rename from tests/idris2/total004/run rename to tests/idris2/total/total004/run index c1204f255d..e3074a9579 100755 --- a/tests/idris2/total004/run +++ b/tests/idris2/total/total004/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 Total.idr < input diff --git a/tests/idris2/total005/Total.idr b/tests/idris2/total/total005/Total.idr similarity index 100% rename from tests/idris2/total005/Total.idr rename to tests/idris2/total/total005/Total.idr diff --git a/tests/idris2/total005/expected b/tests/idris2/total/total005/expected similarity index 100% rename from tests/idris2/total005/expected rename to tests/idris2/total/total005/expected diff --git a/tests/idris2/total005/input b/tests/idris2/total/total005/input similarity index 100% rename from tests/idris2/total005/input rename to tests/idris2/total/total005/input diff --git a/tests/idris2/total/total005/run b/tests/idris2/total/total005/run new file mode 100755 index 0000000000..e3074a9579 --- /dev/null +++ b/tests/idris2/total/total005/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +idris2 Total.idr < input diff --git a/tests/idris2/total006/Total.idr b/tests/idris2/total/total006/Total.idr similarity index 100% rename from tests/idris2/total006/Total.idr rename to tests/idris2/total/total006/Total.idr diff --git a/tests/idris2/total006/expected b/tests/idris2/total/total006/expected similarity index 100% rename from tests/idris2/total006/expected rename to tests/idris2/total/total006/expected diff --git a/tests/idris2/total006/input b/tests/idris2/total/total006/input similarity index 100% rename from tests/idris2/total006/input rename to tests/idris2/total/total006/input diff --git a/tests/idris2/total/total006/run b/tests/idris2/total/total006/run new file mode 100755 index 0000000000..e3074a9579 --- /dev/null +++ b/tests/idris2/total/total006/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +idris2 Total.idr < input diff --git a/tests/idris2/total007/expected b/tests/idris2/total/total007/expected similarity index 100% rename from tests/idris2/total007/expected rename to tests/idris2/total/total007/expected diff --git a/tests/idris2/total007/partial.idr b/tests/idris2/total/total007/partial.idr similarity index 100% rename from tests/idris2/total007/partial.idr rename to tests/idris2/total/total007/partial.idr diff --git a/tests/idris2/total/total007/run b/tests/idris2/total/total007/run new file mode 100755 index 0000000000..ad76332274 --- /dev/null +++ b/tests/idris2/total/total007/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check partial.idr diff --git a/tests/idris2/total008/expected b/tests/idris2/total/total008/expected similarity index 100% rename from tests/idris2/total008/expected rename to tests/idris2/total/total008/expected diff --git a/tests/idris2/total008/partial.idr b/tests/idris2/total/total008/partial.idr similarity index 100% rename from tests/idris2/total008/partial.idr rename to tests/idris2/total/total008/partial.idr diff --git a/tests/idris2/total/total008/run b/tests/idris2/total/total008/run new file mode 100755 index 0000000000..ad76332274 --- /dev/null +++ b/tests/idris2/total/total008/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check partial.idr diff --git a/tests/idris2/total009/expected b/tests/idris2/total/total009/expected similarity index 100% rename from tests/idris2/total009/expected rename to tests/idris2/total/total009/expected diff --git a/tests/idris2/total/total009/run b/tests/idris2/total/total009/run new file mode 100755 index 0000000000..57faa6eb90 --- /dev/null +++ b/tests/idris2/total/total009/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check tree.idr diff --git a/tests/idris2/total009/tree.idr b/tests/idris2/total/total009/tree.idr similarity index 100% rename from tests/idris2/total009/tree.idr rename to tests/idris2/total/total009/tree.idr diff --git a/tests/idris2/total010/PartialWith.idr b/tests/idris2/total/total010/PartialWith.idr similarity index 100% rename from tests/idris2/total010/PartialWith.idr rename to tests/idris2/total/total010/PartialWith.idr diff --git a/tests/idris2/total010/expected b/tests/idris2/total/total010/expected similarity index 100% rename from tests/idris2/total010/expected rename to tests/idris2/total/total010/expected diff --git a/tests/idris2/total/total010/run b/tests/idris2/total/total010/run new file mode 100755 index 0000000000..424ebf7a19 --- /dev/null +++ b/tests/idris2/total/total010/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check PartialWith.idr diff --git a/tests/idris2/total011/Issue1460.idr b/tests/idris2/total/total011/Issue1460.idr similarity index 100% rename from tests/idris2/total011/Issue1460.idr rename to tests/idris2/total/total011/Issue1460.idr diff --git a/tests/idris2/total011/Issue1782.idr b/tests/idris2/total/total011/Issue1782.idr similarity index 100% rename from tests/idris2/total011/Issue1782.idr rename to tests/idris2/total/total011/Issue1782.idr diff --git a/tests/idris2/total011/Issue1828.idr b/tests/idris2/total/total011/Issue1828.idr similarity index 100% rename from tests/idris2/total011/Issue1828.idr rename to tests/idris2/total/total011/Issue1828.idr diff --git a/tests/idris2/total011/Issue1859-2.idr b/tests/idris2/total/total011/Issue1859-2.idr similarity index 100% rename from tests/idris2/total011/Issue1859-2.idr rename to tests/idris2/total/total011/Issue1859-2.idr diff --git a/tests/idris2/total011/Issue1859.idr b/tests/idris2/total/total011/Issue1859.idr similarity index 100% rename from tests/idris2/total011/Issue1859.idr rename to tests/idris2/total/total011/Issue1859.idr diff --git a/tests/idris2/total011/expected b/tests/idris2/total/total011/expected similarity index 100% rename from tests/idris2/total011/expected rename to tests/idris2/total/total011/expected diff --git a/tests/idris2/total011/run b/tests/idris2/total/total011/run similarity index 81% rename from tests/idris2/total011/run rename to tests/idris2/total/total011/run index fa22291fc8..12281e64b6 100755 --- a/tests/idris2/total011/run +++ b/tests/idris2/total/total011/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh check Issue1782.idr check Issue1460.idr diff --git a/tests/idris2/total012/Issue1828.idr b/tests/idris2/total/total012/Issue1828.idr similarity index 100% rename from tests/idris2/total012/Issue1828.idr rename to tests/idris2/total/total012/Issue1828.idr diff --git a/tests/idris2/total012/TotallyTotal.idr b/tests/idris2/total/total012/TotallyTotal.idr similarity index 100% rename from tests/idris2/total012/TotallyTotal.idr rename to tests/idris2/total/total012/TotallyTotal.idr diff --git a/tests/idris2/total012/expected b/tests/idris2/total/total012/expected similarity index 100% rename from tests/idris2/total012/expected rename to tests/idris2/total/total012/expected diff --git a/tests/idris2/total012/run b/tests/idris2/total/total012/run similarity index 94% rename from tests/idris2/total012/run rename to tests/idris2/total/total012/run index aa40ac525d..5bb43bced0 100755 --- a/tests/idris2/total012/run +++ b/tests/idris2/total/total012/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh check Issue1828.idr check Issue1828.idr --total --log totality.requirement:20 | sed -r "s/.([0-9]){10}//g" diff --git a/tests/idris2/total013/Issue1404.idr b/tests/idris2/total/total013/Issue1404.idr similarity index 100% rename from tests/idris2/total013/Issue1404.idr rename to tests/idris2/total/total013/Issue1404.idr diff --git a/tests/idris2/total013/expected b/tests/idris2/total/total013/expected similarity index 100% rename from tests/idris2/total013/expected rename to tests/idris2/total/total013/expected diff --git a/tests/idris2/total/total013/run b/tests/idris2/total/total013/run new file mode 100755 index 0000000000..066473c9ee --- /dev/null +++ b/tests/idris2/total/total013/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Issue1404.idr diff --git a/tests/idris2/total014/FunCompTC.idr b/tests/idris2/total/total014/FunCompTC.idr similarity index 100% rename from tests/idris2/total014/FunCompTC.idr rename to tests/idris2/total/total014/FunCompTC.idr diff --git a/tests/idris2/total014/expected b/tests/idris2/total/total014/expected similarity index 100% rename from tests/idris2/total014/expected rename to tests/idris2/total/total014/expected diff --git a/tests/idris2/total/total014/run b/tests/idris2/total/total014/run new file mode 100755 index 0000000000..b015edfa31 --- /dev/null +++ b/tests/idris2/total/total014/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check FunCompTC.idr diff --git a/tests/idris2/total015/CoveringData.idr b/tests/idris2/total/total015/CoveringData.idr similarity index 100% rename from tests/idris2/total015/CoveringData.idr rename to tests/idris2/total/total015/CoveringData.idr diff --git a/tests/idris2/total015/expected b/tests/idris2/total/total015/expected similarity index 100% rename from tests/idris2/total015/expected rename to tests/idris2/total/total015/expected diff --git a/tests/idris2/total015/run b/tests/idris2/total/total015/run similarity index 50% rename from tests/idris2/total015/run rename to tests/idris2/total/total015/run index c2aaa1a2b4..739009c897 100755 --- a/tests/idris2/total015/run +++ b/tests/idris2/total/total015/run @@ -1,3 +1,3 @@ -. ../../testutils.sh +. ../../../testutils.sh check CoveringData.idr diff --git a/tests/idris2/total016/AssertPositivity.idr b/tests/idris2/total/total016/AssertPositivity.idr similarity index 100% rename from tests/idris2/total016/AssertPositivity.idr rename to tests/idris2/total/total016/AssertPositivity.idr diff --git a/tests/idris2/total016/LazyPositivityCheck.idr b/tests/idris2/total/total016/LazyPositivityCheck.idr similarity index 100% rename from tests/idris2/total016/LazyPositivityCheck.idr rename to tests/idris2/total/total016/LazyPositivityCheck.idr diff --git a/tests/idris2/total016/expected b/tests/idris2/total/total016/expected similarity index 100% rename from tests/idris2/total016/expected rename to tests/idris2/total/total016/expected diff --git a/tests/idris2/total016/run b/tests/idris2/total/total016/run similarity index 70% rename from tests/idris2/total016/run rename to tests/idris2/total/total016/run index f7a0e2fe14..2fe9a91463 100755 --- a/tests/idris2/total016/run +++ b/tests/idris2/total/total016/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh check AssertPositivity.idr check LazyPositivityCheck.idr diff --git a/tests/idris2/total017/Paper.idr b/tests/idris2/total/total017/Paper.idr similarity index 100% rename from tests/idris2/total017/Paper.idr rename to tests/idris2/total/total017/Paper.idr diff --git a/tests/idris2/total017/expected b/tests/idris2/total/total017/expected similarity index 100% rename from tests/idris2/total017/expected rename to tests/idris2/total/total017/expected diff --git a/tests/idris2/total/total017/run b/tests/idris2/total/total017/run new file mode 100755 index 0000000000..d09c9a82df --- /dev/null +++ b/tests/idris2/total/total017/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Paper.idr diff --git a/tests/idris2/total018/Issue2448.idr b/tests/idris2/total/total018/Issue2448.idr similarity index 100% rename from tests/idris2/total018/Issue2448.idr rename to tests/idris2/total/total018/Issue2448.idr diff --git a/tests/idris2/total018/expected b/tests/idris2/total/total018/expected similarity index 100% rename from tests/idris2/total018/expected rename to tests/idris2/total/total018/expected diff --git a/tests/idris2/total/total018/run b/tests/idris2/total/total018/run new file mode 100755 index 0000000000..4e601ee26c --- /dev/null +++ b/tests/idris2/total/total018/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Issue2448.idr diff --git a/tests/idris2/total019/Check.idr b/tests/idris2/total/total019/Check.idr similarity index 100% rename from tests/idris2/total019/Check.idr rename to tests/idris2/total/total019/Check.idr diff --git a/tests/idris2/total019/expected b/tests/idris2/total/total019/expected similarity index 100% rename from tests/idris2/total019/expected rename to tests/idris2/total/total019/expected diff --git a/tests/idris2/total/total019/run b/tests/idris2/total/total019/run new file mode 100755 index 0000000000..7b512c7449 --- /dev/null +++ b/tests/idris2/total/total019/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Check.idr diff --git a/tests/idris2/total020/Check.idr b/tests/idris2/total/total020/Check.idr similarity index 100% rename from tests/idris2/total020/Check.idr rename to tests/idris2/total/total020/Check.idr diff --git a/tests/idris2/total020/expected b/tests/idris2/total/total020/expected similarity index 100% rename from tests/idris2/total020/expected rename to tests/idris2/total/total020/expected diff --git a/tests/idris2/total/total020/run b/tests/idris2/total/total020/run new file mode 100755 index 0000000000..7b512c7449 --- /dev/null +++ b/tests/idris2/total/total020/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Check.idr diff --git a/tests/idris2/total020/test.ipkg b/tests/idris2/total/total020/test.ipkg similarity index 100% rename from tests/idris2/total020/test.ipkg rename to tests/idris2/total/total020/test.ipkg diff --git a/tests/idris2/total005/run b/tests/idris2/total005/run deleted file mode 100755 index c1204f255d..0000000000 --- a/tests/idris2/total005/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -idris2 Total.idr < input diff --git a/tests/idris2/total006/run b/tests/idris2/total006/run deleted file mode 100755 index c1204f255d..0000000000 --- a/tests/idris2/total006/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -idris2 Total.idr < input diff --git a/tests/idris2/total007/run b/tests/idris2/total007/run deleted file mode 100755 index 5c09e30a92..0000000000 --- a/tests/idris2/total007/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check partial.idr diff --git a/tests/idris2/total008/run b/tests/idris2/total008/run deleted file mode 100755 index 5c09e30a92..0000000000 --- a/tests/idris2/total008/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check partial.idr diff --git a/tests/idris2/total009/run b/tests/idris2/total009/run deleted file mode 100755 index e33f3104e3..0000000000 --- a/tests/idris2/total009/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check tree.idr diff --git a/tests/idris2/total010/run b/tests/idris2/total010/run deleted file mode 100755 index e21261e515..0000000000 --- a/tests/idris2/total010/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check PartialWith.idr diff --git a/tests/idris2/total013/run b/tests/idris2/total013/run deleted file mode 100755 index a5c8d5d776..0000000000 --- a/tests/idris2/total013/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check Issue1404.idr diff --git a/tests/idris2/total014/run b/tests/idris2/total014/run deleted file mode 100755 index f2e9b12faa..0000000000 --- a/tests/idris2/total014/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check FunCompTC.idr diff --git a/tests/idris2/total017/run b/tests/idris2/total017/run deleted file mode 100755 index ffc5516b50..0000000000 --- a/tests/idris2/total017/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check Paper.idr diff --git a/tests/idris2/total018/run b/tests/idris2/total018/run deleted file mode 100755 index 320b54769f..0000000000 --- a/tests/idris2/total018/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check Issue2448.idr diff --git a/tests/idris2/total019/run b/tests/idris2/total019/run deleted file mode 100755 index afe2e9700b..0000000000 --- a/tests/idris2/total019/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check Check.idr diff --git a/tests/idris2/total020/run b/tests/idris2/total020/run deleted file mode 100755 index afe2e9700b..0000000000 --- a/tests/idris2/total020/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check Check.idr diff --git a/tests/idris2/unification001/run b/tests/idris2/unification001/run deleted file mode 100755 index 0921a91ff0..0000000000 --- a/tests/idris2/unification001/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -check Issue647.idr diff --git a/tests/idris2/warning001/Holes.idr b/tests/idris2/warning/warning001/Holes.idr similarity index 100% rename from tests/idris2/warning001/Holes.idr rename to tests/idris2/warning/warning001/Holes.idr diff --git a/tests/idris2/warning001/Issue1401.idr b/tests/idris2/warning/warning001/Issue1401.idr similarity index 100% rename from tests/idris2/warning001/Issue1401.idr rename to tests/idris2/warning/warning001/Issue1401.idr diff --git a/tests/idris2/warning001/Issue539.idr b/tests/idris2/warning/warning001/Issue539.idr similarity index 100% rename from tests/idris2/warning001/Issue539.idr rename to tests/idris2/warning/warning001/Issue539.idr diff --git a/tests/idris2/warning001/Issue621.idr b/tests/idris2/warning/warning001/Issue621.idr similarity index 100% rename from tests/idris2/warning001/Issue621.idr rename to tests/idris2/warning/warning001/Issue621.idr diff --git a/tests/idris2/warning001/PR1407.idr b/tests/idris2/warning/warning001/PR1407.idr similarity index 100% rename from tests/idris2/warning001/PR1407.idr rename to tests/idris2/warning/warning001/PR1407.idr diff --git a/tests/idris2/warning001/expected b/tests/idris2/warning/warning001/expected similarity index 100% rename from tests/idris2/warning001/expected rename to tests/idris2/warning/warning001/expected diff --git a/tests/idris2/warning001/run b/tests/idris2/warning/warning001/run similarity index 79% rename from tests/idris2/warning001/run rename to tests/idris2/warning/warning001/run index 2c0748ecfd..4a90ca5a9d 100755 --- a/tests/idris2/warning001/run +++ b/tests/idris2/warning/warning001/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh check Issue539.idr check Issue621.idr diff --git a/tests/idris2/warning002/Foo.idr b/tests/idris2/warning/warning002/Foo.idr similarity index 100% rename from tests/idris2/warning002/Foo.idr rename to tests/idris2/warning/warning002/Foo.idr diff --git a/tests/idris2/warning002/Main.idr b/tests/idris2/warning/warning002/Main.idr similarity index 100% rename from tests/idris2/warning002/Main.idr rename to tests/idris2/warning/warning002/Main.idr diff --git a/tests/idris2/warning002/deprecated.ipkg b/tests/idris2/warning/warning002/deprecated.ipkg similarity index 100% rename from tests/idris2/warning002/deprecated.ipkg rename to tests/idris2/warning/warning002/deprecated.ipkg diff --git a/tests/idris2/warning002/expected b/tests/idris2/warning/warning002/expected similarity index 100% rename from tests/idris2/warning002/expected rename to tests/idris2/warning/warning002/expected diff --git a/tests/idris2/warning002/run b/tests/idris2/warning/warning002/run similarity index 76% rename from tests/idris2/warning002/run rename to tests/idris2/warning/warning002/run index 7542784c76..e77a2159a1 100755 --- a/tests/idris2/warning002/run +++ b/tests/idris2/warning/warning002/run @@ -1,4 +1,4 @@ -. ../../testutils.sh +. ../../../testutils.sh idris2 --build deprecated.ipkg idris2 Foo.idr < pure (Left $ "Failed to open socket: " ++ show fail) - res <- bind sock (Just (Hostname "localhost")) 0 - if res /= 0 - then pure (Left $ "Failed to bind socket with error: " ++ show res) - else do - port <- getSockPort sock - res <- listen sock - if res /= 0 - then pure (Left $ "Failed to listen on socket with error: " ++ show res) - else do - forked <- fork (serve port sock) - pure $ Right (port, forked) - - where - serve : Port -> Socket -> IO () - serve port sock = do - Right (s, _) <- accept sock - | Left err => putStrLn ("Failed to accept on socket with error: " ++ show err) - Right (str, _) <- recv s 1024 - | Left err => putStrLn ("Failed to accept on socket with error: " ++ show err) - putStrLn ("Received: " ++ str) - Right n <- send s ("echo: " ++ str) - | Left err => putStrLn ("Server failed to send data with error: " ++ show err) - pure () - -runClient : Port -> IO () -runClient serverPort = do - Right sock <- socket AF_INET Stream 0 - | Left fail => putStrLn ("Failed to open socket: " ++ show fail) - res <- connect sock (Hostname "localhost") serverPort - if res /= 0 - then putStrLn ("Failed to connect client to port " ++ show serverPort ++ ": " ++ show res) - else do - Right n <- send sock ("hello world!") - | Left err => putStrLn ("Client failed to send data with error: " ++ show err) - Right (str, _) <- recv sock 1024 - | Left err => putStrLn ("Client failed to receive on socket with error: " ++ show err) - -- assuming that stdout buffers get flushed in between system calls, this is "guaranteed" - -- to be printed after the server prints its own message - putStrLn ("Received: " ++ str) - -main : IO () -main = do - Right (serverPort, tid) <- runServer - | Left err => putStrLn $ "[server] " ++ err - runClient serverPort diff --git a/tests/node/node014/expected b/tests/node/node014/expected deleted file mode 100644 index 92aad5a55a..0000000000 --- a/tests/node/node014/expected +++ /dev/null @@ -1,2 +0,0 @@ -Received: hello world! -Received: echo: hello world! diff --git a/tests/node/node014/run b/tests/node/node014/run deleted file mode 100755 index f29a1de4ad..0000000000 --- a/tests/node/node014/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -run --cg node -p network Echo.idr diff --git a/tests/racket/barrier001/Main.idr b/tests/racket/barrier001/Main.idr index add64d4997..d2b27330e5 100644 --- a/tests/racket/barrier001/Main.idr +++ b/tests/racket/barrier001/Main.idr @@ -10,6 +10,6 @@ main = do putStrLn "Hello" barrierWait barrier putStrLn "Goodbye" - for threadIDs $ \threadID => + for_ threadIDs $ \threadID => threadWait threadID sleep 1 From cf73efe8e03ab95e29e8a1eea7920185f26e4b93 Mon Sep 17 00:00:00 2001 From: Robert Wright Date: Fri, 1 Sep 2023 14:37:24 +0100 Subject: [PATCH 07/43] Update linter config --- .github/linters/.chktexrc | 14 ++++++++++++++ .github/linters/.markdown-lint.yml | 1 + 2 files changed, 15 insertions(+) create mode 100644 .github/linters/.chktexrc diff --git a/.github/linters/.chktexrc b/.github/linters/.chktexrc new file mode 100644 index 0000000000..3b3a676c60 --- /dev/null +++ b/.github/linters/.chktexrc @@ -0,0 +1,14 @@ +######################## +######################## +## LaTeX Linter rules ## +######################## +######################## + +CmdLine { + # Command terminated with space. + --nowarn 1 + + # You ought to remove spaces in front of punctuation. + # This gets confused in code blocks, failing on eg. `x : Nat` + --nowarn 26 +} diff --git a/.github/linters/.markdown-lint.yml b/.github/linters/.markdown-lint.yml index 8f250e870c..b1f878d305 100644 --- a/.github/linters/.markdown-lint.yml +++ b/.github/linters/.markdown-lint.yml @@ -32,6 +32,7 @@ MD026: MD029: false # Ordered list item prefix MD033: false # Allow inline HTML MD036: false # Emphasis used instead of a heading +MD041: false # First line in file should be a top level heading MD044: # Proper names should have the correct capitalization code_blocks: false html_elements: false From b4d7bba5507715a147fc233c02b6644207d520bd Mon Sep 17 00:00:00 2001 From: Guillaume Allais Date: Fri, 1 Sep 2023 15:40:26 +0100 Subject: [PATCH 08/43] [ cleanup ] use `default` arguments This simplifies most calls to `testsInDir`. --- libs/test/Test/Golden.idr | 14 +++++--- tests/Main.idr | 76 +++++++++++++++++++-------------------- 2 files changed, 48 insertions(+), 42 deletions(-) diff --git a/libs/test/Test/Golden.idr b/libs/test/Test/Golden.idr index f2c990a75d..aadc10d9d8 100644 --- a/libs/test/Test/Golden.idr +++ b/libs/test/Test/Golden.idr @@ -421,15 +421,21 @@ record TestPool where ||| Find all the test in the given directory. export -testsInDir : (dirName : String) -> (testNameFilter : String -> Bool) -> (poolName : String) -> List Requirement -> Codegen -> IO TestPool -testsInDir dirName testNameFilter poolName reqs cg = do +testsInDir : + (dirName : String) -> + {default (const True) pred : String -> Bool} -> + (poolName : String) -> + {default [] requirements : List Requirement} -> + {default Nothing codegen : Codegen} -> + IO TestPool +testsInDir dirName poolName = do Right names <- listDir dirName | Left e => die $ "failed to list " ++ dirName ++ ": " ++ show e - let names = [n | n <- names, testNameFilter n] + let names = [n | n <- names, pred n] let testNames = [dirName ++ "/" ++ n | n <- names] testNames <- filter testNames when (length testNames == 0) $ die $ "no tests found in " ++ dirName - pure $ MkTestPool poolName reqs cg testNames + pure $ MkTestPool poolName requirements codegen testNames where -- Directory without `run` file is not a test isTest : (path : String) -> IO Bool diff --git a/tests/Main.idr b/tests/Main.idr index e4b96c2ca5..e6933b7c69 100644 --- a/tests/Main.idr +++ b/tests/Main.idr @@ -12,73 +12,73 @@ import Test.Golden -- Test cases ttimpTests : IO TestPool -ttimpTests = testsInDir "ttimp" (const True) "TTImp" [] Nothing +ttimpTests = testsInDir "ttimp" "TTImp" idrisTestsBasic : IO TestPool -idrisTestsBasic = testsInDir "idris2/basic" (const True) "Fundamental language features" [] Nothing +idrisTestsBasic = testsInDir "idris2/basic" "Fundamental language features" idrisTestsDebug : IO TestPool -idrisTestsDebug = testsInDir "idris2/debug" (const True) "Debug features" [] Nothing +idrisTestsDebug = testsInDir "idris2/debug" "Debug features" idrisTestsCoverage : IO TestPool -idrisTestsCoverage = testsInDir "idris2/coverage" (const True) "Coverage checking" [] Nothing +idrisTestsCoverage = testsInDir "idris2/coverage" "Coverage checking" idrisTestsTermination : IO TestPool -idrisTestsTermination = testsInDir "idris2/termination" (const True) "Termination checking" [] Nothing +idrisTestsTermination = testsInDir "idris2/termination" "Termination checking" idrisTestsCasetree : IO TestPool -idrisTestsCasetree = testsInDir "idris2/casetree" (const True) "Case tree building" [] Nothing +idrisTestsCasetree = testsInDir "idris2/casetree" "Case tree building" idrisTestsWarning : IO TestPool -idrisTestsWarning = testsInDir "idris2/warning" (const True) "Warnings" [] Nothing +idrisTestsWarning = testsInDir "idris2/warning" "Warnings" idrisTestsFailing : IO TestPool -idrisTestsFailing = testsInDir "idris2/failing" (const True) "Failing blocks" [] Nothing +idrisTestsFailing = testsInDir "idris2/failing" "Failing blocks" ||| Error messages, including parse errors ("perror") idrisTestsError : IO TestPool -idrisTestsError = testsInDir "idris2/error" (const True) "Error messages" [] Nothing +idrisTestsError = testsInDir "idris2/error" "Error messages" idrisTestsInteractive : IO TestPool -idrisTestsInteractive = testsInDir "idris2/interactive" (const True) "Interactive editing" [] Nothing +idrisTestsInteractive = testsInDir "idris2/interactive" "Interactive editing" idrisTestsInterface : IO TestPool -idrisTestsInterface = testsInDir "idris2/interface" (const True) "Interface" [] Nothing +idrisTestsInterface = testsInDir "idris2/interface" "Interface" ||| QTT and linearity related idrisTestsLinear : IO TestPool -idrisTestsLinear = testsInDir "idris2/linear" (const True) "Quantities" [] Nothing +idrisTestsLinear = testsInDir "idris2/linear" "Quantities" idrisTestsLiterate : IO TestPool -idrisTestsLiterate = testsInDir "idris2/literate" (const True) "Literate programming" [] Nothing +idrisTestsLiterate = testsInDir "idris2/literate" "Literate programming" ||| Performance: things which have been slow in the past, or which ||| pose interesting challenges for the elaborator idrisTestsPerformance : IO TestPool -idrisTestsPerformance = testsInDir "idris2/perf" (const True) "Performance" [] Nothing +idrisTestsPerformance = testsInDir "idris2/perf" "Performance" idrisTestsRegression : IO TestPool -idrisTestsRegression = testsInDir "idris2/reg" (const True) "Various regressions" [] Nothing +idrisTestsRegression = testsInDir "idris2/reg" "Various regressions" ||| Data types, including records idrisTestsData : IO TestPool -idrisTestsData = testsInDir "idris2/data" (const True) "Data and record types" [] Nothing +idrisTestsData = testsInDir "idris2/data" "Data and record types" ||| %builtin related tests for the frontend (type-checking) idrisTestsBuiltin : IO TestPool -idrisTestsBuiltin = testsInDir "idris2/builtin" (const True) "Builtin types and functions" [] Nothing +idrisTestsBuiltin = testsInDir "idris2/builtin" "Builtin types and functions" ||| Evaluator, REPL, specialisation idrisTestsEvaluator : IO TestPool -idrisTestsEvaluator = testsInDir "idris2/evaluator" (const True) "Evaluation" [] Nothing +idrisTestsEvaluator = testsInDir "idris2/evaluator" "Evaluation" idrisTestsREPL : IO TestPool -idrisTestsREPL = testsInDir "idris2/repl" (const True) "REPL commands and help" [] Nothing +idrisTestsREPL = testsInDir "idris2/repl" "REPL commands and help" idrisTestsAllSchemes : Requirement -> IO TestPool -idrisTestsAllSchemes cg = testsInDir "allschemes" (const True) +idrisTestsAllSchemes cg = testsInDir "allschemes" ("Test across all scheme backends: " ++ show cg ++ " instance") - [] (Just cg) + {codegen = Just cg} idrisTestsAllBackends : Requirement -> TestPool idrisTestsAllBackends cg = MkTestPool @@ -98,22 +98,22 @@ idrisTestsAllBackends cg = MkTestPool ||| Totality checking, including positivity idrisTestsTotality : IO TestPool -idrisTestsTotality = testsInDir "idris2/total" (const True) "Totality checking" [] Nothing +idrisTestsTotality = testsInDir "idris2/total" "Totality checking" -- This will only work with an Idris compiled via Chez or Racket, but at -- least for the moment we're not officially supporting self hosting any -- other way. If we do, we'll need to have a way to disable these. idrisTestsSchemeEval : IO TestPool -idrisTestsSchemeEval = testsInDir "idris2/schemeeval" (const True) "Scheme Evaluator" [] Nothing +idrisTestsSchemeEval = testsInDir "idris2/schemeeval" "Scheme Evaluator" idrisTestsReflection : IO TestPool -idrisTestsReflection = testsInDir "idris2/reflection" (const True) "Quotation and Reflection" [] Nothing +idrisTestsReflection = testsInDir "idris2/reflection" "Quotation and Reflection" idrisTestsWith : IO TestPool -idrisTestsWith = testsInDir "idris2/with" (const True) "With abstraction" [] Nothing +idrisTestsWith = testsInDir "idris2/with" "With abstraction" idrisTestsIPKG : IO TestPool -idrisTestsIPKG = testsInDir "idris2/pkg" (const True) "Package and .ipkg files" [] Nothing +idrisTestsIPKG = testsInDir "idris2/pkg" "Package and .ipkg files" idrisTests : TestPool idrisTests = MkTestPool "Misc" [] Nothing @@ -150,31 +150,31 @@ idrisTests = MkTestPool "Misc" [] Nothing ] typeddTests : IO TestPool -typeddTests = testsInDir "typedd-book" (const True) "Type Driven Development" [] Nothing +typeddTests = testsInDir "typedd-book" "Type Driven Development" chezTests : IO TestPool -chezTests = testsInDir "chez" (const True) "Chez backend" [] (Just Chez) +chezTests = testsInDir "chez" "Chez backend" {codegen = Just Chez} refcTests : IO TestPool -refcTests = testsInDir "refc" (const True) "Reference counting C backend" [] (Just C) +refcTests = testsInDir "refc" "Reference counting C backend" {codegen = Just C} racketTests : IO TestPool -racketTests = testsInDir "racket" (const True) "Racket backend" [] (Just Racket) +racketTests = testsInDir "racket" "Racket backend" {codegen = Just Racket} nodeTests : IO TestPool -nodeTests = testsInDir "node" (const True) "Node backend" [] (Just Node) +nodeTests = testsInDir "node" "Node backend" {codegen = Just Node} vmcodeInterpTests : IO TestPool -vmcodeInterpTests = testsInDir "vmcode" (const True) "VMCode interpreter" [] Nothing +vmcodeInterpTests = testsInDir "vmcode" "VMCode interpreter" ideModeTests : IO TestPool -ideModeTests = testsInDir "ideMode" (const True) "IDE mode" [] Nothing +ideModeTests = testsInDir "ideMode" "IDE mode" preludeTests : IO TestPool -preludeTests = testsInDir "prelude" (const True) "Prelude library" [] Nothing +preludeTests = testsInDir "prelude" "Prelude library" templateTests : IO TestPool -templateTests = testsInDir "templates" (const True) "Test templates" [] Nothing +templateTests = testsInDir "templates" "Test templates" -- base library tests are run against -- each codegen supported and to keep @@ -182,14 +182,14 @@ templateTests = testsInDir "templates" (const True) "Test templates" [] Nothing -- that only runs if all backends are -- available. baseLibraryTests : IO TestPool -baseLibraryTests = testsInDir "base" (const True) "Base library" [Chez, Node] Nothing +baseLibraryTests = testsInDir "base" "Base library" {requirements = [Chez, Node]} -- same behavior as `baseLibraryTests` contribLibraryTests : IO TestPool -contribLibraryTests = testsInDir "contrib" (const True) "Contrib library" [Chez, Node] Nothing +contribLibraryTests = testsInDir "contrib" "Contrib library" {requirements = [Chez, Node]} codegenTests : IO TestPool -codegenTests = testsInDir "codegen" (const True) "Code generation" [] Nothing +codegenTests = testsInDir "codegen" "Code generation" main : IO () main = runner $ From 8c14e9527b48e1d7e18c4bb7e69c62dce71bc836 Mon Sep 17 00:00:00 2001 From: Steve Dunham Date: Thu, 7 Sep 2023 20:26:18 -0700 Subject: [PATCH 09/43] [ test ] move test to correct location --- tests/idris2/{total020 => total/total021}/Issue-3030.idr | 0 tests/idris2/{total020 => total/total021}/Issue-3030b.idr | 0 tests/idris2/{total020 => total/total021}/Issue-524.idr | 0 tests/idris2/{total020 => total/total021}/expected | 0 tests/idris2/{total020 => total/total021}/run | 0 5 files changed, 0 insertions(+), 0 deletions(-) rename tests/idris2/{total020 => total/total021}/Issue-3030.idr (100%) rename tests/idris2/{total020 => total/total021}/Issue-3030b.idr (100%) rename tests/idris2/{total020 => total/total021}/Issue-524.idr (100%) rename tests/idris2/{total020 => total/total021}/expected (100%) rename tests/idris2/{total020 => total/total021}/run (100%) diff --git a/tests/idris2/total020/Issue-3030.idr b/tests/idris2/total/total021/Issue-3030.idr similarity index 100% rename from tests/idris2/total020/Issue-3030.idr rename to tests/idris2/total/total021/Issue-3030.idr diff --git a/tests/idris2/total020/Issue-3030b.idr b/tests/idris2/total/total021/Issue-3030b.idr similarity index 100% rename from tests/idris2/total020/Issue-3030b.idr rename to tests/idris2/total/total021/Issue-3030b.idr diff --git a/tests/idris2/total020/Issue-524.idr b/tests/idris2/total/total021/Issue-524.idr similarity index 100% rename from tests/idris2/total020/Issue-524.idr rename to tests/idris2/total/total021/Issue-524.idr diff --git a/tests/idris2/total020/expected b/tests/idris2/total/total021/expected similarity index 100% rename from tests/idris2/total020/expected rename to tests/idris2/total/total021/expected diff --git a/tests/idris2/total020/run b/tests/idris2/total/total021/run similarity index 100% rename from tests/idris2/total020/run rename to tests/idris2/total/total021/run From 9d083154a5a061e4f9758108f8b121547943545e Mon Sep 17 00:00:00 2001 From: CodingCellist Date: Wed, 13 Sep 2023 18:37:49 +0200 Subject: [PATCH 10/43] Test pack and lsp during CI (#3067) --- .../{ci-idris2.yml => ci-idris2-and-libs.yml} | 107 ++++++++++++++++-- 1 file changed, 99 insertions(+), 8 deletions(-) rename .github/workflows/{ci-idris2.yml => ci-idris2-and-libs.yml} (86%) diff --git a/.github/workflows/ci-idris2.yml b/.github/workflows/ci-idris2-and-libs.yml similarity index 86% rename from .github/workflows/ci-idris2.yml rename to .github/workflows/ci-idris2-and-libs.yml index 9fdf02e50b..f7d5fe3223 100644 --- a/.github/workflows/ci-idris2.yml +++ b/.github/workflows/ci-idris2-and-libs.yml @@ -1,4 +1,4 @@ -name: Idris2 +name: Idris2 and External Libs on: push: @@ -487,13 +487,18 @@ jobs: shell: bash ###################################################################### - # Ubuntu testing some libraries. + # Testing some libraries on Ubuntu and pack. + # # We are particularly interested in libraries that are heavily using # dependent types, that are prone to find bugs and regressions in the # compiler. ###################################################################### - ubuntu-test-collie: + ###################################################################### + # Test that we can build Collie + ###################################################################### + + ub-test-collie: needs: [initialise, ubuntu-self-host-previous-version] runs-on: ubuntu-latest if: | @@ -521,7 +526,11 @@ jobs: run: | make - ubuntu-test-frex: + ######################################################################## + # Test that we can build Frex + ######################################################################## + + ub-test-frex: needs: [initialise, ubuntu-self-host-previous-version] runs-on: ubuntu-latest if: | @@ -550,7 +559,13 @@ jobs: make make test - ubuntu-test-elab: + ###################################################################### + # Test that we can build Stefan Höck's elab-util and pack + ###################################################################### + + # ELAB-UTIL + + ub-test-elab-util: needs: [initialise, ubuntu-self-host-previous-version] runs-on: ubuntu-latest if: | @@ -578,12 +593,88 @@ jobs: run: | make all + # PACK + # + # N.B. instead of bootstrapping pack, we use the dockerimage where it is + # already, and then use pack to rebuild pack with this job's idris2 sources + + ub-pack-test-pack: + needs: [initialise, ubuntu-self-host-previous-version] + runs-on: ubuntu-latest + if: | + !contains(needs.initialise.outputs.commit_message, '[ci:') + || contains(needs.initialise.outputs.commit_message, '[ci: libs]') + env: + IDRIS2_CG: chez + PACK_DIR: /root/.pack + strategy: + fail-fast: false + # N.B.: + container: ghcr.io/stefan-hoeck/idris2-pack:latest + steps: + - name: Checkout + uses: actions/checkout@v3 + with: + repository: 'stefan-hoeck/idris2-pack' + # by default, pack uses the main idris2 head, not the current job's one + - name: Toml setup + run: | + { echo "[idris2]" + echo "url = \"https://github.com/${GITHUB_REPOSITORY}\"" + echo "commit = \"latest:${GITHUB_REF_NAME}\"" + echo "bootstrap = true" + } > pack.toml + - name: Build idris2-pack + run: | + git config --global --add safe.directory "${PWD}" + git fetch "https://github.com/${GITHUB_REPOSITORY}" "${GITHUB_REF}:${GITHUB_REF_NAME}" + pack install pack + ###################################################################### - # Ubuntu using katla to build html doc of the libs + # Test that we can build the LSP ###################################################################### - ubuntu-katla: - needs: [initialise, ubuntu-test-collie] + ub-pack-test-lsp: + needs: [initialise, ubuntu-self-host-previous-version] + runs-on: ubuntu-latest + if: | + !contains(needs.initialise.outputs.commit_message, '[ci:') + || contains(needs.initialise.outputs.commit_message, '[ci: libs]') + env: + IDRIS2_CG: chez + PACK_DIR: /root/.pack + # LSP is vastly easier to build using pack, so do that + container: ghcr.io/stefan-hoeck/idris2-pack:latest + strategy: + fail-fast: false + steps: + - name: Checkout + uses: actions/checkout@v3 + with: + repository: 'idris-community/idris2-lsp' + - name: Toml setup + run: | + { echo "[idris2]" + echo "url = \"https://github.com/${GITHUB_REPOSITORY}\"" + echo "commit = \"latest:${GITHUB_REF_NAME}\"" + echo "bootstrap = true" + } > pack.toml + # make sure pack is running the PR's Idris2 before building LSP + - name: Build pack with PR-Idris + run: | + git config --global --add safe.directory "${PWD}" + git fetch "https://github.com/${GITHUB_REPOSITORY}" "${GITHUB_REF}:${GITHUB_REF_NAME}" + pack install pack + - name: Build+install idris2-lsp + run: | + pack --no-prompt install-app idris2-lsp + + ###################################################################### + # Test that we can use Katla to build html doc of the libs + ###################################################################### + + ub-test-katla-and-html: + needs: [initialise, ubuntu-self-host-previous-version] runs-on: ubuntu-latest if: | !contains(needs.initialise.outputs.commit_message, '[ci:') From 0029257eece5a935e9903a193b39dd05c0ea5e93 Mon Sep 17 00:00:00 2001 From: Steve Dunham Date: Thu, 14 Sep 2023 06:29:41 -0700 Subject: [PATCH 11/43] [ fix ] consider nest when guessing scrutinee (#3070) --- src/Idris/Elab/Implementation.idr | 4 +--- src/TTImp/Elab/App.idr | 4 +--- src/TTImp/Elab/Case.idr | 3 ++- src/TTImp/Elab/Local.idr | 22 ++++++++-------------- src/TTImp/TTImp.idr | 7 +++++++ tests/idris2/basic/case002/WhereData.idr | 11 +++++++++++ tests/idris2/basic/case002/expected | 1 + tests/idris2/basic/case002/run | 3 +++ 8 files changed, 34 insertions(+), 21 deletions(-) create mode 100644 tests/idris2/basic/case002/WhereData.idr create mode 100644 tests/idris2/basic/case002/expected create mode 100755 tests/idris2/basic/case002/run diff --git a/src/Idris/Elab/Implementation.idr b/src/Idris/Elab/Implementation.idr index d52bb147c3..1dc4301f4a 100644 --- a/src/Idris/Elab/Implementation.idr +++ b/src/Idris/Elab/Implementation.idr @@ -128,9 +128,7 @@ elabImplementation : {vars : _} -> elabImplementation {vars} ifc vis opts_in pass env nest is cons iname ps named impName_in nusing mbody = do -- let impName_in = maybe (mkImplName fc iname ps) id impln -- If we're in a nested block, update the name - let impName_nest = case lookup impName_in (names nest) of - Just (Just n', _) => n' - _ => impName_in + let impName_nest = mapNestedName nest impName_in impName <- inCurrentNS impName_nest -- The interface name might be qualified, so check if it's an -- alias for something diff --git a/src/TTImp/Elab/App.idr b/src/TTImp/Elab/App.idr index 06141ff62b..31e4f5fc1d 100644 --- a/src/TTImp/Elab/App.idr +++ b/src/TTImp/Elab/App.idr @@ -836,9 +836,7 @@ checkApp rig elabinfo nest env fc (IVar fc' n) expargs autoargs namedargs exp " to " ++ show expargs ++ "\n\tFunction type " ++ (show !(toFullNames fnty)) ++ "\n\tExpected app type " ++ show exptyt)) - let fn = case lookup n (names nest) of - Just (Just n', _) => n' - _ => n + let fn = mapNestedName nest n normalisePrims prims env !(checkAppWith rig elabinfo nest env fc ntm nty (Just fn, arglen) expargs autoargs namedargs False exp) where diff --git a/src/TTImp/Elab/Case.idr b/src/TTImp/Elab/Case.idr index ba5e03f348..95ca225856 100644 --- a/src/TTImp/Elab/Case.idr +++ b/src/TTImp/Elab/Case.idr @@ -13,6 +13,7 @@ import Core.Value import Idris.Syntax import Idris.REPL.Opts +import TTImp.Elab.App import TTImp.Elab.Check import TTImp.Elab.Delayed import TTImp.Elab.ImplicitBind @@ -460,7 +461,7 @@ checkCase rig elabinfo nest env fc opts scr scrty_in alts exp = case getFn x of IVar _ n => do defs <- get Ctxt - [(n', (_, ty))] <- lookupTyName n (gamma defs) + [(_, (_, ty))] <- lookupTyName (mapNestedName nest n) (gamma defs) | _ => guessScrType xs Just (tyn, tyty) <- getRetTy defs !(nf defs [] ty) | _ => guessScrType xs diff --git a/src/TTImp/Elab/Local.idr b/src/TTImp/Elab/Local.idr index 6107e8d1eb..048208b39d 100644 --- a/src/TTImp/Elab/Local.idr +++ b/src/TTImp/Elab/Local.idr @@ -103,44 +103,38 @@ localHelper {vars} nest env nestdecls_in func -- Update the names in the declarations to the new 'nested' names. -- When we encounter the names in elaboration, we'll update to an -- application of the nested name. - newName : NestedNames vars -> Name -> Name - newName nest n - = case lookup n (names nest) of - Just (Just n', _) => n' - _ => n - updateTyName : NestedNames vars -> ImpTy -> ImpTy updateTyName nest (MkImpTy loc' nameLoc n ty) - = MkImpTy loc' nameLoc (newName nest n) ty + = MkImpTy loc' nameLoc (mapNestedName nest n) ty updateDataName : NestedNames vars -> ImpData -> ImpData updateDataName nest (MkImpData loc' n tycons dopts dcons) - = MkImpData loc' (newName nest n) tycons dopts + = MkImpData loc' (mapNestedName nest n) tycons dopts (map (updateTyName nest) dcons) updateDataName nest (MkImpLater loc' n tycons) - = MkImpLater loc' (newName nest n) tycons + = MkImpLater loc' (mapNestedName nest n) tycons updateFieldName : NestedNames vars -> IField -> IField updateFieldName nest (MkIField fc rigc piinfo n rawimp) - = MkIField fc rigc piinfo (newName nest n) rawimp + = MkIField fc rigc piinfo (mapNestedName nest n) rawimp updateRecordName : NestedNames vars -> ImpRecord -> ImpRecord updateRecordName nest (MkImpRecord fc n params opts conName fields) - = MkImpRecord fc (newName nest n) + = MkImpRecord fc (mapNestedName nest n) params opts - (newName nest conName) + (mapNestedName nest conName) (map (updateFieldName nest) fields) updateRecordNS : NestedNames vars -> Maybe String -> Maybe String updateRecordNS _ Nothing = Nothing - updateRecordNS nest (Just ns) = Just $ show $ newName nest (UN $ mkUserName ns) + updateRecordNS nest (Just ns) = Just $ show $ mapNestedName nest (UN $ mkUserName ns) updateName : NestedNames vars -> ImpDecl -> ImpDecl updateName nest (IClaim loc' r vis fnopts ty) = IClaim loc' r vis fnopts (updateTyName nest ty) updateName nest (IDef loc' n cs) - = IDef loc' (newName nest n) cs + = IDef loc' (mapNestedName nest n) cs updateName nest (IData loc' vis mbt d) = IData loc' vis mbt (updateDataName nest d) updateName nest (IRecord loc' ns vis mbt imprecord) diff --git a/src/TTImp/TTImp.idr b/src/TTImp/TTImp.idr index 71f6857940..a358f41d4e 100644 --- a/src/TTImp/TTImp.idr +++ b/src/TTImp/TTImp.idr @@ -38,6 +38,13 @@ Weaken NestedNames where wknName (n, (mn, vars, rep)) = (n, (mn, map weaken vars, \fc, nt => weaken (rep fc nt))) +-- replace nested name with full name +export +mapNestedName : NestedNames vars -> Name -> Name +mapNestedName nest n = case lookup n (names nest) of + (Just (Just n', _)) => n' + _ => n + -- Unchecked terms, with implicit arguments -- This is the raw, elaboratable form. -- Higher level expressions (e.g. case, pattern matching let, where blocks, diff --git a/tests/idris2/basic/case002/WhereData.idr b/tests/idris2/basic/case002/WhereData.idr new file mode 100644 index 0000000000..a00d5444e8 --- /dev/null +++ b/tests/idris2/basic/case002/WhereData.idr @@ -0,0 +1,11 @@ +module Main + +foo : Int -> Int +foo x = case isLT of + Yes => x*2 + No => x*4 + where + data MyLT = Yes | No + + isLT : MyLT + isLT = if x < 20 then Yes else No diff --git a/tests/idris2/basic/case002/expected b/tests/idris2/basic/case002/expected new file mode 100644 index 0000000000..28f2441181 --- /dev/null +++ b/tests/idris2/basic/case002/expected @@ -0,0 +1 @@ +1/1: Building WhereData (WhereData.idr) diff --git a/tests/idris2/basic/case002/run b/tests/idris2/basic/case002/run new file mode 100755 index 0000000000..43b46c6ae5 --- /dev/null +++ b/tests/idris2/basic/case002/run @@ -0,0 +1,3 @@ +rm -rf build + +$1 --no-color --console-width 0 --no-banner --check WhereData.idr From dc79c6dd059e06ed24ab9c7041bec5267b71bb90 Mon Sep 17 00:00:00 2001 From: "Thomas E. Hansen" Date: Fri, 15 Sep 2023 10:36:37 +0200 Subject: [PATCH 12/43] [ ci ] re 3067: fix CI on main branch The extra CI jobs introduced in #3067 work fine as long as 'main' is not the checked out branch. This is due to the fetch to a new branch, which git does (reasonably) not allow when you're trying to fetch 'main' into a new branch that's also called 'main'. In this case, we should just `git pull origin main`, which is what the script now (hopefully) does. --- .github/workflows/ci-idris2-and-libs.yml | 17 +++++++++++++++-- 1 file changed, 15 insertions(+), 2 deletions(-) diff --git a/.github/workflows/ci-idris2-and-libs.yml b/.github/workflows/ci-idris2-and-libs.yml index f7d5fe3223..9cc21f6268 100644 --- a/.github/workflows/ci-idris2-and-libs.yml +++ b/.github/workflows/ci-idris2-and-libs.yml @@ -627,7 +627,14 @@ jobs: - name: Build idris2-pack run: | git config --global --add safe.directory "${PWD}" - git fetch "https://github.com/${GITHUB_REPOSITORY}" "${GITHUB_REF}:${GITHUB_REF_NAME}" + + # only fetch to a new branch if we're not on main (otherwise, git complains) + if [[ $(git branch --show-current) != 'main' ]] + then git fetch "https://github.com/${GITHUB_REPOSITORY}" "${GITHUB_REF}:${GITHUB_REF_NAME}" + else git pull origin main + fi + + # rebuild pack with the fetched Idris2 pack install pack ###################################################################### @@ -663,7 +670,13 @@ jobs: - name: Build pack with PR-Idris run: | git config --global --add safe.directory "${PWD}" - git fetch "https://github.com/${GITHUB_REPOSITORY}" "${GITHUB_REF}:${GITHUB_REF_NAME}" + + # only fetch to a new branch if we're not on main (otherwise, git complains) + if [[ $(git branch --show-current) != 'main' ]] + then git fetch "https://github.com/${GITHUB_REPOSITORY}" "${GITHUB_REF}:${GITHUB_REF_NAME}" + else git pull origin main + fi + pack install pack - name: Build+install idris2-lsp run: | From f6c000e27e3d0a1247ea07e1370a982fdf5ce343 Mon Sep 17 00:00:00 2001 From: Joel Berkeley <16429957+joelberkeley@users.noreply.github.com> Date: Tue, 19 Sep 2023 18:40:05 +0100 Subject: [PATCH 13/43] Fix typo in namespace for [bi]traversable composition --- libs/prelude/Prelude/Interfaces.idr | 3 +-- 1 file changed, 1 insertion(+), 2 deletions(-) diff --git a/libs/prelude/Prelude/Interfaces.idr b/libs/prelude/Prelude/Interfaces.idr index e7ae329e37..f4c2ab7880 100644 --- a/libs/prelude/Prelude/Interfaces.idr +++ b/libs/prelude/Prelude/Interfaces.idr @@ -583,8 +583,7 @@ namespace Traversable using Foldable.Compose Functor.Compose where traverse = traverse . traverse -namespace Bitraveresable - +namespace Bitraversable ||| Composition of a bitraversable and a traversable is bitraversable. public export %tcinline [Compose] (l : Traversable t) => (r : Bitraversable p) => Bitraversable (t .: p) From a643e3af627eafb75ffb47f68a0a012264a7458c Mon Sep 17 00:00:00 2001 From: Denis Buzdalov Date: Thu, 21 Sep 2023 12:29:08 +0300 Subject: [PATCH 14/43] [ elab ] Print script's FC in the bad elaboration script error --- src/Idris/Error.idr | 4 ++++ .../reflection022/BadElabScript.idr | 14 +++++++++++++ .../idris2/reflection/reflection022/expected | 21 +++++++++++++++++++ tests/idris2/reflection/reflection022/run | 3 +++ 4 files changed, 42 insertions(+) create mode 100644 tests/idris2/reflection/reflection022/BadElabScript.idr create mode 100644 tests/idris2/reflection/reflection022/expected create mode 100755 tests/idris2/reflection/reflection022/run diff --git a/src/Idris/Error.idr b/src/Idris/Error.idr index 892cdf5372..ac6932e77f 100644 --- a/src/Idris/Error.idr +++ b/src/Idris/Error.idr @@ -603,6 +603,10 @@ perrorRaw (BadRunElab fc env script desc) = pure $ errorDesc (reflow "Bad elaborator script" <++> code !(pshow env script) <++> parens (pretty0 desc) <+> dot) <+> line <+> !(ploc fc) + <+> !(let scriptFC = getLoc script in + if isJust (isNonEmptyFC scriptFC) + then pure $ line <+> reflow "Stuck place in the script:" <+> line <+> !(ploc scriptFC) + else pure emptyDoc) perrorRaw (RunElabFail e) = pure $ reflow "Error during reflection" <+> colon <++> !(perrorRaw e) perrorRaw (GenericMsg fc str) = pure $ pretty0 str <+> line <+> !(ploc fc) diff --git a/tests/idris2/reflection/reflection022/BadElabScript.idr b/tests/idris2/reflection/reflection022/BadElabScript.idr new file mode 100644 index 0000000000..233fcfc29e --- /dev/null +++ b/tests/idris2/reflection/reflection022/BadElabScript.idr @@ -0,0 +1,14 @@ +module BadElabScript + +import Language.Reflection + +%language ElabReflection + +x : Elab Nat +x = do + ignore $ check {expected=Type} `(Nat) + ?someHole + check `(%search) + +y : Nat +y = %runElab x diff --git a/tests/idris2/reflection/reflection022/expected b/tests/idris2/reflection/reflection022/expected new file mode 100644 index 0000000000..d06e7095a2 --- /dev/null +++ b/tests/idris2/reflection/reflection022/expected @@ -0,0 +1,21 @@ +1/1: Building BadElabScript (BadElabScript.idr) +Error: While processing right hand side of y. Bad elaborator script ?someHole [no locals in scope] (script is not a data value). + +BadElabScript:14:5--14:15 + 10 | ?someHole + 11 | check `(%search) + 12 | + 13 | y : Nat + 14 | y = %runElab x + ^^^^^^^^^^ + +Stuck place in the script: + +BadElabScript:10:3--10:12 + 06 | + 07 | x : Elab Nat + 08 | x = do + 09 | ignore $ check {expected=Type} `(Nat) + 10 | ?someHole + ^^^^^^^^^ + diff --git a/tests/idris2/reflection/reflection022/run b/tests/idris2/reflection/reflection022/run new file mode 100755 index 0000000000..51d6e5de9a --- /dev/null +++ b/tests/idris2/reflection/reflection022/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check BadElabScript.idr From 6dc06cd67d44d4ef2fab31b3a6f42b76f22f4527 Mon Sep 17 00:00:00 2001 From: Denis Buzdalov Date: Sat, 23 Sep 2023 18:06:55 +0300 Subject: [PATCH 15/43] [ base ] Add update functions to sorted maps --- CHANGELOG.md | 2 ++ libs/base/Data/SortedMap.idr | 22 ++++++++++++++++++++++ libs/base/Data/SortedMap/Dependent.idr | 22 ++++++++++++++++++++++ 3 files changed, 46 insertions(+) diff --git a/CHANGELOG.md b/CHANGELOG.md index 8c15e01c97..3f1a293ad2 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -203,6 +203,8 @@ * Adds `infixOfBy` and `isInfixOfBy` into `Data.List`. +* Adds updating functions to `SortedMap` and `SortedDMap`. + #### System * Changes `getNProcessors` to return the number of online processors rather than diff --git a/libs/base/Data/SortedMap.idr b/libs/base/Data/SortedMap.idr index b369c7e1ed..6aaba88720 100644 --- a/libs/base/Data/SortedMap.idr +++ b/libs/base/Data/SortedMap.idr @@ -37,6 +37,28 @@ export delete : k -> SortedMap k v -> SortedMap k v delete k = M . delete k . unM +||| Updates or deletes a value based on the decision function +||| +||| The decision function takes information about the presence of the value, +||| and the value itself, if it is present. +||| It returns a new value or the fact that there should be no value as the result. +||| +||| The current implementation performs up to two traversals of the original map +export +update : (Maybe v -> Maybe v) -> k -> SortedMap k v -> SortedMap k v +update f k m = case f $ lookup k m of + Just v => insert k v m + Nothing => delete k m + +||| Updates existing value, if it is present, and does nothing otherwise +||| +||| The current implementation performs up to two traversals of the original map +export +updateExisting : (v -> v) -> k -> SortedMap k v -> SortedMap k v +updateExisting f k m = case lookup k m of + Just v => insert k (f v) m + Nothing => m + export fromList : Ord k => List (k, v) -> SortedMap k v fromList = flip insertFrom empty diff --git a/libs/base/Data/SortedMap/Dependent.idr b/libs/base/Data/SortedMap/Dependent.idr index 5dd83a51ab..9233c7075f 100644 --- a/libs/base/Data/SortedMap/Dependent.idr +++ b/libs/base/Data/SortedMap/Dependent.idr @@ -254,6 +254,28 @@ delete k (M (S n) t) = Left t' => (M _ t') Right t' => (M _ t') +||| Updates or deletes a value based on the decision function +||| +||| The decision function takes information about the presence of the value, +||| and the value itself, if it is present. +||| It returns a new value or the fact that there should be no value as the result. +||| +||| The current implementation performs up to two traversals of the original map +export +update : DecEq k => (x : k) -> (Maybe (v x) -> Maybe (v x)) -> SortedDMap k v -> SortedDMap k v +update k f m = case f $ lookupPrecise k m of + Just v => insert k v m + Nothing => delete k m + +||| Updates existing value, if it is present, and does nothing otherwise +||| +||| The current implementation performs up to two traversals of the original map +export +updateExisting : DecEq k => (x : k) -> (v x -> v x) -> SortedDMap k v -> SortedDMap k v +updateExisting k f m = case lookupPrecise k m of + Just v => insert k (f v) m + Nothing => m + export fromList : Ord k => List (x : k ** v x) -> SortedDMap k v fromList = foldl (flip (uncurry insert)) empty From 3886200d298f65f04c47451e4832c196c58e1a4e Mon Sep 17 00:00:00 2001 From: Denis Buzdalov Date: Fri, 15 Sep 2023 16:49:17 +0300 Subject: [PATCH 16/43] [ fix ] Make `traverse` and friends lazy for `LazyList` --- CHANGELOG.md | 4 ++++ libs/contrib/Data/List/Lazy.idr | 8 ++++---- tests/Main.idr | 2 +- tests/idris2/misc/lazy004/LazyFor.idr | 26 ++++++++++++++++++++++++++ tests/idris2/misc/lazy004/expected | 11 +++++++++++ tests/idris2/misc/lazy004/input | 2 ++ tests/idris2/misc/lazy004/run | 3 +++ 7 files changed, 51 insertions(+), 5 deletions(-) create mode 100644 tests/idris2/misc/lazy004/LazyFor.idr create mode 100644 tests/idris2/misc/lazy004/expected create mode 100644 tests/idris2/misc/lazy004/input create mode 100755 tests/idris2/misc/lazy004/run diff --git a/CHANGELOG.md b/CHANGELOG.md index 3f1a293ad2..ba2866023d 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -228,6 +228,10 @@ about the modulo operator's upper bound, which can be useful when working with indices (for example). +* Existing specialised variants of functions from the `Traversable` for `LazyList` + were made to be indeed lazy by the effect, but their requirements were strengthened + from `Applicative` to `Monad`. + #### Papers * In `Control.DivideAndConquer`: a port of the paper diff --git a/libs/contrib/Data/List/Lazy.idr b/libs/contrib/Data/List/Lazy.idr index 7cb22d7d59..92218cab87 100644 --- a/libs/contrib/Data/List/Lazy.idr +++ b/libs/contrib/Data/List/Lazy.idr @@ -124,16 +124,16 @@ Monad LazyList where -- The result of a traversal will be a non-lazy list in general -- (you can't delay the "effect" of an applicative functor). public export -traverse : Applicative f => (a -> f b) -> LazyList a -> f (List b) +traverse : Monad f => (a -> f b) -> LazyList a -> f (List b) traverse g [] = pure [] -traverse g (x :: xs) = [| g x :: traverse g xs |] +traverse g (x::xs) = pure $ !(g x) :: !(traverse g xs) public export %inline -for : Applicative f => LazyList a -> (a -> f b) -> f (List b) +for : Monad f => LazyList a -> (a -> f b) -> f (List b) for = flip traverse public export %inline -sequence : Applicative f => LazyList (f a) -> f (List a) +sequence : Monad f => LazyList (f a) -> f (List a) sequence = traverse id -- Traverse a lazy list with lazy effect lazily diff --git a/tests/Main.idr b/tests/Main.idr index e6933b7c69..eb6ed85b8a 100644 --- a/tests/Main.idr +++ b/tests/Main.idr @@ -125,7 +125,7 @@ idrisTests = MkTestPool "Misc" [] Nothing "import001", "import002", "import003", "import004", "import005", "import006", "import007", "import008", "import009", -- Implicit laziness, lazy evaluation - "lazy001", "lazy002", "lazy003", + "lazy001", "lazy002", "lazy003", "lazy004", -- Namespace blocks "namespace001", "namespace002", "namespace003", "namespace004", "namespace005", -- Parameters blocks diff --git a/tests/idris2/misc/lazy004/LazyFor.idr b/tests/idris2/misc/lazy004/LazyFor.idr new file mode 100644 index 0000000000..9f77d1099d --- /dev/null +++ b/tests/idris2/misc/lazy004/LazyFor.idr @@ -0,0 +1,26 @@ +import Data.List.Lazy + +import Debug.Trace + +%default total + +xs : LazyList Nat +xs = iterateN 5 (traceValBy (("xs: " ++) . show) . S) Z + +ys : Maybe (List Nat) +ys = for xs $ \n => if n >= 3 then Nothing else Just (10 * n) + +xs' : LazyList Nat +xs' = iterateN 5 (traceValBy (("xs': " ++) . show) . S) Z + +for' : Monad f => LazyList a -> (a -> f b) -> f $ List b +for' [] _ = pure [] +for' (x::xs) g = pure $ !(g x) :: !(for' xs g) + +ys' : Maybe (List Nat) +ys' = for' xs' $ \n => if n >= 3 then Nothing else Just (10 * n) + +main : IO () +main = do + putStrLn $ show ys + putStrLn $ show ys' diff --git a/tests/idris2/misc/lazy004/expected b/tests/idris2/misc/lazy004/expected new file mode 100644 index 0000000000..57f0d0658d --- /dev/null +++ b/tests/idris2/misc/lazy004/expected @@ -0,0 +1,11 @@ +1/1: Building LazyFor (LazyFor.idr) +Main> xs: 1 +xs: 2 +xs: 3 +Nothing +Main> xs': 1 +xs': 2 +xs': 3 +Nothing +Main> +Bye for now! diff --git a/tests/idris2/misc/lazy004/input b/tests/idris2/misc/lazy004/input new file mode 100644 index 0000000000..dc9382e437 --- /dev/null +++ b/tests/idris2/misc/lazy004/input @@ -0,0 +1,2 @@ +:exec putStrLn $ show ys +:exec putStrLn $ show ys' diff --git a/tests/idris2/misc/lazy004/run b/tests/idris2/misc/lazy004/run new file mode 100755 index 0000000000..77d9e01f90 --- /dev/null +++ b/tests/idris2/misc/lazy004/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +idris2 -p contrib LazyFor.idr < input From 46483fd120211148807683396079a32ca2fe16e9 Mon Sep 17 00:00:00 2001 From: Steve Dunham Date: Mon, 25 Sep 2023 14:25:26 -0700 Subject: [PATCH 17/43] [ fix ] support .lidr.md and .lidr.tex extensions (#3071) --- src/Core/Directory.idr | 2 +- src/Idris/Package/Init.idr | 2 +- src/Idris/REPL.idr | 2 +- src/Libraries/Utils/Path.idr | 33 +++++++++++++++++++ src/Parser/Unlit.idr | 5 ++- .../idris2/literate/literate018/Test.lidr.md | 11 +++++++ tests/idris2/literate/literate018/expected | 1 + tests/idris2/literate/literate018/run | 3 ++ 8 files changed, 55 insertions(+), 4 deletions(-) create mode 100644 tests/idris2/literate/literate018/Test.lidr.md create mode 100644 tests/idris2/literate/literate018/expected create mode 100755 tests/idris2/literate/literate018/run diff --git a/src/Core/Directory.idr b/src/Core/Directory.idr index 8037e32754..777f9e4fe2 100644 --- a/src/Core/Directory.idr +++ b/src/Core/Directory.idr @@ -216,7 +216,7 @@ mbPathToNS wdir sdir fname = sdir = fromMaybe "" sdir base = if isAbsolute fname then wdir sdir else sdir in - unsafeFoldModuleIdent . reverse . splitPath . Path.dropExtension + unsafeFoldModuleIdent . reverse . splitPath . Path.dropExtensions <$> (Path.dropBase `on` cleanPath) base fname export diff --git a/src/Idris/Package/Init.idr b/src/Idris/Package/Init.idr index 607cc0d1e1..156d856e31 100644 --- a/src/Idris/Package/Init.idr +++ b/src/Idris/Package/Init.idr @@ -54,7 +54,7 @@ findModules start = do go acc ((path, (root ** iot)) :: iots) = do t <- liftIO iot let mods = flip map t.files $ \ entry => - let fname = fst (splitFileName (fileName entry)) in + let fname = fst (splitExtensions (fileName entry)) in let mod = unsafeFoldModuleIdent (fname :: path) in let fp = toFilePath entry in (mod, fp) diff --git a/src/Idris/REPL.idr b/src/Idris/REPL.idr index 1eb14070b0..8a43944aec 100644 --- a/src/Idris/REPL.idr +++ b/src/Idris/REPL.idr @@ -1083,7 +1083,7 @@ process (ImportPackage package) = do tree <- coreLift $ explore packageDirPath fentries <- coreLift $ toPaths (toRelative tree) errs <- for fentries $ \entry => do - let entry' = dropExtension entry + let entry' = dropExtensions entry let sp = forget $ split (== dirSeparator) entry' let ns = concat $ intersperse "." sp let ns' = mkNamespace ns diff --git a/src/Libraries/Utils/Path.idr b/src/Libraries/Utils/Path.idr index 21b09ac847..4031ad4f7f 100644 --- a/src/Libraries/Utils/Path.idr +++ b/src/Libraries/Utils/Path.idr @@ -21,6 +21,7 @@ import System.File infixl 5 , /> infixr 7 <.> +infixr 7 <..> ||| The character that separates directories in the path. @@ -595,11 +596,43 @@ export show $ setFileName' (stem ++ ext) path' Nothing => path +||| Appends a extension to the path, replacing multiple extensions. +||| +||| If there is no file name, the path will not change; +||| if the path has no extension, the extension will be appended; +||| if the given extension is empty, all of the extensions of the path will be dropped; +||| otherwise, the extensions will be replaced. +||| +||| ```idris example +||| "/tmp/Foo.lidr.md" <.> "idr" == "/tmp/Foo.idr" +||| ``` +export +(<..>) : String -> (extension : String) -> String +(<..>) path ext = + let + path' = parse path + ext = pack $ dropWhile (\char => char == '.' || isSpace char) (unpack ext) + ext = if ext == "" then "" else "." ++ ext + in + case fileName' path' of + Just name => + let (stem, _) = splitExtensions name in + show $ setFileName' (stem ++ ext) path' + Nothing => path + + ||| Drops the extension of the path. export dropExtension : String -> String dropExtension path = path <.> "" +||| Drops all of the extensions of a path. +||| ```idris example +||| "/tmp/Foo.lidr.md" == "/tmp/Foo" +export +dropExtensions : String -> String +dropExtensions path = path <..> "" + ||| Looks up an executable from a list of candidate names in the PATH export pathLookup : List String -> IO (Maybe String) diff --git a/src/Parser/Unlit.idr b/src/Parser/Unlit.idr index f38d2e47dd..067281a72c 100644 --- a/src/Parser/Unlit.idr +++ b/src/Parser/Unlit.idr @@ -47,12 +47,15 @@ styleTeX = MkLitStyle export listOfExtensionsLiterate : List String listOfExtensionsLiterate - = concatMap file_extensions + = do let exts = concatMap file_extensions [ styleBird , styleOrg , styleCMark , styleTeX ] + pfx <- [ "", ".idr", ".lidr"] + ext <- exts + pure (pfx ++ ext) ||| Are we dealing with a valid literate file name, if so return the base name and used extension. export diff --git a/tests/idris2/literate/literate018/Test.lidr.md b/tests/idris2/literate/literate018/Test.lidr.md new file mode 100644 index 0000000000..170cb06dd1 --- /dev/null +++ b/tests/idris2/literate/literate018/Test.lidr.md @@ -0,0 +1,11 @@ +# Markdown test + +Idris should allow multiple extensions on literate files with a module declaration. + +```idris +module Test + +main : IO () +main = printLn "hello" +``` + diff --git a/tests/idris2/literate/literate018/expected b/tests/idris2/literate/literate018/expected new file mode 100644 index 0000000000..5ba7f61e5d --- /dev/null +++ b/tests/idris2/literate/literate018/expected @@ -0,0 +1 @@ +1/1: Building Test (Test.lidr.md) diff --git a/tests/idris2/literate/literate018/run b/tests/idris2/literate/literate018/run new file mode 100755 index 0000000000..50b704f883 --- /dev/null +++ b/tests/idris2/literate/literate018/run @@ -0,0 +1,3 @@ +rm -rf build + +$1 --no-color --console-width 0 --no-banner --check Test.lidr.md From 0c40a76c2c5c2176a2707cc86d3c7d5f7c2fdd8d Mon Sep 17 00:00:00 2001 From: Denis Buzdalov Date: Wed, 27 Sep 2023 20:05:12 +0300 Subject: [PATCH 18/43] [ re #2884 ] Move existing test to an appropriate category --- tests/{contrib => base}/sortedmap_001/SortedMapTest.idr | 0 tests/{contrib => base}/sortedmap_001/expected | 0 tests/base/sortedmap_001/run | 3 +++ tests/contrib/sortedmap_001/run | 3 --- 4 files changed, 3 insertions(+), 3 deletions(-) rename tests/{contrib => base}/sortedmap_001/SortedMapTest.idr (100%) rename tests/{contrib => base}/sortedmap_001/expected (100%) create mode 100644 tests/base/sortedmap_001/run delete mode 100644 tests/contrib/sortedmap_001/run diff --git a/tests/contrib/sortedmap_001/SortedMapTest.idr b/tests/base/sortedmap_001/SortedMapTest.idr similarity index 100% rename from tests/contrib/sortedmap_001/SortedMapTest.idr rename to tests/base/sortedmap_001/SortedMapTest.idr diff --git a/tests/contrib/sortedmap_001/expected b/tests/base/sortedmap_001/expected similarity index 100% rename from tests/contrib/sortedmap_001/expected rename to tests/base/sortedmap_001/expected diff --git a/tests/base/sortedmap_001/run b/tests/base/sortedmap_001/run new file mode 100644 index 0000000000..55ff550ddf --- /dev/null +++ b/tests/base/sortedmap_001/run @@ -0,0 +1,3 @@ +. ../../testutils.sh + +run SortedMapTest.idr diff --git a/tests/contrib/sortedmap_001/run b/tests/contrib/sortedmap_001/run deleted file mode 100644 index 69738df4bf..0000000000 --- a/tests/contrib/sortedmap_001/run +++ /dev/null @@ -1,3 +0,0 @@ -. ../../testutils.sh - -run -p contrib SortedMapTest.idr From 276d41d86c7dc4d0a2f3caa26fa404e088d2b243 Mon Sep 17 00:00:00 2001 From: Denis Buzdalov Date: Wed, 27 Sep 2023 21:25:11 +0300 Subject: [PATCH 19/43] [ new ] Implement `Sized` for `Seq`s --- CHANGELOG.md | 2 ++ libs/contrib/Data/Seq/Sized.idr | 4 ++++ libs/contrib/Data/Seq/Unsized.idr | 4 ++++ 3 files changed, 10 insertions(+) diff --git a/CHANGELOG.md b/CHANGELOG.md index ba2866023d..55ef74a42e 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -232,6 +232,8 @@ were made to be indeed lazy by the effect, but their requirements were strengthened from `Applicative` to `Monad`. +* Implements `Sized` for `Data.Seq.Sized` and `Data.Seq.Unsized`. + #### Papers * In `Control.DivideAndConquer`: a port of the paper diff --git a/libs/contrib/Data/Seq/Sized.idr b/libs/contrib/Data/Seq/Sized.idr index 7483638440..43578cbd76 100644 --- a/libs/contrib/Data/Seq/Sized.idr +++ b/libs/contrib/Data/Seq/Sized.idr @@ -207,3 +207,7 @@ public export implementation {n : Nat} -> Applicative (Seq n) where pure = replicate n (<*>) = zipWith ($) + +public export +implementation Sized (Seq n a) where + size (MkSeq s) = size s diff --git a/libs/contrib/Data/Seq/Unsized.idr b/libs/contrib/Data/Seq/Unsized.idr index b4b850fd74..cb0eb28f53 100644 --- a/libs/contrib/Data/Seq/Unsized.idr +++ b/libs/contrib/Data/Seq/Unsized.idr @@ -218,3 +218,7 @@ public export public export implementation Monad Seq where xs >>= f = foldMap f xs + +public export +implementation Sized (Seq a) where + size (MkSeq s) = size s From f3eff838b2d97d09b0eaa85f40b31d42f5e97761 Mon Sep 17 00:00:00 2001 From: Denis Buzdalov Date: Wed, 27 Sep 2023 19:50:59 +0300 Subject: [PATCH 20/43] [ perf ] Do not split tree too early --- libs/contrib/Data/Seq/Internal.idr | 7 +++---- 1 file changed, 3 insertions(+), 4 deletions(-) diff --git a/libs/contrib/Data/Seq/Internal.idr b/libs/contrib/Data/Seq/Internal.idr index 5d9fb2158a..b9cf10664f 100644 --- a/libs/contrib/Data/Seq/Internal.idr +++ b/libs/contrib/Data/Seq/Internal.idr @@ -441,10 +441,9 @@ export split : Nat -> FingerTree (Elem a) -> (FingerTree (Elem a), FingerTree (Elem a)) split _ Empty = (Empty, Empty) split i xs = - let MkSplit l x r = splitTree i xs - in if size xs > i - then (l, x `consTree` r) - else (xs, Empty) + if size xs > i + then let MkSplit l x r = splitTree i xs in (l, x `consTree` r) + else (xs, Empty) -- Concatenation From 46a2dc1c1f7b195a093715a4146566a43cba027e Mon Sep 17 00:00:00 2001 From: Denis Buzdalov Date: Thu, 22 Jun 2023 17:21:08 +0300 Subject: [PATCH 21/43] [ doc, tiny ] Correct wrong directive for unbound implicits turning off --- docs/source/faq/faq.rst | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/docs/source/faq/faq.rst b/docs/source/faq/faq.rst index b24a36e5ef..f25d008be5 100644 --- a/docs/source/faq/faq.rst +++ b/docs/source/faq/faq.rst @@ -269,7 +269,7 @@ directive: .. code-block:: idris - %auto_implicits off + %unbound_implicits off In this case, you can bind ``n`` and ``m`` as implicits, but not ``ty``, as follows: From 847b525189df2f0dd251ad7759f4ac9569fae12d Mon Sep 17 00:00:00 2001 From: Denis Buzdalov Date: Wed, 4 Oct 2023 12:28:47 +0300 Subject: [PATCH 22/43] [ re #3067, readme ] Fix the build badge in the readme --- README.md | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/README.md b/README.md index 9a42e520bd..0c6f090a6e 100644 --- a/README.md +++ b/README.md @@ -2,7 +2,7 @@ Idris 2 ======= [![Documentation Status](https://readthedocs.org/projects/idris2/badge/?version=latest)](https://idris2.readthedocs.io/en/latest/?badge=latest) -[![Build Status](https://github.com/idris-lang/Idris2/actions/workflows/ci-idris2.yml/badge.svg)](https://github.com/idris-lang/Idris2/actions/workflows/ci-idris2.yml) +[![Build Status](https://github.com/idris-lang/Idris2/actions/workflows/ci-idris2-and-libs.yml/badge.svg?branch=main)](https://github.com/idris-lang/Idris2/actions/workflows/ci-idris2-and-libs.yml?query=branch%3Amain) [Idris 2](https://idris-lang.org/) is a purely functional programming language with first class types. From 567f019230d995bccdbdeed9faf8d05992d80420 Mon Sep 17 00:00:00 2001 From: Denis Buzdalov Date: Tue, 3 Oct 2023 15:48:39 +0300 Subject: [PATCH 23/43] [ test ] Set `IDRIS2_PREFIX` to a local dir when testing --- .gitignore | 1 + tests/idris2/pkg/pkg010/run | 1 - tests/idris2/pkg/pkg016/run | 4 ---- tests/testutils.sh | 23 +++++++++++++++++++++++ 4 files changed, 24 insertions(+), 5 deletions(-) diff --git a/.gitignore b/.gitignore index 0f4b631159..8c797335df 100644 --- a/.gitignore +++ b/.gitignore @@ -18,6 +18,7 @@ idris2docs_venv /ipkg/build /tests/**/build +/tests/**/prefix /tests/**/output* /tests/**/*.so /tests/**/*.dylib diff --git a/tests/idris2/pkg/pkg010/run b/tests/idris2/pkg/pkg010/run index 1135ebaaa6..1a4d91f79d 100755 --- a/tests/idris2/pkg/pkg010/run +++ b/tests/idris2/pkg/pkg010/run @@ -1,6 +1,5 @@ . ../../../testutils.sh -export IDRIS2_PACKAGE_PATH=$IDRIS2_PREFIX/$NAME_VERSION export IDRIS2_PREFIX=$test_dir/currently/nonexistent/dir/ export IDRIS2_INC_CGS= diff --git a/tests/idris2/pkg/pkg016/run b/tests/idris2/pkg/pkg016/run index 7b26a4531a..367a8ff683 100644 --- a/tests/idris2/pkg/pkg016/run +++ b/tests/idris2/pkg/pkg016/run @@ -6,7 +6,3 @@ idris2 --install baz.ipkg > /dev/null idris2 --build test.ipkg build/exec/test - -rm -r "${IDRIS2_PREFIX}"/idris2-*/foo-0 -rm -r "${IDRIS2_PREFIX}"/idris2-*/bar-0 -rm -r "${IDRIS2_PREFIX}"/idris2-*/baz-0 diff --git a/tests/testutils.sh b/tests/testutils.sh index 6dcb4179b9..7f64eb64fc 100755 --- a/tests/testutils.sh +++ b/tests/testutils.sh @@ -9,6 +9,7 @@ idris2="$1" # Delete build files between runs to prevent unexpected differences. # As this is at the top-level, this is run when this script is imported. rm -rf build +rm -rf prefix idris2() { $idris2 --no-banner --console-width 0 --no-color "$@" @@ -32,8 +33,30 @@ sed_literal() { # Folder containing the currently running test if [ "$OS" = "windows" ]; then test_dir="$(cygpath -m "$(pwd)")" + SEP=";" else test_dir="$(pwd)" + SEP=":" +fi + +# Set variables for hygiene testing +if [ -z "$PREFIX_CHANGED" ] && [ -n "$IDRIS2_PREFIX" ]; then + OLD_PREFIX="$IDRIS2_PREFIX" + NEW_PREFIX="$test_dir/prefix" + + OLD_PP="$OLD_PREFIX/$NAME_VERSION" + NEW_PP="$NEW_PREFIX/$NAME_VERSION" + + # Set where to look to installed stuff + export IDRIS2_PACKAGE_PATH="$OLD_PP$SEP$NEW_PP" + export IDRIS2_LIBS="$OLD_PP/libs$SEP$NEW_PP/libs" + export IDRIS2_DATA="$OLD_PP/support$SEP$NEW_PP/support" + + # Set where to install stuff + export IDRIS2_PREFIX="$NEW_PREFIX" + + # Save from re-sourcing this file several times + export PREFIX_CHANGED=1 fi # Remove test directory from output From a5b02747b6a97bc32b27df7cc9ccce3832c493ab Mon Sep 17 00:00:00 2001 From: Denis Buzdalov Date: Tue, 3 Oct 2023 17:23:35 +0300 Subject: [PATCH 24/43] [ re #3066 ] Make the rest of tests to use the same form as the others --- tests/idris2/basic/basic071/run | 7 ++++--- tests/idris2/basic/case001/run | 6 +++--- tests/idris2/basic/case002/run | 4 ++-- tests/idris2/basic/literals001/run | 4 ++-- tests/idris2/literate/literate018/run | 4 ++-- tests/idris2/total/total021/run | 9 ++++----- 6 files changed, 17 insertions(+), 17 deletions(-) diff --git a/tests/idris2/basic/basic071/run b/tests/idris2/basic/basic071/run index d2f5648f3f..e5944da27a 100755 --- a/tests/idris2/basic/basic071/run +++ b/tests/idris2/basic/basic071/run @@ -1,5 +1,6 @@ -rm -rf build -$1 --no-color --console-width 0 --no-banner --check B.idr +. ../../../testutils.sh + +check B.idr # Set very close time for A and B TTC files touch A.idr @@ -13,4 +14,4 @@ sync build/ttc/*/B.tt* echo "-- this should be the last line of output --" -$1 --no-color --console-width 0 --no-banner --check B.idr +check B.idr diff --git a/tests/idris2/basic/case001/run b/tests/idris2/basic/case001/run index 167500c1de..c56b10dc1c 100755 --- a/tests/idris2/basic/case001/run +++ b/tests/idris2/basic/case001/run @@ -1,4 +1,4 @@ -rm -rf build +. ../../../testutils.sh -$1 --no-color --console-width 0 --no-banner --codegen chez InlineCase.idr -o inline-case -grep "define InlineCase-product" build/exec/inline-case_app/inline-case.ss \ No newline at end of file +idris2 --codegen chez InlineCase.idr -o inline-case +grep "define InlineCase-product" build/exec/inline-case_app/inline-case.ss diff --git a/tests/idris2/basic/case002/run b/tests/idris2/basic/case002/run index 43b46c6ae5..49bc952d9d 100755 --- a/tests/idris2/basic/case002/run +++ b/tests/idris2/basic/case002/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../../testutils.sh -$1 --no-color --console-width 0 --no-banner --check WhereData.idr +check WhereData.idr diff --git a/tests/idris2/basic/literals001/run b/tests/idris2/basic/literals001/run index 71e882e509..d35387bb39 100755 --- a/tests/idris2/basic/literals001/run +++ b/tests/idris2/basic/literals001/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../../testutils.sh -$1 --no-color --console-width 0 --no-banner -c Test.idr +check Test.idr diff --git a/tests/idris2/literate/literate018/run b/tests/idris2/literate/literate018/run index 50b704f883..34dfa92ae6 100755 --- a/tests/idris2/literate/literate018/run +++ b/tests/idris2/literate/literate018/run @@ -1,3 +1,3 @@ -rm -rf build +. ../../../testutils.sh -$1 --no-color --console-width 0 --no-banner --check Test.lidr.md +check Test.lidr.md diff --git a/tests/idris2/total/total021/run b/tests/idris2/total/total021/run index 1f977a6226..e6961b3439 100755 --- a/tests/idris2/total/total021/run +++ b/tests/idris2/total/total021/run @@ -1,6 +1,5 @@ -rm -rf build - -$1 --no-color --console-width 0 --no-banner --check Issue-3030.idr -$1 --no-color --console-width 0 --no-banner --check Issue-3030b.idr -$1 --no-color --console-width 0 --no-banner --check Issue-524.idr +. ../../../testutils.sh +check Issue-3030.idr +check Issue-3030b.idr +check Issue-524.idr From 1256ded11082e08855c22ee290dbf5b353a51aad Mon Sep 17 00:00:00 2001 From: Denis Buzdalov Date: Tue, 3 Oct 2023 13:29:28 +0300 Subject: [PATCH 25/43] [ elab ] Implement `Ord` for `Count` --- CHANGELOG.md | 2 ++ libs/base/Language/Reflection/TT.idr | 10 ++++++++++ 2 files changed, 12 insertions(+) diff --git a/CHANGELOG.md b/CHANGELOG.md index 55ef74a42e..e951804d71 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -205,6 +205,8 @@ * Adds updating functions to `SortedMap` and `SortedDMap`. +* Implements `Ord` for `Count` from `Language.Reflection`. + #### System * Changes `getNProcessors` to return the number of online processors rather than diff --git a/libs/base/Language/Reflection/TT.idr b/libs/base/Language/Reflection/TT.idr index c95f7db81a..60b685ab59 100644 --- a/libs/base/Language/Reflection/TT.idr +++ b/libs/base/Language/Reflection/TT.idr @@ -386,6 +386,16 @@ public export Ord Namespace where compare (MkNS ms) (MkNS ns) = compare ms ns +public export +Ord Count where + compare M0 M0 = EQ + compare M0 _ = LT + compare _ M0 = GT + compare M1 M1 = EQ + compare MW MW = EQ + compare MW M1 = GT + compare M1 MW = LT + usernameTag : UserName -> Int usernameTag (Basic _) = 0 usernameTag (Field _) = 1 From 8d5caaa137456adb36d813a0a5241f8c5bdb3edd Mon Sep 17 00:00:00 2001 From: 0xd34df00d <0xd34df00d@gmail.com> Date: Sun, 8 Oct 2023 15:01:30 -0500 Subject: [PATCH 26/43] [ base ] Add `anyToFin` converting a Vect's `Any` to its index --- CHANGELOG.md | 3 +++ libs/base/Data/Vect/Quantifiers.idr | 6 ++++++ 2 files changed, 9 insertions(+) diff --git a/CHANGELOG.md b/CHANGELOG.md index e951804d71..629d5c569e 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -186,6 +186,9 @@ * Add `zipPropertyWith` to `Data.Vect.Quantifiers.All.All`. +* Add `anyToFin` to `Data.Vect.Quantifiers.Any`, + converting the `Any` witness to the index into the corresponding element. + * Implemented `Ord` for `Language.Reflection.TT.Name`, `Language.Reflection.TT.Namespace` and `Language.Reflection.TT.UserName`. diff --git a/libs/base/Data/Vect/Quantifiers.idr b/libs/base/Data/Vect/Quantifiers.idr index 7ee5ca0200..574b146ab2 100644 --- a/libs/base/Data/Vect/Quantifiers.idr +++ b/libs/base/Data/Vect/Quantifiers.idr @@ -61,6 +61,12 @@ namespace Any toExists (Here prf) = Evidence _ prf toExists (There prf) = toExists prf + ||| Get the bounded numeric position of the element satisfying the predicate + public export + anyToFin : {0 xs : Vect n a} -> Any p xs -> Fin n + anyToFin (Here _) = FZ + anyToFin (There later) = FS (anyToFin later) + namespace All ||| A proof that all elements of a vector satisfy a property. It is a list of ||| proofs, corresponding element-wise to the `Vect`. From 32b639ca3c7fc8ef374e35469bdec51199a063e6 Mon Sep 17 00:00:00 2001 From: 0xd34df00d <0xd34df00d@gmail.com> Date: Sun, 8 Oct 2023 15:11:30 -0500 Subject: [PATCH 27/43] [ base ] Prove `anyToFin` preserves the property witnessed by `Any` --- libs/base/Data/Vect/Quantifiers.idr | 8 ++++++++ 1 file changed, 8 insertions(+) diff --git a/libs/base/Data/Vect/Quantifiers.idr b/libs/base/Data/Vect/Quantifiers.idr index 574b146ab2..6157bb27df 100644 --- a/libs/base/Data/Vect/Quantifiers.idr +++ b/libs/base/Data/Vect/Quantifiers.idr @@ -67,6 +67,14 @@ namespace Any anyToFin (Here _) = FZ anyToFin (There later) = FS (anyToFin later) + ||| `anyToFin`'s return type satisfies the predicate + export + anyToFinCorrect : {0 xs : Vect n a} -> + (witness : Any p xs) -> + p (anyToFin witness `index` xs) + anyToFinCorrect (Here prf) = prf + anyToFinCorrect (There later) = anyToFinCorrect later + namespace All ||| A proof that all elements of a vector satisfy a property. It is a list of ||| proofs, corresponding element-wise to the `Vect`. From cbbd0c8caa4e6732318e20b3d4e4f1dba0be1bda Mon Sep 17 00:00:00 2001 From: Denis Buzdalov Date: Tue, 1 Aug 2023 09:52:19 +0300 Subject: [PATCH 28/43] [ elab ] Make `%macro`-function be callable without the extension --- CHANGELOG.md | 5 +++-- src/Core/Binary.idr | 4 ++-- src/Idris/Desugar.idr | 2 +- src/Idris/Resugar.idr | 2 +- src/TTImp/Elab/Ambiguity.idr | 2 +- src/TTImp/Elab/RunElab.idr | 6 +++--- src/TTImp/Elab/Term.idr | 4 ++-- src/TTImp/Reflect.idr | 2 +- src/TTImp/TTImp.idr | 10 +++++----- src/TTImp/TTImp/Functor.idr | 4 ++-- src/TTImp/TTImp/TTC.idr | 8 ++++---- src/TTImp/TTImp/Traversals.idr | 2 +- src/TTImp/Utils.idr | 2 +- .../idris2/reflection/reflection023/DeclMacro.idr | 11 +++++++++++ .../reflection023/UseMacroWithoutExtension.idr | 14 ++++++++++++++ tests/idris2/reflection/reflection023/expected | 2 ++ tests/idris2/reflection/reflection023/run | 3 +++ tests/idris2/reflection/reflection023/test.ipkg | 1 + 18 files changed, 58 insertions(+), 26 deletions(-) create mode 100644 tests/idris2/reflection/reflection023/DeclMacro.idr create mode 100644 tests/idris2/reflection/reflection023/UseMacroWithoutExtension.idr create mode 100644 tests/idris2/reflection/reflection023/expected create mode 100755 tests/idris2/reflection/reflection023/run create mode 100644 tests/idris2/reflection/reflection023/test.ipkg diff --git a/CHANGELOG.md b/CHANGELOG.md index 629d5c569e..66b5ee0c77 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -17,8 +17,9 @@ [#1066](https://github.com/idris-lang/idris2/issues/1066)). * `%hide` directives can now hide conflicting fixities from other modules. * Fixity declarations can now be kept private with export modifiers. -* New fromTTImp, fromName, and fromDecls names for custom TTImp, Name, and Decls - literals. +* New `fromTTImp`, `fromName`, and `fromDecls` names for custom `TTImp`, + `Name`, and `Decls` literals. +* Call to `%macro`-functions do not require the `ElabReflection` extension. ### REPL changes diff --git a/src/Core/Binary.idr b/src/Core/Binary.idr index 087af85692..5183501663 100644 --- a/src/Core/Binary.idr +++ b/src/Core/Binary.idr @@ -25,11 +25,11 @@ import public Libraries.Utils.Binary %default covering ||| TTC files can only be compatible if the version number is the same -||| Update with the current date in YYYYMMDD format, or bump the auxiliary +||| Update with the current date in YYYY_MM_DD_00 format, or bump the auxiliary ||| version number if you're changing the version more than once in the same day. export ttcVersion : Int -ttcVersion = 20230829 * 100 + 0 +ttcVersion = 2023_09_04_00 export checkTTCVersion : String -> Int -> Int -> Core () diff --git a/src/Idris/Desugar.idr b/src/Idris/Desugar.idr index 140b046f58..d8aa6e32df 100644 --- a/src/Idris/Desugar.idr +++ b/src/Idris/Desugar.idr @@ -372,7 +372,7 @@ mutual desugarB side ps (PUnquote fc tm) = pure $ IUnquote fc !(desugarB side ps tm) desugarB side ps (PRunElab fc tm) - = pure $ IRunElab fc !(desugarB side ps tm) + = pure $ IRunElab fc True !(desugarB side ps tm) desugarB side ps (PHole fc br holename) = do when br $ update Syn { bracketholes $= ((UN (Basic holename)) ::) } pure $ IHole fc holename diff --git a/src/Idris/Resugar.idr b/src/Idris/Resugar.idr index 85ccd0b6bb..47b02b54bb 100644 --- a/src/Idris/Resugar.idr +++ b/src/Idris/Resugar.idr @@ -389,7 +389,7 @@ mutual = do ds' <- traverse toPDecl ds pure $ PQuoteDecl fc (catMaybes ds') toPTerm p (IUnquote fc tm) = pure (PUnquote fc !(toPTerm argPrec tm)) - toPTerm p (IRunElab fc tm) = pure (PRunElab fc !(toPTerm argPrec tm)) + toPTerm p (IRunElab fc _ tm) = pure (PRunElab fc !(toPTerm argPrec tm)) toPTerm p (IUnifyLog fc _ tm) = toPTerm p tm toPTerm p (Implicit fc True) = pure (PImplicit fc) diff --git a/src/TTImp/Elab/Ambiguity.idr b/src/TTImp/Elab/Ambiguity.idr index 3e45e772cd..09567214c7 100644 --- a/src/TTImp/Elab/Ambiguity.idr +++ b/src/TTImp/Elab/Ambiguity.idr @@ -145,7 +145,7 @@ expandAmbigName mode nest env orig args (IVar fc x) exp = if (Context.Macro `elem` flags def) && notLHS mode then alternativeFirstSuccess $ reverse $ allSplits args <&> \(macroArgs, extArgs) => - (IRunElab fc $ ICoerced fc $ IVar fc n `buildAlt` macroArgs) `buildAlt` extArgs + (IRunElab fc False $ ICoerced fc $ IVar fc n `buildAlt` macroArgs) `buildAlt` extArgs else wrapDot prim est mode n (map (snd . snd) args) (definition def) (buildAlt (IVar fc n) args) where diff --git a/src/TTImp/Elab/RunElab.idr b/src/TTImp/Elab/RunElab.idr index f3edfbd4fd..05515a3d56 100644 --- a/src/TTImp/Elab/RunElab.idr +++ b/src/TTImp/Elab/RunElab.idr @@ -245,12 +245,12 @@ checkRunElab : {vars : _} -> {auto o : Ref ROpts REPLOpts} -> RigCount -> ElabInfo -> NestedNames vars -> Env Term vars -> - FC -> RawImp -> Maybe (Glued vars) -> + FC -> (requireExtension : Bool) -> RawImp -> Maybe (Glued vars) -> Core (Term vars, Glued vars) -checkRunElab rig elabinfo nest env fc script exp +checkRunElab rig elabinfo nest env fc reqExt script exp = do expected <- mkExpected exp defs <- get Ctxt - unless (isExtension ElabReflection defs) $ + unless (not reqExt || isExtension ElabReflection defs) $ throw (GenericMsg fc "%language ElabReflection not enabled") let n = NS reflectionNS (UN $ Basic "Elab") elabtt <- appCon fc defs n [expected] diff --git a/src/TTImp/Elab/Term.idr b/src/TTImp/Elab/Term.idr index a552d20f39..b0a49bc7f8 100644 --- a/src/TTImp/Elab/Term.idr +++ b/src/TTImp/Elab/Term.idr @@ -204,8 +204,8 @@ checkTerm rig elabinfo nest env (IQuoteDecl fc ds) exp = checkQuoteDecl rig elabinfo nest env fc ds exp checkTerm rig elabinfo nest env (IUnquote fc tm) exp = throw (GenericMsg fc "Can't escape outside a quoted term") -checkTerm rig elabinfo nest env (IRunElab fc tm) exp - = checkRunElab rig elabinfo nest env fc tm exp +checkTerm rig elabinfo nest env (IRunElab fc re tm) exp + = checkRunElab rig elabinfo nest env fc re tm exp checkTerm {vars} rig elabinfo nest env (IPrimVal fc c) exp = do let (cval, cty) = checkPrim {vars} fc c checkExp rig elabinfo env fc cval (gnf env cty) exp diff --git a/src/TTImp/Reflect.idr b/src/TTImp/Reflect.idr index cc007cbf0d..edf9d4ce17 100644 --- a/src/TTImp/Reflect.idr +++ b/src/TTImp/Reflect.idr @@ -607,7 +607,7 @@ mutual = pure (Ref tfc Bound t) reflect fc defs lhs env (IUnquote tfc t) = throw (InternalError "Can't reflect an unquote: escapes should be lifted out") - reflect fc defs lhs env (IRunElab tfc t) + reflect fc defs lhs env (IRunElab tfc _ t) = throw (InternalError "Can't reflect a %runElab") reflect fc defs lhs env (IPrimVal tfc t) = do fc' <- reflect fc defs lhs env tfc diff --git a/src/TTImp/TTImp.idr b/src/TTImp/TTImp.idr index a358f41d4e..12a3bb1710 100644 --- a/src/TTImp/TTImp.idr +++ b/src/TTImp/TTImp.idr @@ -119,7 +119,7 @@ mutual IQuoteName : FC -> Name -> RawImp' nm IQuoteDecl : FC -> List (ImpDecl' nm) -> RawImp' nm IUnquote : FC -> RawImp' nm -> RawImp' nm - IRunElab : FC -> RawImp' nm -> RawImp' nm + IRunElab : FC -> (requireExtension : Bool) -> RawImp' nm -> RawImp' nm IPrimVal : FC -> (c : Constant) -> RawImp' nm IType : FC -> RawImp' nm @@ -207,7 +207,7 @@ mutual show (IQuoteName fc tm) = "(%quotename " ++ show tm ++ ")" show (IQuoteDecl fc tm) = "(%quotedecl " ++ show tm ++ ")" show (IUnquote fc tm) = "(%unquote " ++ show tm ++ ")" - show (IRunElab fc tm) = "(%runelab " ++ show tm ++ ")" + show (IRunElab fc _ tm) = "(%runelab " ++ show tm ++ ")" show (IPrimVal fc c) = show c show (IHole _ x) = "?" ++ x show (IUnifyLog _ lvl x) = "(%logging " ++ show lvl ++ " " ++ show x ++ ")" @@ -608,7 +608,7 @@ findIBinds (IDelay fc tm) = findIBinds tm findIBinds (IForce fc tm) = findIBinds tm findIBinds (IQuote fc tm) = findIBinds tm findIBinds (IUnquote fc tm) = findIBinds tm -findIBinds (IRunElab fc tm) = findIBinds tm +findIBinds (IRunElab fc _ tm) = findIBinds tm findIBinds (IBindHere _ _ tm) = findIBinds tm findIBinds (IBindVar _ n) = [n] findIBinds (IUpdate fc updates tm) @@ -644,7 +644,7 @@ findImplicits (IDelay fc tm) = findImplicits tm findImplicits (IForce fc tm) = findImplicits tm findImplicits (IQuote fc tm) = findImplicits tm findImplicits (IUnquote fc tm) = findImplicits tm -findImplicits (IRunElab fc tm) = findImplicits tm +findImplicits (IRunElab fc _ tm) = findImplicits tm findImplicits (IBindVar _ n) = [n] findImplicits (IUpdate fc updates tm) = findImplicits tm ++ concatMap (findImplicits . getFieldUpdateTerm) updates @@ -877,7 +877,7 @@ getFC (IQuote x _) = x getFC (IQuoteName x _) = x getFC (IQuoteDecl x _) = x getFC (IUnquote x _) = x -getFC (IRunElab x _) = x +getFC (IRunElab x _ _) = x getFC (IAs x _ _ _ _) = x getFC (Implicit x _) = x getFC (IWithUnambigNames x _ _) = x diff --git a/src/TTImp/TTImp/Functor.idr b/src/TTImp/TTImp/Functor.idr index 8504c72f39..55452ed4d1 100644 --- a/src/TTImp/TTImp/Functor.idr +++ b/src/TTImp/TTImp/Functor.idr @@ -62,8 +62,8 @@ mutual = IQuoteDecl fc (map (map f) ds) map f (IUnquote fc t) = IUnquote fc (map f t) - map f (IRunElab fc t) - = IRunElab fc (map f t) + map f (IRunElab fc re t) + = IRunElab fc re (map f t) map f (IPrimVal fc c) = IPrimVal fc c map f (IType fc) diff --git a/src/TTImp/TTImp/TTC.idr b/src/TTImp/TTImp/TTC.idr index 4548995c57..c4521b194c 100644 --- a/src/TTImp/TTImp/TTC.idr +++ b/src/TTImp/TTImp/TTC.idr @@ -72,8 +72,8 @@ mutual = do tag 23; toBuf b fc; toBuf b t toBuf b (IUnquote fc t) = do tag 24; toBuf b fc; toBuf b t - toBuf b (IRunElab fc t) - = do tag 25; toBuf b fc; toBuf b t + toBuf b (IRunElab fc re t) + = do tag 25; toBuf b fc; toBuf b re; toBuf b t toBuf b (IPrimVal fc y) = do tag 26; toBuf b fc; toBuf b y @@ -164,8 +164,8 @@ mutual pure (IQuoteDecl fc y) 24 => do fc <- fromBuf b; y <- fromBuf b pure (IUnquote fc y) - 25 => do fc <- fromBuf b; y <- fromBuf b - pure (IRunElab fc y) + 25 => do fc <- fromBuf b; re <- fromBuf b; y <- fromBuf b + pure (IRunElab fc re y) 26 => do fc <- fromBuf b; y <- fromBuf b pure (IPrimVal fc y) diff --git a/src/TTImp/TTImp/Traversals.idr b/src/TTImp/TTImp/Traversals.idr index 110d684213..8f63e5fe8d 100644 --- a/src/TTImp/TTImp/Traversals.idr +++ b/src/TTImp/TTImp/Traversals.idr @@ -119,7 +119,7 @@ parameters (f : RawImp' nm -> RawImp' nm) mapTTImp (IQuoteName fc n) = f $ IQuoteName fc n mapTTImp (IQuoteDecl fc xs) = f $ IQuoteDecl fc (assert_total $ map mapImpDecl xs) mapTTImp (IUnquote fc t) = f $ IUnquote fc (mapTTImp t) - mapTTImp (IRunElab fc t) = f $ IRunElab fc (mapTTImp t) + mapTTImp (IRunElab fc re t) = f $ IRunElab fc re (mapTTImp t) mapTTImp (IPrimVal fc c) = f $ IPrimVal fc c mapTTImp (IType fc) = f $ IType fc mapTTImp (IHole fc str) = f $ IHole fc str diff --git a/src/TTImp/Utils.idr b/src/TTImp/Utils.idr index 5fa57d8f33..9f7f562c1f 100644 --- a/src/TTImp/Utils.idr +++ b/src/TTImp/Utils.idr @@ -187,7 +187,7 @@ findBindableNamesQuot env used (IUnifyLog fc k x) findBindableNamesQuot env used (IQuote fc x) = [] findBindableNamesQuot env used (IQuoteName fc x) = [] findBindableNamesQuot env used (IQuoteDecl fc xs) = [] -findBindableNamesQuot env used (IRunElab fc x) = [] +findBindableNamesQuot env used (IRunElab fc _ x) = [] export findUniqueBindableNames : diff --git a/tests/idris2/reflection/reflection023/DeclMacro.idr b/tests/idris2/reflection/reflection023/DeclMacro.idr new file mode 100644 index 0000000000..17c0685244 --- /dev/null +++ b/tests/idris2/reflection/reflection023/DeclMacro.idr @@ -0,0 +1,11 @@ +module DeclMacro + +import public Language.Reflection + +export %macro +macroFun : Nat -> Elab Nat +macroFun x = pure $ x + 1 + +export +justScript : Nat -> Elab Nat +justScript x = pure $ x + 2 diff --git a/tests/idris2/reflection/reflection023/UseMacroWithoutExtension.idr b/tests/idris2/reflection/reflection023/UseMacroWithoutExtension.idr new file mode 100644 index 0000000000..53807aba5c --- /dev/null +++ b/tests/idris2/reflection/reflection023/UseMacroWithoutExtension.idr @@ -0,0 +1,14 @@ +module UseMacroWithoutExtension + +import DeclMacro + +useMacro : Nat +useMacro = macroFun 5 + +useMacroCorr : UseMacroWithoutExtension.useMacro = 6 +useMacroCorr = Refl + +failing "%language ElabReflection not enabled" + + stillCan'tUseRunElabWithoutExtension : Nat + stillCan'tUseRunElabWithoutExtension = %runElab justScript 4 diff --git a/tests/idris2/reflection/reflection023/expected b/tests/idris2/reflection/reflection023/expected new file mode 100644 index 0000000000..4f5eed7b64 --- /dev/null +++ b/tests/idris2/reflection/reflection023/expected @@ -0,0 +1,2 @@ +1/2: Building DeclMacro (DeclMacro.idr) +2/2: Building UseMacroWithoutExtension (UseMacroWithoutExtension.idr) diff --git a/tests/idris2/reflection/reflection023/run b/tests/idris2/reflection/reflection023/run new file mode 100755 index 0000000000..d6269bb92f --- /dev/null +++ b/tests/idris2/reflection/reflection023/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check UseMacroWithoutExtension.idr diff --git a/tests/idris2/reflection/reflection023/test.ipkg b/tests/idris2/reflection/reflection023/test.ipkg new file mode 100644 index 0000000000..2c5b5ab52d --- /dev/null +++ b/tests/idris2/reflection/reflection023/test.ipkg @@ -0,0 +1 @@ +package a-test From 6815aefbe0da34106fa83ad8c48dc4dd8358bc4e Mon Sep 17 00:00:00 2001 From: Denis Buzdalov Date: Tue, 10 Oct 2023 21:53:41 +0300 Subject: [PATCH 29/43] [ elab ] Implement file operations, e.g. applicable for type providers --- CHANGELOG.md | 2 + libs/base/Language/Reflection.idr | 41 +++++++++- src/Core/Name/Namespace.idr | 5 ++ src/TTImp/Elab/RunElab.idr | 66 ++++++++++++++++ .../reflection/reflection024/.gitignore | 3 + .../idris2/reflection/reflection024/expected | 61 +++++++++++++++ tests/idris2/reflection/reflection024/run | 77 +++++++++++++++++++ .../reflection024/src/Inside/PrintDirs.idr | 19 +++++ .../reflection024/src/LessSimpleRW.idr | 37 +++++++++ .../reflection/reflection024/src/SimpleRW.idr | 33 ++++++++ .../reflection024/src/TypeProviders.idr | 74 ++++++++++++++++++ .../reflection024/src/existentToRead | 2 + .../reflection024/src/existentToWrite | 1 + .../reflection024/src/fancy-record.txt | 2 + .../idris2/reflection/reflection024/test.ipkg | 5 ++ 15 files changed, 427 insertions(+), 1 deletion(-) create mode 100644 tests/idris2/reflection/reflection024/.gitignore create mode 100644 tests/idris2/reflection/reflection024/expected create mode 100755 tests/idris2/reflection/reflection024/run create mode 100644 tests/idris2/reflection/reflection024/src/Inside/PrintDirs.idr create mode 100644 tests/idris2/reflection/reflection024/src/LessSimpleRW.idr create mode 100644 tests/idris2/reflection/reflection024/src/SimpleRW.idr create mode 100644 tests/idris2/reflection/reflection024/src/TypeProviders.idr create mode 100644 tests/idris2/reflection/reflection024/src/existentToRead create mode 100644 tests/idris2/reflection/reflection024/src/existentToWrite create mode 100644 tests/idris2/reflection/reflection024/src/fancy-record.txt create mode 100644 tests/idris2/reflection/reflection024/test.ipkg diff --git a/CHANGELOG.md b/CHANGELOG.md index 66b5ee0c77..a4941be8b6 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -20,6 +20,8 @@ * New `fromTTImp`, `fromName`, and `fromDecls` names for custom `TTImp`, `Name`, and `Decls` literals. * Call to `%macro`-functions do not require the `ElabReflection` extension. +* Elaborator scripts were made to be able to access project files, + allowing the support for type providers and similar stuff. ### REPL changes diff --git a/libs/base/Language/Reflection.idr b/libs/base/Language/Reflection.idr index 512c576053..97ab3b361d 100644 --- a/libs/base/Language/Reflection.idr +++ b/libs/base/Language/Reflection.idr @@ -11,6 +11,23 @@ import public Control.Monad.Trans --- Elaboration data structure --- ---------------------------------- +public export +data LookupDir = + ||| The dir of the `ipkg`-file, or the current dir if there is no one + ProjectDir | + ||| The source dir set in the `ipkg`-file, or the current dir if there is no one + SourceDir | + ||| The dir where the current module is located + ||| + ||| For the module `Language.Reflection` it would be `/Language/` + CurrentModuleDir | + ||| The dir where submodules of the current module are located + ||| + ||| For the module `Language.Reflection` it would be `/Language/Reflection/` + SubmodulesDir | + ||| The dir where built files are located, set in the `ipkg`-file and defaulted to `build` + BuildDir + ||| Elaboration scripts ||| Where types/terms are returned, binders will have unique, if not ||| necessarily human readabe, names @@ -73,6 +90,13 @@ data Elab : Type -> Type where -- Check a group of top level declarations Declare : List Decl -> Elab () + -- Read the contents of a file, if it is present + ReadFile : LookupDir -> (path : String) -> Elab $ Maybe String + -- Writes to a file, replacing existing contents, if were present + WriteFile : LookupDir -> (path : String) -> (contents : String) -> Elab () + -- Returns the specified type of dir related to the current idris project + IdrisDir : LookupDir -> Elab String + export Functor Elab where map f e = Bind e $ Pure . f @@ -148,7 +172,7 @@ interface Monad m => Elaboration m where ||| Given a possibly ambiguous name, get all the matching names and their types getType : Name -> m (List (Name, TTImp)) - ||| Get the metadata associated with a name. Returns all matching namea and their types + ||| Get the metadata associated with a name. Returns all matching names and their types getInfo : Name -> m (List (Name, NameInfo)) ||| Get the type of a local variable @@ -160,6 +184,15 @@ interface Monad m => Elaboration m where ||| Make some top level declarations declare : List Decl -> m () + ||| Read the contents of a file, if it is present + readFile : LookupDir -> (path : String) -> m $ Maybe String + + ||| Writes to a file, replacing existing contents, if were present + writeFile : LookupDir -> (path : String) -> (contents : String) -> m () + + ||| Returns the specified type of dir related to the current idris project + idrisDir : LookupDir -> m String + export %inline ||| Report an error in elaboration fail : Elaboration m => String -> m a @@ -195,6 +228,9 @@ Elaboration Elab where getLocalType = GetLocalType getCons = GetCons declare = Declare + readFile = ReadFile + writeFile = WriteFile + idrisDir = IdrisDir public export Elaboration m => MonadTrans t => Monad (t m) => Elaboration (t m) where @@ -216,6 +252,9 @@ Elaboration m => MonadTrans t => Monad (t m) => Elaboration (t m) where getLocalType = lift . getLocalType getCons = lift . getCons declare = lift . declare + readFile = lift .: readFile + writeFile d = lift .: writeFile d + idrisDir = lift . idrisDir ||| Catch failures and use the `Maybe` monad instead export diff --git a/src/Core/Name/Namespace.idr b/src/Core/Name/Namespace.idr index 1e8f8f7b21..7e683b5ef0 100644 --- a/src/Core/Name/Namespace.idr +++ b/src/Core/Name/Namespace.idr @@ -124,6 +124,11 @@ namespace ModuleIdent toPath : ModuleIdent -> String toPath = joinPath . reverse . unsafeUnfoldModuleIdent + export + parent : ModuleIdent -> Maybe ModuleIdent + parent (MkMI (_::rest)) = Just $ MkMI rest + parent _ = Nothing + ------------------------------------------------------------------------------------- -- HIERARCHICAL STRUCTURE ------------------------------------------------------------------------------------- diff --git a/src/TTImp/Elab/RunElab.idr b/src/TTImp/Elab/RunElab.idr index 05515a3d56..4158990107 100644 --- a/src/TTImp/Elab/RunElab.idr +++ b/src/TTImp/Elab/RunElab.idr @@ -3,6 +3,7 @@ module TTImp.Elab.RunElab import Core.Context import Core.Context.Log import Core.Core +import Core.Directory import Core.Env import Core.Metadata import Core.Normalise @@ -16,6 +17,8 @@ import Idris.Resugar import Idris.REPL.Opts import Idris.Syntax +import Libraries.Utils.Path + import TTImp.Elab.Check import TTImp.Elab.Delayed import TTImp.Reflect @@ -23,6 +26,8 @@ import TTImp.TTImp import TTImp.TTImp.Functor import TTImp.Unelab +import System.File.Meta + %default covering record NameInfo where @@ -85,6 +90,47 @@ elabScript rig fc nest env script@(NDCon nfc nm t ar args) exp EmptyFC => fc x => x + -- parses and resolves `Language.Reflection.LookupDir` + lookupDir : Defs -> NF vars -> Core String + lookupDir defs (NDCon _ conName _ _ []) + = do defs <- get Ctxt + NS ns (UN (Basic n)) <- toFullNames conName + | fnm => failWith defs $ "bad lookup dir fullnames " ++ show fnm + let True = ns == reflectionNS + | False => failWith defs $ "bad lookup dir namespace " ++ show ns + let wd = defs.options.dirs.working_dir + let sd = maybe wd (\sd => joinPath [wd, sd]) defs.options.dirs.source_dir + let modDir : Bool -> Core String := \doParent => do + syn <- get Syn + let parentIfNeeded = if doParent then parent else pure + let moduleSuffix = toList $ (head' syn.saveMod >>= parentIfNeeded) <&> toPath + pure $ joinPath $ sd :: moduleSuffix + case n of + "ProjectDir" => pure wd + "SourceDir" => pure sd + "CurrentModuleDir" => modDir True + "SubmodulesDir" => modDir False + "BuildDir" => pure $ joinPath [wd, defs.options.dirs.build_dir] + _ => failWith defs $ "bad lookup dir value" + lookupDir defs lk + = do defs <- get Ctxt + empty <- clearDefs defs + throw (BadRunElab fc env !(quote empty env lk) "lookup dir is not a data value") + + validatePath : Defs -> String -> Core () + validatePath defs path = do + let True = isRelative path + | False => failWith defs $ "path must be relative" + let True = pathDoesNotEscape 0 $ splitPath path + | False => failWith defs $ "path must not escape the directory" + pure () + where + pathDoesNotEscape : (depth : Nat) -> List String -> Bool + pathDoesNotEscape _ [] = True + pathDoesNotEscape Z (".."::rest) = False + pathDoesNotEscape (S n) (".."::rest) = pathDoesNotEscape n rest + pathDoesNotEscape n (_ ::rest) = pathDoesNotEscape (S n) rest + elabCon : Defs -> String -> List (Closure vars) -> Core (NF vars) elabCon defs "Pure" [_,val] = do empty <- clearDefs defs @@ -228,6 +274,26 @@ elabScript rig fc nest env script@(NDCon nfc nm t ar args) exp decls <- reify defs d' traverse_ (processDecl [] (MkNested []) []) decls scriptRet () + elabCon defs "ReadFile" [lk, pth] + = do pathPrefix <- lookupDir defs !(evalClosure defs lk) + path <- reify defs !(evalClosure defs pth) + validatePath defs path + let fullPath = joinPath [pathPrefix, path] + True <- coreLift $ exists fullPath + | False => scriptRet $ Nothing {ty=String} + contents <- readFile fullPath + scriptRet $ Just contents + elabCon defs "WriteFile" [lk, pth, contents] + = do pathPrefix <- lookupDir defs !(evalClosure defs lk) + path <- reify defs !(evalClosure defs pth) + validatePath defs path + contents <- reify defs !(evalClosure defs contents) + let fullPath = joinPath [pathPrefix, path] + whenJust (parent fullPath) ensureDirectoryExists + writeFile fullPath contents + scriptRet () + elabCon defs "IdrisDir" [lk] + = do evalClosure defs lk >>= lookupDir defs >>= scriptRet elabCon defs n args = failWith defs $ "unexpected Elab constructor " ++ n ++ ", or incorrect count of arguments: " ++ show (length args) elabScript rig fc nest env script exp diff --git a/tests/idris2/reflection/reflection024/.gitignore b/tests/idris2/reflection/reflection024/.gitignore new file mode 100644 index 0000000000..ddbaeea8c5 --- /dev/null +++ b/tests/idris2/reflection/reflection024/.gitignore @@ -0,0 +1,3 @@ +src/*xistent* +src/..a-dot-named-file +src/another-fancy-record.json diff --git a/tests/idris2/reflection/reflection024/expected b/tests/idris2/reflection/reflection024/expected new file mode 100644 index 0000000000..8da12f7b0c --- /dev/null +++ b/tests/idris2/reflection/reflection024/expected @@ -0,0 +1,61 @@ +--- Print dirs with ipkg --- +1/1: Building Inside.PrintDirs (src/Inside/PrintDirs.idr) +LOG elab:0: project dir: __TEST_DIR__ +LOG elab:0: source dir: __TEST_DIR__/src +LOG elab:0: current module dir: __TEST_DIR__/src/Inside +LOG elab:0: submodules dir: __TEST_DIR__/src/Inside/PrintDirs +LOG elab:0: build dir: __TEST_DIR__/build + +--- Print dirs with ipkg (with changed pwd) --- +1/1: Building Inside.PrintDirs (src/Inside/PrintDirs.idr) +LOG elab:0: project dir: __TEST_DIR__ +LOG elab:0: source dir: __TEST_DIR__/src +LOG elab:0: current module dir: __TEST_DIR__/src/Inside +LOG elab:0: submodules dir: __TEST_DIR__/src/Inside/PrintDirs +LOG elab:0: build dir: __TEST_DIR__/build + +--- Print dirs without ipkg (with changed pwd) --- +1/1: Building Inside.PrintDirs (Inside/PrintDirs.idr) +LOG elab:0: project dir: __TEST_DIR__ +LOG elab:0: source dir: __TEST_DIR__ +LOG elab:0: current module dir: __TEST_DIR__/Inside +LOG elab:0: submodules dir: __TEST_DIR__/Inside/PrintDirs +LOG elab:0: build dir: __TEST_DIR__/build + +--- Simple reads and writes --- +1/1: Building SimpleRW (src/SimpleRW.idr) +LOG elab:0: reading existentToRead: +LOG elab:0: contents: +existent to read +second line + +LOG elab:0: reading nonExistentToRead: +LOG elab:0: FILE DOES NOT EXIST +LOG elab:0: written to existentToWrite +LOG elab:0: written to nonExistentToWrite +existent to read +second line +cat: src/nonExistentToRead: No such file or directory +WRITTEN CONTENTS +LA-LA-LA +WRITTEN CONTENTS +LA-LA-LA + +--- A little but less simple reads and writes --- +1/1: Building LessSimpleRW (src/LessSimpleRW.idr) +LOG elab:0: written to nonExistentOriginally/a-generated-file +LOG elab:0: reading nonExistentOriginally/../nonExistentOriginally/a-generated-file: +LOG elab:0: contents: +WRITTEN CONTENTS +LA-LA-LA + +LOG elab:0: written to ..a-dot-named-file +WRITTEN CONTENTS +LA-LA-LA +WRITTEN CONTENTS +LA-LA-LA + +--- Type providers --- +1/1: Building TypeProviders (src/TypeProviders.idr) +Derived: +{"veryStringField":"String","veryIntegerField":"Integer","varyNatField":"Prelude.Types.Nat"} diff --git a/tests/idris2/reflection/reflection024/run b/tests/idris2/reflection/reflection024/run new file mode 100755 index 0000000000..823e58d1f4 --- /dev/null +++ b/tests/idris2/reflection/reflection024/run @@ -0,0 +1,77 @@ +. ../../../testutils.sh + +echo "--- Print dirs with ipkg ---" +check --find-ipkg src/Inside/PrintDirs.idr | filter_test_dir + +############################## + +echo +. ../../../testutils.sh + +( +cd src +echo "--- Print dirs with ipkg (with changed pwd) ---" +check --find-ipkg Inside/PrintDirs.idr | filter_test_dir +) + +############################## + +echo + +( +cd src +. ../../../../testutils.sh # this is here because we will pollute with `build` dir here +echo "--- Print dirs without ipkg (with changed pwd) ---" +check Inside/PrintDirs.idr | filter_test_dir +) + +############################## + +echo +. ../../../testutils.sh + +echo "--- Simple reads and writes ---" + +rm -rf src/nonExistent* +echo "non-overwritten existent to write" > src/existentToWrite + +check --find-ipkg src/SimpleRW.idr + +cat src/existentToRead +cat src/nonExistentToRead 2>&1 +cat src/existentToWrite +cat src/nonExistentToWrite 2>&1 + +rm -rf src/nonExistent* +echo "non-overwritten existent to write" > src/existentToWrite + +############################## + +echo +. ../../../testutils.sh + +echo "--- A little but less simple reads and writes ---" + +rm -rf src/nonExistent* +rm -f src/..a-dot-named-file + +check --find-ipkg src/LessSimpleRW.idr + +cat src/nonExistentOriginally/a-generated-file 2>&1 +cat src/..a-dot-named-file 2>&1 + +rm -rf src/nonExistent* +rm -f src/..a-dot-named-file + +############################## + +echo +. ../../../testutils.sh + +echo "--- Type providers ---" +rm -rf src/another-fancy-record.json +check --find-ipkg src/TypeProviders.idr + +echo "Derived:" +cat src/another-fancy-record.json +echo diff --git a/tests/idris2/reflection/reflection024/src/Inside/PrintDirs.idr b/tests/idris2/reflection/reflection024/src/Inside/PrintDirs.idr new file mode 100644 index 0000000000..1929045909 --- /dev/null +++ b/tests/idris2/reflection/reflection024/src/Inside/PrintDirs.idr @@ -0,0 +1,19 @@ +module Inside.PrintDirs + +import Language.Reflection + +dirs : List (String, LookupDir) +dirs = + [ ("project dir" , ProjectDir ) + , ("source dir" , SourceDir ) + , ("current module dir", CurrentModuleDir) + , ("submodules dir" , SubmodulesDir ) + , ("build dir" , BuildDir ) + ] + +logAllDirs : Elab () +logAllDirs = for_ dirs $ \(msg, lk) => logMsg "elab" 0 "\{msg}: \{!(idrisDir lk)}" + +%language ElabReflection + +%runElab logAllDirs diff --git a/tests/idris2/reflection/reflection024/src/LessSimpleRW.idr b/tests/idris2/reflection/reflection024/src/LessSimpleRW.idr new file mode 100644 index 0000000000..bb6336deb8 --- /dev/null +++ b/tests/idris2/reflection/reflection024/src/LessSimpleRW.idr @@ -0,0 +1,37 @@ +module LessSimpleRW + +import Language.Reflection + +%default total + +readAndLog : (path : String) -> Elab () +readAndLog path = do + v <- readFile CurrentModuleDir path + logMsg "elab" 0 "reading \{path}:" + logMsg "elab" 0 $ (" " ++) $ maybe "FILE DOES NOT EXIST" ("contents:\n" ++) v + +writeAndLog : (path : String) -> Elab () +writeAndLog path = do + writeFile CurrentModuleDir path "WRITTEN CONTENTS\nLA-LA-LA\n" + logMsg "elab" 0 "written to \{path}" + +%language ElabReflection + +-- Check that we can write to a dir that was not previously existent +%runElab writeAndLog "nonExistentOriginally/a-generated-file" + +-- Check that '..' in path are okay unless they escape the lookup dir +%runElab readAndLog "nonExistentOriginally/../nonExistentOriginally/a-generated-file" + +-- Check double dots are allowed inside file names +%runElab writeAndLog "..a-dot-named-file" + +failing "path must not escape the directory" + + -- Check that we cannot escape + %runElab readAndLog "../whatever" + +failing "path must not escape the directory" + + -- Check a slightly more complicated case of escaping + %runElab readAndLog "nonExistentOriginally/../../whatever" diff --git a/tests/idris2/reflection/reflection024/src/SimpleRW.idr b/tests/idris2/reflection/reflection024/src/SimpleRW.idr new file mode 100644 index 0000000000..315c071dc5 --- /dev/null +++ b/tests/idris2/reflection/reflection024/src/SimpleRW.idr @@ -0,0 +1,33 @@ +module SimpleRW + +import Language.Reflection + +%default total + +existentToRead, nonExistentToRead, existentToWrite, nonExistentToWrite : String +existentToRead = "existentToRead" +nonExistentToRead = "nonExistentToRead" +existentToWrite = "existentToWrite" +nonExistentToWrite = "nonExistentToWrite" + +readAndLog : (path : String) -> Elab () +readAndLog path = do + v <- readFile CurrentModuleDir path + logMsg "elab" 0 "reading \{path}:" + logMsg "elab" 0 $ (" " ++) $ maybe "FILE DOES NOT EXIST" ("contents:\n" ++) v + +writeAndLog : (path : String) -> Elab () +writeAndLog path = do + writeFile CurrentModuleDir path "WRITTEN CONTENTS\nLA-LA-LA\n" + logMsg "elab" 0 "written to \{path}" + +go : Elab () +go = do + readAndLog existentToRead + readAndLog nonExistentToRead + writeAndLog existentToWrite + writeAndLog nonExistentToWrite + +%language ElabReflection + +%runElab go diff --git a/tests/idris2/reflection/reflection024/src/TypeProviders.idr b/tests/idris2/reflection/reflection024/src/TypeProviders.idr new file mode 100644 index 0000000000..b5acc75907 --- /dev/null +++ b/tests/idris2/reflection/reflection024/src/TypeProviders.idr @@ -0,0 +1,74 @@ +||| This module is purely for illustration purposes of the mechanism a-la original type providers +module TypeProviders + +import Data.List1 + +import Language.JSON.Data + +import Language.Reflection + +%default total + +%language ElabReflection + +||| Mode one: source of truth in an external file, so generate Idris definitions using a file + +declareRecordFromSimpleText : LookupDir -> (path : String) -> (tyName, consName : Name) -> Elab () +declareRecordFromSimpleText lk textPath dataName consName = do + let errDesc : Elab String := do pure "text at \{textPath} in \{!(idrisDir lk)}" + Just text <- readFile lk textPath + | Nothing => do fail "Did not found \{!errDesc}" + let fs = lines text + fs <- for fs $ \line => case words line of + [name, type] => pure (name, type) + _ => fail "Expected two strings in line in \{!errDesc}" + let dataDecl : Decl = IRecord EmptyFC Nothing Public Nothing $ + MkRecord EmptyFC dataName [] [] consName $ + fs <&> \(name, type) => do + let (ns, n) = unsnoc $ split (== '.') type + let type = if null ns then UN $ Basic n + else NS (MkNS $ reverse ns) (UN $ Basic n) + MkIField EmptyFC MW ExplicitArg (UN $ Basic name) (IVar EmptyFC type) + declare [dataDecl] + +-- Declare a type from a definitions in a file +%runElab declareRecordFromSimpleText CurrentModuleDir "fancy-record.txt" `{FancyRecord} `{MkFancyRecord} + +-- Use newly created data type +frVal : FancyRecord +frVal = MkFancyRecord {natField = S Z, boolField = True} + +||| Mode two: source of truth in Idris module, so generate an external file by Idris definition + +saveRecordToSimpleJSON : LookupDir -> (path : String) -> (type : Name) -> Elab () +saveRecordToSimpleJSON lk jsonPath type = do + infos <- getInfo type + let [(type, _)] = flip filter infos $ \case + (name, MkNameInfo $ TyCon _ _) => True + _ => False + | [] => fail "Did not found any type matching \{show type}" + | (_::_) => fail "Too many types matching \{show type}" + [con] <- getCons type + | [] => fail "Did not found any constructors of \{show type}" + | (_::_) => fail "Too many constructors of \{show type}" + [(_, conType)] <- getType con + | _ => fail "Error while getting the type of constructor \{show con}" + args <- unPi conType + let args = args <&> map (JString . show) + writeFile lk jsonPath $ show $ JObject args + + where + + unPi : TTImp -> Elab $ List (String, TTImp) + unPi $ IPi _ MW ExplicitArg (Just $ UN $ Basic nm) ty rest = ((nm, ty) ::) <$> unPi rest + unPi $ IPi fc _ _ _ _ _ = failAt fc "Unsupported parameter (we support only omega explicit simply named ones)" + unPi _ = pure [] + +-- Declare a type which would be a source of truth +record AnotherFancyRecord where + veryStringField : String + veryIntegerField : Integer + varyNatField : Nat + +-- Generate an external description to a data type +%runElab saveRecordToSimpleJSON CurrentModuleDir "another-fancy-record.json" `{AnotherFancyRecord} diff --git a/tests/idris2/reflection/reflection024/src/existentToRead b/tests/idris2/reflection/reflection024/src/existentToRead new file mode 100644 index 0000000000..2f83a54454 --- /dev/null +++ b/tests/idris2/reflection/reflection024/src/existentToRead @@ -0,0 +1,2 @@ +existent to read +second line diff --git a/tests/idris2/reflection/reflection024/src/existentToWrite b/tests/idris2/reflection/reflection024/src/existentToWrite new file mode 100644 index 0000000000..a75a770286 --- /dev/null +++ b/tests/idris2/reflection/reflection024/src/existentToWrite @@ -0,0 +1 @@ +non-overwritten existent to write diff --git a/tests/idris2/reflection/reflection024/src/fancy-record.txt b/tests/idris2/reflection/reflection024/src/fancy-record.txt new file mode 100644 index 0000000000..39fa6add12 --- /dev/null +++ b/tests/idris2/reflection/reflection024/src/fancy-record.txt @@ -0,0 +1,2 @@ +natField Prelude.Types.Nat +boolField Bool diff --git a/tests/idris2/reflection/reflection024/test.ipkg b/tests/idris2/reflection/reflection024/test.ipkg new file mode 100644 index 0000000000..86bdfc4b73 --- /dev/null +++ b/tests/idris2/reflection/reflection024/test.ipkg @@ -0,0 +1,5 @@ +package a-test + +sourcedir = "src" + +depends = contrib From f7d4b7f4ed64b209d230346ee57a12f2e6416cac Mon Sep 17 00:00:00 2001 From: Denis Buzdalov Date: Mon, 9 Oct 2023 15:45:47 +0300 Subject: [PATCH 30/43] [ base ] Add a bridge between `MonadState` and `Ref` --- CHANGELOG.md | 3 +++ libs/base/Data/Ref.idr | 11 ++++++++++ .../data_ref_monadstate/RefMonadState.idr | 20 +++++++++++++++++++ tests/base/data_ref_monadstate/a-test.ipkg | 1 + tests/base/data_ref_monadstate/expected | 5 +++++ tests/base/data_ref_monadstate/run | 3 +++ 6 files changed, 43 insertions(+) create mode 100644 tests/base/data_ref_monadstate/RefMonadState.idr create mode 100644 tests/base/data_ref_monadstate/a-test.ipkg create mode 100644 tests/base/data_ref_monadstate/expected create mode 100644 tests/base/data_ref_monadstate/run diff --git a/CHANGELOG.md b/CHANGELOG.md index a4941be8b6..59f0fa6e3c 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -213,6 +213,9 @@ * Implements `Ord` for `Count` from `Language.Reflection`. +* Implements `MonadState` for `Data.Ref` with a named implementation requiring + a particular reference. + #### System * Changes `getNProcessors` to return the number of online processors rather than diff --git a/libs/base/Data/Ref.idr b/libs/base/Data/Ref.idr index 4f44d32e4a..7bb99cc5f5 100644 --- a/libs/base/Data/Ref.idr +++ b/libs/base/Data/Ref.idr @@ -2,6 +2,7 @@ module Data.Ref import public Data.IORef import public Control.Monad.ST +import Control.Monad.State.Interface %default total @@ -22,3 +23,13 @@ Ref (ST s) (STRef s) where newRef = newSTRef readRef = readSTRef writeRef = writeSTRef + +namespace MonadState + + export + ForRef : Ref m r => Monad m => r a -> MonadState a m + ForRef ref = MS where + %inline + [MS] MonadState a m where + get = readRef ref + put = writeRef ref diff --git a/tests/base/data_ref_monadstate/RefMonadState.idr b/tests/base/data_ref_monadstate/RefMonadState.idr new file mode 100644 index 0000000000..43c1956ee4 --- /dev/null +++ b/tests/base/data_ref_monadstate/RefMonadState.idr @@ -0,0 +1,20 @@ +module RefMonadState + +import Control.Monad.State.Interface + +import Data.Ref + +%default total + +fancy : MonadState Nat m => HasIO m => m String +fancy = do + n <- get + putStrLn "current n: \{show n}" + put $ 20 + n + pure "`\{show n}`" + +main : IO () +main = do + ref <- newRef 18 + for_ [1..4] $ const $ fancy @{ForRef ref} + putStrLn "resulting value: \{show !(readRef ref)}" diff --git a/tests/base/data_ref_monadstate/a-test.ipkg b/tests/base/data_ref_monadstate/a-test.ipkg new file mode 100644 index 0000000000..2c5b5ab52d --- /dev/null +++ b/tests/base/data_ref_monadstate/a-test.ipkg @@ -0,0 +1 @@ +package a-test diff --git a/tests/base/data_ref_monadstate/expected b/tests/base/data_ref_monadstate/expected new file mode 100644 index 0000000000..b77294f68d --- /dev/null +++ b/tests/base/data_ref_monadstate/expected @@ -0,0 +1,5 @@ +current n: 18 +current n: 38 +current n: 58 +current n: 78 +resulting value: 98 diff --git a/tests/base/data_ref_monadstate/run b/tests/base/data_ref_monadstate/run new file mode 100644 index 0000000000..ad4e6126fa --- /dev/null +++ b/tests/base/data_ref_monadstate/run @@ -0,0 +1,3 @@ +. ../../testutils.sh + +run RefMonadState.idr From 7fbbb030df0cfd115e3f3e62b71479c3c44df571 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Stefan=20H=C3=B6ck?= Date: Fri, 13 Oct 2023 14:48:15 +0200 Subject: [PATCH 31/43] [ new ] add Data.List.grouped function (#3089) --- CHANGELOG.md | 2 ++ libs/base/Data/List.idr | 18 ++++++++++++++++++ 2 files changed, 20 insertions(+) diff --git a/CHANGELOG.md b/CHANGELOG.md index 59f0fa6e3c..5f2a5ae4b8 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -211,6 +211,8 @@ * Adds updating functions to `SortedMap` and `SortedDMap`. +* Adds `grouped` function to `Data.List` for splitting a list into equal-sized slices. + * Implements `Ord` for `Count` from `Language.Reflection`. * Implements `MonadState` for `Data.Ref` with a named implementation requiring diff --git a/libs/base/Data/List.idr b/libs/base/Data/List.idr index 9e6c8d2f48..2ff1608b19 100644 --- a/libs/base/Data/List.idr +++ b/libs/base/Data/List.idr @@ -925,6 +925,24 @@ public export groupAllWith : Ord b => (a -> b) -> List a -> List (List1 a) groupAllWith f = groupWith f . sortBy (comparing f) +||| Partitions a list into fixed sized sublists. +||| +||| Note: The last list in the result might be shorter than the rest if +||| the input cannot evenly be split into groups of the same size. +||| +||| ```idris example +||| grouped 3 [1..10] === [[1,2,3],[4,5,6],[7,8,9],[10]] +||| ``` +public export +grouped : (n : Nat) -> {auto 0 p : IsSucc n} -> List a -> List (List a) +grouped _ [] = [] +grouped (S m) (x::xs) = go [<] [ SnocList a -> Nat -> List a -> List (List a) + go sxs sx c [] = sxs <>> [sx <>> []] + go sxs sx 0 (x :: xs) = go (sxs :< (sx <>> [])) [ Date: Fri, 13 Oct 2023 09:26:24 -0500 Subject: [PATCH 32/43] [ base ] Add `Data.Vect.Quantifiers.All.remember`, the inverse to `forget` (#3096) Co-authored-by: Guillaume Allais --- CHANGELOG.md | 2 +- libs/base/Data/Vect.idr | 9 +++++++++ libs/base/Data/Vect/Quantifiers.idr | 26 +++++++++++++++++++++++++- 3 files changed, 35 insertions(+), 2 deletions(-) diff --git a/CHANGELOG.md b/CHANGELOG.md index 5f2a5ae4b8..fa1e22fc7a 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -187,7 +187,7 @@ * Generalized `imapProperty` in `Data.List.Quantifiers.All.All` and `Data.Vect.Quantifiers.All.All`. -* Add `zipPropertyWith` to `Data.Vect.Quantifiers.All.All`. +* Add `zipPropertyWith` and `remember` to `Data.Vect.Quantifiers.All.All`. * Add `anyToFin` to `Data.Vect.Quantifiers.Any`, converting the `Any` witness to the index into the corresponding element. diff --git a/libs/base/Data/Vect.idr b/libs/base/Data/Vect.idr index 92433bca93..413959e42d 100644 --- a/libs/base/Data/Vect.idr +++ b/libs/base/Data/Vect.idr @@ -49,6 +49,11 @@ Biinjective Vect.(::) where -- Indexing into vectors -------------------------------------------------------------------------------- +export +invertVectZ : (xs : Vect Z a) -> xs === [] +invertVectZ [] = Refl + + ||| All but the first element of the vector ||| ||| ```idris example @@ -67,6 +72,10 @@ public export head : Vect (S len) elem -> elem head (x::_) = x +export +invertVectS : (xs : Vect (S n) a) -> xs === head xs :: tail xs +invertVectS (_ :: _) = Refl + ||| The last element of the vector ||| ||| ```idris example diff --git a/libs/base/Data/Vect/Quantifiers.idr b/libs/base/Data/Vect/Quantifiers.idr index 6157bb27df..b8ed0b5965 100644 --- a/libs/base/Data/Vect/Quantifiers.idr +++ b/libs/base/Data/Vect/Quantifiers.idr @@ -136,11 +136,36 @@ namespace All imapProperty _ _ [] = [] imapProperty i f @{ix :: ixs} (x::xs) = f @{ix} x :: imapProperty i f @{ixs} xs + ||| If `All` witnesses a property that does not depend on the vector `xs` + ||| it's indexed by, then it is really a `Vect`. public export forget : All (const p) {n} xs -> Vect n p forget [] = [] forget (x::xs) = x :: forget xs + ||| Any `Vect` can be lifted to become an `All` + ||| witnessing the presence of elements of the `Vect`'s type. + public export + remember : (xs : Vect n ty) -> All (const ty) xs + remember [] = [] + remember (x :: xs) = x :: remember xs + + export + forgetRememberId : (xs : Vect n ty) -> forget (remember xs) = xs + forgetRememberId [] = Refl + forgetRememberId (x :: xs) = cong (x ::) (forgetRememberId xs) + + public export + castAllConst : {0 xs, ys : Vect n a} -> All (const ty) xs -> All (const ty) ys + castAllConst [] = rewrite invertVectZ ys in [] + castAllConst (x :: xs) = rewrite invertVectS ys in x :: castAllConst xs + + export + rememberForgetId : (vs : All (const ty) xs) -> + castAllConst (remember (forget vs)) === vs + rememberForgetId [] = Refl + rememberForgetId (x :: xs) = cong (x ::) (rememberForgetId xs) + export zipPropertyWith : (f : {0 x : a} -> p x -> q x -> r x) -> All p xs -> All q xs -> All r xs @@ -218,4 +243,3 @@ namespace All drop' 0 ys = rewrite minusZeroRight k in ys drop' (S k) [] = [] drop' (S k) (y :: ys) = drop' k ys - From 419a440dad1ca4c26f3c7a8ac2fb6c77d2e03812 Mon Sep 17 00:00:00 2001 From: Denis Buzdalov Date: Fri, 13 Oct 2023 17:26:42 +0300 Subject: [PATCH 33/43] [ impl ] Support `default` implicits in named implementations (#3100) --- CHANGELOG.md | 1 + src/Idris/Desugar.idr | 21 ++++++---- src/Idris/Elab/Implementation.idr | 29 ++++++++----- src/Idris/Parser.idr | 41 +++++++++++-------- src/Idris/Syntax.idr | 2 +- src/Idris/Syntax/Traversals.idr | 12 +++--- .../interface030/DefaultImplicitsInImpls.idr | 20 +++++++++ tests/idris2/interface/interface030/expected | 1 + tests/idris2/interface/interface030/run | 3 ++ 9 files changed, 88 insertions(+), 42 deletions(-) create mode 100644 tests/idris2/interface/interface030/DefaultImplicitsInImpls.idr create mode 100644 tests/idris2/interface/interface030/expected create mode 100644 tests/idris2/interface/interface030/run diff --git a/CHANGELOG.md b/CHANGELOG.md index fa1e22fc7a..7525577ded 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -20,6 +20,7 @@ * New `fromTTImp`, `fromName`, and `fromDecls` names for custom `TTImp`, `Name`, and `Decls` literals. * Call to `%macro`-functions do not require the `ElabReflection` extension. +* Default implicits are supported for named implementations. * Elaborator scripts were made to be able to access project files, allowing the support for type providers and similar stuff. diff --git a/src/Idris/Desugar.idr b/src/Idris/Desugar.idr index d8aa6e32df..b206b3fd98 100644 --- a/src/Idris/Desugar.idr +++ b/src/Idris/Desugar.idr @@ -911,6 +911,15 @@ mutual = do tms' <- traverse (desugar AnyExpr ps) tms pure (ForeignExport tms') + %inline + mapDesugarPiInfo : {auto s : Ref Syn SyntaxInfo} -> + {auto c : Ref Ctxt Defs} -> + {auto u : Ref UST UState} -> + {auto m : Ref MD Metadata} -> + {auto o : Ref ROpts REPLOpts} -> + List Name -> PiInfo PTerm -> Core (PiInfo RawImp) + mapDesugarPiInfo ps = PiInfo.traverse (desugar AnyExpr ps) + -- Given a high level declaration, return a list of TTImp declarations -- which process it, and update any necessary state on the way. export @@ -1019,9 +1028,10 @@ mutual desugarDecl ps (PImplementation fc vis fnopts pass is cons tn params impln nusing body) = do opts <- traverse (desugarFnOpt ps) fnopts - is' <- for is $ \ (fc, c,n,tm) => + is' <- for is $ \ (fc, c, n, pi, tm) => do tm' <- desugar AnyExpr ps tm - pure (fc, c, n, tm') + pi' <- mapDesugarPiInfo ps pi + pure (fc, c, n, pi', tm') cons' <- for cons $ \ (n, tm) => do tm' <- desugar AnyExpr ps tm pure (n, tm') @@ -1034,13 +1044,13 @@ mutual $ findUniqueBindableNames fc True ps [] let paramsb = map (doBind bnames) params' - let isb = map (\ (info, r, n, tm) => (info, r, n, doBind bnames tm)) is' + let isb = map (\ (info, r, n, p, tm) => (info, r, n, p, doBind bnames tm)) is' let consb = map (\(n, tm) => (n, doBind bnames tm)) cons' body' <- maybe (pure Nothing) (\b => do b' <- traverse (desugarDecl ps) b pure (Just (concat b'))) body - -- calculate the name of the interface, if it's not explicitly + -- calculate the name of the implementation, if it's not explicitly -- given. let impname = maybe (mkImplName fc tn paramsb) id impln @@ -1102,9 +1112,6 @@ mutual NS ns (DN str (MN ("__mk" ++ str) 0)) mkConName n = DN (show n) (MN ("__mk" ++ show n) 0) - mapDesugarPiInfo : List Name -> PiInfo PTerm -> Core (PiInfo RawImp) - mapDesugarPiInfo ps = traverse (desugar AnyExpr ps) - desugarDecl ps (PFixity fc vis fix prec opName) = do ctx <- get Ctxt -- We update the context of fixities by adding a namespaced fixity diff --git a/src/Idris/Elab/Implementation.idr b/src/Idris/Elab/Implementation.idr index 1dc4301f4a..b263072629 100644 --- a/src/Idris/Elab/Implementation.idr +++ b/src/Idris/Elab/Implementation.idr @@ -47,10 +47,10 @@ bindConstraints fc p [] ty = ty bindConstraints fc p ((n, ty) :: rest) sc = IPi fc top p n ty (bindConstraints fc p rest sc) -bindImpls : List (FC, RigCount, Name, RawImp) -> RawImp -> RawImp +bindImpls : List (FC, RigCount, Name, PiInfo RawImp, RawImp) -> RawImp -> RawImp bindImpls [] ty = ty -bindImpls ((fc, r, n, ty) :: rest) sc - = IPi fc r Implicit (Just n) ty (bindImpls rest sc) +bindImpls ((fc, r, n, p, ty) :: rest) sc + = IPi fc r p (Just n) ty (bindImpls rest sc) addDefaults : FC -> Name -> (params : List (Name, RawImp)) -> -- parameters have been specialised, use them! @@ -99,11 +99,16 @@ addDefaults fc impName params allms defs body getMethImps : {vars : _} -> {auto c : Ref Ctxt Defs} -> Env Term vars -> Term vars -> - Core (List (Name, RigCount, RawImp)) + Core (List (Name, RigCount, Maybe RawImp, RawImp)) getMethImps env (Bind fc x (Pi fc' c Implicit ty) sc) = do rty <- map (map rawName) $ unelabNoSugar env ty ts <- getMethImps (Pi fc' c Implicit ty :: env) sc - pure ((x, c, rty) :: ts) + pure ((x, c, Nothing, rty) :: ts) +getMethImps env (Bind fc x (Pi fc' c (DefImplicit def) ty) sc) + = do rty <- map (map rawName) $ unelabNoSugar env ty + rdef <- map (map rawName) $ unelabNoSugar env def + ts <- getMethImps (Pi fc' c (DefImplicit def) ty :: env) sc + pure ((x, c, Just rdef, rty) :: ts) getMethImps env tm = pure [] export @@ -115,7 +120,7 @@ elabImplementation : {vars : _} -> {auto o : Ref ROpts REPLOpts} -> FC -> Visibility -> List FnOpt -> Pass -> Env Term vars -> NestedNames vars -> - (implicits : List (FC, RigCount, Name, RawImp)) -> + (implicits : List (FC, RigCount, Name, PiInfo RawImp, RawImp)) -> (constraints : List (Maybe Name, RawImp)) -> Name -> (ps : List RawImp) -> @@ -346,7 +351,7 @@ elabImplementation {vars} ifc vis opts_in pass env nest is cons iname ps named i -- When applying the method in the field for the record, eta expand -- the expected arguments based on the field type, so that implicits get -- inserted in the right place - mkMethField : List (Name, RigCount, RawImp) -> + mkMethField : List (Name, a) -> List (Name, List (Name, RigCount, PiInfo RawImp)) -> (Name, Name, List (String, String), RigCount, Maybe TotalReq, RawImp) -> RawImp mkMethField methImps fldTys (topn, n, upds, c, treq, ty) @@ -378,16 +383,18 @@ elabImplementation {vars} ifc vis opts_in pass env nest is cons iname ps named i = do mn <- inCurrentNS (methName n) pure (dropNS n, IVar vfc mn) - bindImps : List (Name, RigCount, RawImp) -> RawImp -> RawImp + bindImps : List (Name, RigCount, Maybe RawImp, RawImp) -> RawImp -> RawImp bindImps [] ty = ty - bindImps ((n, c, t) :: ts) ty + bindImps ((n, c, Just def, t) :: ts) ty + = IPi vfc c (DefImplicit def) (Just n) t (bindImps ts ty) + bindImps ((n, c, Nothing, t) :: ts) ty = IPi vfc c Implicit (Just n) t (bindImps ts ty) -- Return method name, specialised method name, implicit name updates, -- and method type. Also return how the method name should be updated -- in later method types (specifically, putting implicits in) topMethType : List (Name, RawImp) -> - Name -> List (Name, RigCount, RawImp) -> + Name -> List (Name, RigCount, Maybe RawImp, RawImp) -> List String -> List Name -> List Name -> List Name -> Method -> Core ((Name, Name, List (String, String), RigCount, Maybe TotalReq, RawImp), @@ -445,7 +452,7 @@ elabImplementation {vars} ifc vis opts_in pass env nest is cons iname ps named i pure ((meth.name, n, upds, meth.count, meth.totalReq, mty), methupds') topMethTypes : List (Name, RawImp) -> - Name -> List (Name, RigCount, RawImp) -> + Name -> List (Name, RigCount, Maybe RawImp, RawImp) -> List String -> List Name -> List Name -> List Name -> List Method -> diff --git a/src/Idris/Parser.idr b/src/Idris/Parser.idr index d31360de44..30a7644653 100644 --- a/src/Idris/Parser.idr +++ b/src/Idris/Parser.idr @@ -1594,6 +1594,16 @@ recordConstructor fname n <- mustWork $ decoratedDataConstructorName fname pure (doc, n) +autoImplicitField : OriginDesc -> IndentInfo -> Rule (PiInfo t) +autoImplicitField fname _ = AutoImplicit <$ decoratedKeyword fname "auto" + +defImplicitField : OriginDesc -> IndentInfo -> Rule (PiInfo PTerm) +defImplicitField fname indents = do + decoratedKeyword fname "default" + commit + t <- simpleExpr fname indents + pure (DefImplicit t) + constraints : OriginDesc -> IndentInfo -> EmptyRule (List (Maybe Name, PTerm)) constraints fname indents = do tm <- appExpr pdef fname indents @@ -1610,15 +1620,22 @@ constraints fname indents pure ((Just n, tm) :: more) <|> pure [] -implBinds : OriginDesc -> IndentInfo -> EmptyRule (List (FC, RigCount, Name, PTerm)) -implBinds fname indents = concatMap (map adjust) <$> go where +implBinds : OriginDesc -> IndentInfo -> (namedImpl : Bool) -> EmptyRule (List (FC, RigCount, Name, PiInfo PTerm, PTerm)) +implBinds fname indents namedImpl = concatMap (map adjust) <$> go where - adjust : (RigCount, WithBounds Name, PTerm) -> (FC, RigCount, Name, PTerm) + adjust : (RigCount, WithBounds Name, a) -> (FC, RigCount, Name, a) adjust (r, wn, ty) = (virtualiseFC (boundToFC fname wn), r, wn.val, ty) - go : EmptyRule (List (List (RigCount, WithBounds Name, PTerm))) + isDefaultImplicit : PiInfo a -> Bool + isDefaultImplicit (DefImplicit _) = True + isDefaultImplicit _ = False + + go : EmptyRule (List (List (RigCount, WithBounds Name, PiInfo PTerm, PTerm))) go = do decoratedSymbol fname "{" - ns <- pibindListName fname indents + piInfo <- bounds $ option Implicit $ defImplicitField fname indents + when (not namedImpl && isDefaultImplicit piInfo.val) $ + fatalLoc piInfo.bounds "Default implicits are allowed only for named implementations" + ns <- map @{Compose} (\(rc, wb, n) => (rc, wb, piInfo.val, n)) $ pibindListName fname indents commitSymbol fname "}" commitSymbol fname "->" more <- go @@ -1667,7 +1684,7 @@ implDecl fname indents iname <- optional $ decoratedSymbol fname "[" *> decorate fname Function name <* decoratedSymbol fname "]" - impls <- implBinds fname indents + impls <- implBinds fname indents (isJust iname) cons <- constraints fname indents n <- decorate fname Typ name params <- many (continue indents *> simpleExpr fname indents) @@ -1685,7 +1702,7 @@ fieldDecl fname indents = do doc <- optDocumentation fname decoratedSymbol fname "{" commit - impl <- option Implicit (autoImplicitField <|> defImplicitField) + impl <- option Implicit (autoImplicitField fname indents <|> defImplicitField fname indents) fs <- fieldBody doc impl decoratedSymbol fname "}" atEnd indents @@ -1695,16 +1712,6 @@ fieldDecl fname indents atEnd indents pure fs where - autoImplicitField : Rule (PiInfo t) - autoImplicitField = AutoImplicit <$ decoratedKeyword fname "auto" - - defImplicitField : Rule (PiInfo PTerm) - defImplicitField = do - decoratedKeyword fname "default" - commit - t <- simpleExpr fname indents - pure (DefImplicit t) - fieldBody : String -> PiInfo PTerm -> Rule (List PField) fieldBody doc p = do b <- bounds (do diff --git a/src/Idris/Syntax.idr b/src/Idris/Syntax.idr index 8e746a0b3d..3abdb3d9ea 100644 --- a/src/Idris/Syntax.idr +++ b/src/Idris/Syntax.idr @@ -447,7 +447,7 @@ mutual PDecl' nm PImplementation : FC -> Visibility -> List PFnOpt -> Pass -> - (implicits : List (FC, RigCount, Name, PTerm' nm)) -> + (implicits : List (FC, RigCount, Name, PiInfo (PTerm' nm), PTerm' nm)) -> (constraints : List (Maybe Name, PTerm' nm)) -> Name -> (params : List (PTerm' nm)) -> diff --git a/src/Idris/Syntax/Traversals.idr b/src/Idris/Syntax/Traversals.idr index f57c13de63..b0648c5cbe 100644 --- a/src/Idris/Syntax/Traversals.idr +++ b/src/Idris/Syntax/Traversals.idr @@ -327,11 +327,11 @@ mapPTermM f = goPTerm where (::) . (\ c => (a, b, c)) <$> goPTerm t <*> go3TupledPTerms ts - goImplicits : List (x, y, z, PTerm' nm) -> - Core (List (x, y, z, PTerm' nm)) + goImplicits : List (x, y, z, PiInfo (PTerm' nm), PTerm' nm) -> + Core (List (x, y, z, PiInfo (PTerm' nm), PTerm' nm)) goImplicits [] = pure [] - goImplicits ((a, b, c, t) :: ts) = - ((::) . (a,b,c,)) <$> goPTerm t + goImplicits ((a, b, c, p, t) :: ts) = + ((::) . (a,b,c,)) <$> ((,) <$> goPiInfo p <*> goPTerm t) <*> goImplicits ts go4TupledPTerms : List (x, y, PiInfo (PTerm' nm), PTerm' nm) -> @@ -576,9 +576,9 @@ mapPTerm f = goPTerm where go3TupledPTerms [] = [] go3TupledPTerms ((a, b, t) :: ts) = (a, b, goPTerm t) :: go3TupledPTerms ts - goImplicits : List (x, y, z, PTerm' nm) -> List (x, y, z, PTerm' nm) + goImplicits : List (x, y, z, PiInfo (PTerm' nm), PTerm' nm) -> List (x, y, z, PiInfo (PTerm' nm), PTerm' nm) goImplicits [] = [] - goImplicits ((a, b, c, t) :: ts) = (a,b,c, goPTerm t) :: goImplicits ts + goImplicits ((a, b, c, p, t) :: ts) = (a,b,c, goPiInfo p, goPTerm t) :: goImplicits ts go4TupledPTerms : List (x, y, PiInfo (PTerm' nm), PTerm' nm) -> List (x, y, PiInfo (PTerm' nm), PTerm' nm) diff --git a/tests/idris2/interface/interface030/DefaultImplicitsInImpls.idr b/tests/idris2/interface/interface030/DefaultImplicitsInImpls.idr new file mode 100644 index 0000000000..0e196149f6 --- /dev/null +++ b/tests/idris2/interface/interface030/DefaultImplicitsInImpls.idr @@ -0,0 +1,20 @@ +module DefaultImplicitsInImpls + +import Data.Vect + +%default total + +[TunableImpl] {default False fancy : Bool} -> Show Nat where + show x = if fancy then "!!! hohoho !!!" else "no fancy" + +X0 : String +X0 = show @{TunableImpl} 5 + +x0Corr : X0 === "no fancy" +x0Corr = Refl + +X1 : String +X1 = show @{TunableImpl {fancy=True}} 5 + +x1Corr : X1 === "!!! hohoho !!!" +x1Corr = Refl diff --git a/tests/idris2/interface/interface030/expected b/tests/idris2/interface/interface030/expected new file mode 100644 index 0000000000..2eac914b83 --- /dev/null +++ b/tests/idris2/interface/interface030/expected @@ -0,0 +1 @@ +1/1: Building DefaultImplicitsInImpls (DefaultImplicitsInImpls.idr) diff --git a/tests/idris2/interface/interface030/run b/tests/idris2/interface/interface030/run new file mode 100644 index 0000000000..dcb75f8ca3 --- /dev/null +++ b/tests/idris2/interface/interface030/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check DefaultImplicitsInImpls.idr From c04404a95ba6bee1a60a1cbc2a7e31cc313eaf61 Mon Sep 17 00:00:00 2001 From: Steve Dunham Date: Fri, 13 Oct 2023 11:02:58 -0700 Subject: [PATCH 34/43] [ fix #3097 ] Fix issues parsing %logging followed by named impls (#3098) Co-authored-by: G. Allais --- src/Idris/Desugar.idr | 3 --- src/Idris/Parser.idr | 14 ++++---------- src/Idris/Syntax.idr | 3 --- src/Idris/Syntax/Traversals.idr | 2 -- tests/idris2/error/perror030/LoggingParse.idr | 7 +++++++ tests/idris2/error/perror030/expected | 1 + tests/idris2/error/perror030/run | 3 +++ 7 files changed, 15 insertions(+), 18 deletions(-) create mode 100644 tests/idris2/error/perror030/LoggingParse.idr create mode 100644 tests/idris2/error/perror030/expected create mode 100644 tests/idris2/error/perror030/run diff --git a/src/Idris/Desugar.idr b/src/Idris/Desugar.idr index b206b3fd98..f9196dd34c 100644 --- a/src/Idris/Desugar.idr +++ b/src/Idris/Desugar.idr @@ -979,9 +979,6 @@ mutual uds' <- traverse (desugarDecl ps) uds update Syn { usingImpl := oldu } pure (concat uds') - desugarDecl ps (PReflect fc tm) - = throw (GenericMsg fc "Reflection not implemented yet") --- pure [IReflect fc !(desugar AnyExpr ps tm)] desugarDecl ps (PInterface fc vis cons_in tn doc params det conname body) = do addDocString tn doc let paramNames = map fst params diff --git a/src/Idris/Parser.idr b/src/Idris/Parser.idr index 30a7644653..9fdfe742d1 100644 --- a/src/Idris/Parser.idr +++ b/src/Idris/Parser.idr @@ -1835,14 +1835,8 @@ fixDecl fname indents directiveDecl : OriginDesc -> IndentInfo -> Rule PDecl directiveDecl fname indents - = do b <- bounds ((do d <- directive fname indents - pure (\fc : FC => PDirective fc d)) - <|> - (do decoratedPragma fname "runElab" - tm <- expr pdef fname indents - atEnd indents - pure (\fc : FC => PReflect fc tm))) - pure (b.val (boundToFC fname b)) + = do b <- bounds (directive fname indents) + pure (PDirective (boundToFC fname b) b.val) -- Declared at the top -- topDecl : OriginDesc -> IndentInfo -> Rule (List PDecl) @@ -1855,6 +1849,8 @@ topDecl fname indents pure [d] <|> do ds <- claims fname indents pure (forget ds) + <|> do d <- directiveDecl fname indents + pure [d] <|> do d <- implDecl fname indents pure [d] <|> do d <- definition fname indents @@ -1880,8 +1876,6 @@ topDecl fname indents pure [d] <|> do d <- transformDecl fname indents pure [d] - <|> do d <- directiveDecl fname indents - pure [d] <|> do dstr <- bounds (terminal "Expected CG directive" (\case CGDirective d => Just d diff --git a/src/Idris/Syntax.idr b/src/Idris/Syntax.idr index 3abdb3d9ea..f273e8efbb 100644 --- a/src/Idris/Syntax.idr +++ b/src/Idris/Syntax.idr @@ -434,7 +434,6 @@ mutual List (PDecl' nm) -> PDecl' nm PUsing : FC -> List (Maybe Name, PTerm' nm) -> List (PDecl' nm) -> PDecl' nm - PReflect : FC -> PTerm' nm -> PDecl' nm PInterface : FC -> Visibility -> (constraints : List (Maybe Name, PTerm' nm)) -> @@ -482,7 +481,6 @@ mutual getPDeclLoc (PData fc _ _ _ _) = fc getPDeclLoc (PParameters fc _ _) = fc getPDeclLoc (PUsing fc _ _) = fc - getPDeclLoc (PReflect fc _) = fc getPDeclLoc (PInterface fc _ _ _ _ _ _ _ _) = fc getPDeclLoc (PImplementation fc _ _ _ _ _ _ _ _ _ _) = fc getPDeclLoc (PRecord fc _ _ _ _) = fc @@ -1055,7 +1053,6 @@ Show PDecl where show (PData{}) = "PData" show (PParameters{}) = "PParameters" show (PUsing{}) = "PUsing" - show (PReflect{}) = "PReflect" show (PInterface{}) = "PInterface" show (PImplementation{}) = "PImplementation" show (PRecord{}) = "PRecord" diff --git a/src/Idris/Syntax/Traversals.idr b/src/Idris/Syntax/Traversals.idr index b0648c5cbe..9db5d8fa22 100644 --- a/src/Idris/Syntax/Traversals.idr +++ b/src/Idris/Syntax/Traversals.idr @@ -240,7 +240,6 @@ mapPTermM f = goPTerm where goPDecl (PUsing fc mnts ps) = PUsing fc <$> goPairedPTerms mnts <*> goPDecls ps - goPDecl (PReflect fc t) = PReflect fc <$> goPTerm t goPDecl (PInterface fc v mnts n doc nrts ns mn ps) = PInterface fc v <$> goPairedPTerms mnts <*> pure n @@ -523,7 +522,6 @@ mapPTerm f = goPTerm where = PParameters fc (go4TupledPTerms nts) (goPDecl <$> ps) goPDecl (PUsing fc mnts ps) = PUsing fc (goPairedPTerms mnts) (goPDecl <$> ps) - goPDecl (PReflect fc t) = PReflect fc $ goPTerm t goPDecl (PInterface fc v mnts n doc nrts ns mn ps) = PInterface fc v (goPairedPTerms mnts) n doc (go3TupledPTerms nrts) ns mn (goPDecl <$> ps) goPDecl (PImplementation fc v opts p is cs n ts mn ns mps) diff --git a/tests/idris2/error/perror030/LoggingParse.idr b/tests/idris2/error/perror030/LoggingParse.idr new file mode 100644 index 0000000000..8d380151d5 --- /dev/null +++ b/tests/idris2/error/perror030/LoggingParse.idr @@ -0,0 +1,7 @@ +import Data.Vect + +%logging "elab" 1 + +[Named] {n : Nat} -> Show (Vect n a) where + show x = "whatever" + diff --git a/tests/idris2/error/perror030/expected b/tests/idris2/error/perror030/expected new file mode 100644 index 0000000000..342800e6fb --- /dev/null +++ b/tests/idris2/error/perror030/expected @@ -0,0 +1 @@ +1/1: Building LoggingParse (LoggingParse.idr) diff --git a/tests/idris2/error/perror030/run b/tests/idris2/error/perror030/run new file mode 100644 index 0000000000..193bf1acc2 --- /dev/null +++ b/tests/idris2/error/perror030/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check LoggingParse.idr From 3e5d8a54d4871ae1cb9dc96b2fd27c7ad7c592c4 Mon Sep 17 00:00:00 2001 From: Aleksei Volkov Date: Sun, 15 Oct 2023 23:30:48 +0300 Subject: [PATCH 35/43] [ fix ] Prevent relative path traversal in elaborator scripts --- src/TTImp/Elab/RunElab.idr | 1 + tests/idris2/reflection/reflection024/src/LessSimpleRW.idr | 5 +++++ 2 files changed, 6 insertions(+) diff --git a/src/TTImp/Elab/RunElab.idr b/src/TTImp/Elab/RunElab.idr index 4158990107..05b33be196 100644 --- a/src/TTImp/Elab/RunElab.idr +++ b/src/TTImp/Elab/RunElab.idr @@ -129,6 +129,7 @@ elabScript rig fc nest env script@(NDCon nfc nm t ar args) exp pathDoesNotEscape _ [] = True pathDoesNotEscape Z (".."::rest) = False pathDoesNotEscape (S n) (".."::rest) = pathDoesNotEscape n rest + pathDoesNotEscape n ("." ::rest) = pathDoesNotEscape n rest pathDoesNotEscape n (_ ::rest) = pathDoesNotEscape (S n) rest elabCon : Defs -> String -> List (Closure vars) -> Core (NF vars) diff --git a/tests/idris2/reflection/reflection024/src/LessSimpleRW.idr b/tests/idris2/reflection/reflection024/src/LessSimpleRW.idr index bb6336deb8..1543918fa0 100644 --- a/tests/idris2/reflection/reflection024/src/LessSimpleRW.idr +++ b/tests/idris2/reflection/reflection024/src/LessSimpleRW.idr @@ -35,3 +35,8 @@ failing "path must not escape the directory" -- Check a slightly more complicated case of escaping %runElab readAndLog "nonExistentOriginally/../../whatever" + +failing "path must not escape the directory" + + -- Check that '.' does not allow escaping + %runElab readAndLog "./../whatever" From 7c8076c1499fb268703d2fa1d6fb4412338a282c Mon Sep 17 00:00:00 2001 From: 0xd34df00d <0xd34df00d@gmail.com> Date: Sun, 15 Oct 2023 15:32:02 -0500 Subject: [PATCH 36/43] [ base ] Relevant and irrelevant traversals for Data.Vect.Quantifiers.All --- CHANGELOG.md | 3 ++- libs/base/Data/Vect/Quantifiers.idr | 22 ++++++++++++++++++++++ 2 files changed, 24 insertions(+), 1 deletion(-) diff --git a/CHANGELOG.md b/CHANGELOG.md index 7525577ded..10d4cd295b 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -188,7 +188,8 @@ * Generalized `imapProperty` in `Data.List.Quantifiers.All.All` and `Data.Vect.Quantifiers.All.All`. -* Add `zipPropertyWith` and `remember` to `Data.Vect.Quantifiers.All.All`. +* Add `zipPropertyWith`, `traverseProperty`, `traversePropertyRelevant` and `remember` + to `Data.Vect.Quantifiers.All.All`. * Add `anyToFin` to `Data.Vect.Quantifiers.Any`, converting the `Any` witness to the index into the corresponding element. diff --git a/libs/base/Data/Vect/Quantifiers.idr b/libs/base/Data/Vect/Quantifiers.idr index b8ed0b5965..e29183fdf7 100644 --- a/libs/base/Data/Vect/Quantifiers.idr +++ b/libs/base/Data/Vect/Quantifiers.idr @@ -173,6 +173,28 @@ namespace All zipPropertyWith f (px :: pxs) (qx :: qxs) = f px qx :: zipPropertyWith f pxs qxs + ||| A `Traversable`'s `traverse` for `All`, + ||| for traversals that don't care about the values of the associated `Vect`. + export + traverseProperty : Applicative f => + {0 xs : Vect n a} -> + (forall x. p x -> f (q x)) -> + All p xs -> + f (All q xs) + traverseProperty f [] = pure [] + traverseProperty f (x :: xs) = [| f x :: traverseProperty f xs |] + + ||| A `Traversable`'s `traverse` for `All`, + ||| in case the elements of the `Vect` that the `All` is proving `p` about are also needed. + export + traversePropertyRelevant : Applicative f => + {xs : Vect n a} -> + ((x : a) -> p x -> f (q x)) -> + All p xs -> + f (All q xs) + traversePropertyRelevant f [] = pure [] + traversePropertyRelevant f (x :: xs) = [| f _ x :: traversePropertyRelevant f xs |] + export All (Show . p) xs => Show (All p xs) where show pxs = "[" ++ show' "" pxs ++ "]" From 2358a74a292ca63dbf61938482aa3dec51f6257d Mon Sep 17 00:00:00 2001 From: Denis Buzdalov Date: Mon, 16 Oct 2023 20:21:00 +0300 Subject: [PATCH 37/43] [ base ] Implement `Zippable` for several standard types + small cleanup --- CHANGELOG.md | 4 ++++ libs/base/Data/Colist.idr | 19 +++--------------- libs/base/Data/Colist1.idr | 20 +++--------------- libs/base/Data/Either.idr | 11 ++++++++++ libs/base/Data/Maybe.idr | 13 ++++++++++++ libs/base/Data/SortedMap.idr | 8 ++++++++ libs/base/Data/Zippable.idr | 39 ++++++++++++++++++++++++++++++++++++ libs/contrib/Data/IMaybe.idr | 16 +++++++++++++++ 8 files changed, 97 insertions(+), 33 deletions(-) diff --git a/CHANGELOG.md b/CHANGELOG.md index 10d4cd295b..4cefe942d7 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -220,6 +220,10 @@ * Implements `MonadState` for `Data.Ref` with a named implementation requiring a particular reference. +* Adds implementations of `Zippable` to `Either`, `Pair`, `Maybe`, `SortedMap`. + +* Adds a `Compose` and `FromApplicative` named implementations for `Zippable`. + #### System * Changes `getNProcessors` to return the number of online processors rather than diff --git a/libs/base/Data/Colist.idr b/libs/base/Data/Colist.idr index edc86853ad..c8958b75c6 100644 --- a/libs/base/Data/Colist.idr +++ b/libs/base/Data/Colist.idr @@ -249,19 +249,6 @@ Applicative Colist where _ <*> [] = [] f :: fs <*> a :: as = f a :: (fs <*> as) -public export -Zippable Colist where - zipWith f as bs = [| f as bs |] - - zipWith3 f as bs cs = [| f as bs cs |] - - unzip xs = (map fst xs, map snd xs) - - unzip3 xs = ( map (\(a,_,_) => a) xs - , map (\(_,b,_) => b) xs - , map (\(_,_,c) => c) xs - ) - - unzipWith f = unzip . map f - - unzipWith3 f = unzip3 . map f +public export %hint %inline +ZippableColist : Zippable Colist +ZippableColist = FromApplicative diff --git a/libs/base/Data/Colist1.idr b/libs/base/Data/Colist1.idr index caef8dac49..6b30a91c60 100644 --- a/libs/base/Data/Colist1.idr +++ b/libs/base/Data/Colist1.idr @@ -196,23 +196,9 @@ Applicative Colist1 where (f ::: fs) <*> (a ::: as) = f a ::: (fs <*> as) -public export -Zippable Colist1 where - zipWith f (x ::: xs) (y ::: ys) = f x y ::: zipWith f xs ys - - zipWith3 f (x ::: xs) (y ::: ys) (z ::: zs) = - f x y z ::: zipWith3 f xs ys zs - - unzip xs = (map fst xs, map snd xs) - - unzip3 xs = ( map (\(a,_,_) => a) xs - , map (\(_,b,_) => b) xs - , map (\(_,_,c) => c) xs - ) - - unzipWith f = unzip . map f - - unzipWith3 f = unzip3 . map f +public export %hint %inline +ZippableColist1 : Zippable Colist1 +ZippableColist1 = FromApplicative -------------------------------------------------------------------------------- -- Interleavings diff --git a/libs/base/Data/Either.idr b/libs/base/Data/Either.idr index fa583eadc1..c0b17dd6a6 100644 --- a/libs/base/Data/Either.idr +++ b/libs/base/Data/Either.idr @@ -106,6 +106,17 @@ mirror : Either a b -> Either b a mirror (Left x) = Right x mirror (Right x) = Left x +public export +Zippable (Either a) where + zipWith f x y = [| f x y |] + zipWith3 f x y z = [| f x y z |] + + unzipWith f (Left e) = (Left e, Left e) + unzipWith f (Right xy) = let (x, y) = f xy in (Right x, Right y) + + unzipWith3 f (Left e) = (Left e, Left e, Left e) + unzipWith3 f (Right xyz) = let (x, y, z) = f xyz in (Right x, Right y, Right z) + -------------------------------------------------------------------------------- -- Bifunctor diff --git a/libs/base/Data/Maybe.idr b/libs/base/Data/Maybe.idr index c5a16edace..4c6e228f31 100644 --- a/libs/base/Data/Maybe.idr +++ b/libs/base/Data/Maybe.idr @@ -2,6 +2,8 @@ module Data.Maybe import Control.Function +import Data.Zippable + %default total public export @@ -84,3 +86,14 @@ namespace Monoid public export [Deep] Semigroup a => Monoid (Maybe a) using Semigroup.Deep where neutral = Nothing + +public export +Zippable Maybe where + zipWith f x y = [| f x y |] + zipWith3 f x y z = [| f x y z |] + + unzipWith f Nothing = (Nothing, Nothing) + unzipWith f (Just x) = let (a, b) = f x in (Just a, Just b) + + unzipWith3 f Nothing = (Nothing, Nothing, Nothing) + unzipWith3 f (Just x) = let (a, b, c) = f x in (Just a, Just b, Just c) diff --git a/libs/base/Data/SortedMap.idr b/libs/base/Data/SortedMap.idr index 6aaba88720..6f20f4f9af 100644 --- a/libs/base/Data/SortedMap.idr +++ b/libs/base/Data/SortedMap.idr @@ -1,6 +1,7 @@ module Data.SortedMap import Data.SortedMap.Dependent +import Data.Zippable %hide Prelude.toList @@ -94,6 +95,13 @@ export implementation Traversable (SortedMap k) where traverse f = map M . traverse f . unM +export +implementation Ord k => Zippable (SortedMap k) where + zipWith f mx my = fromList $ mapMaybe (\(k, x) => (k,) . f x <$> lookup k my) $ toList mx + zipWith3 f mx my mz = fromList $ mapMaybe (\(k, x) => (k,) .: f x <$> lookup k my <*> lookup k mz) $ toList mx + unzipWith f m = let m' = map f m in (map fst m', map snd m') + unzipWith3 f m = let m' = map f m in (map fst m', map (fst . snd) m', map (snd . snd) m') + ||| Merge two maps. When encountering duplicate keys, using a function to combine the values. ||| Uses the ordering of the first map given. export diff --git a/libs/base/Data/Zippable.idr b/libs/base/Data/Zippable.idr index d8ec1ebfb2..fcb74512e2 100644 --- a/libs/base/Data/Zippable.idr +++ b/libs/base/Data/Zippable.idr @@ -55,3 +55,42 @@ interface Zippable z where ||| @ z the parameterised type unzip3 : z (a, b, c) -> (z a, z b, z c) unzip3 = unzipWith3 id + +public export +[Compose] Zippable f => Zippable g => Zippable (f . g) where + zipWith f = zipWith $ zipWith f + zipWith3 f = zipWith3 $ zipWith3 f + unzipWith f = unzipWith $ unzipWith f + unzipWith3 f = unzipWith3 $ unzipWith3 f + + zip = zipWith zip + zip3 = zipWith3 zip3 + unzip = unzipWith unzip + unzip3 = unzipWith3 unzip3 + +||| Variant of composing and decomposing using existing `Applicative` implementation. +||| +||| This implementation is lazy during unzipping. +||| Caution! It may repeat the whole original work, each time the unzipped elements are used. +||| +||| Please be aware that for every `Applicative` has the same semantics as `Zippable`. +||| Consider `List` as a simple example. +||| However, this implementation is applicable for lazy data types or deferred computations. +public export +[FromApplicative] Applicative f => Zippable f where + zipWith f x y = [| f x y |] + zipWith3 f x y z = [| f x y z |] + + unzip u = (map fst u, map snd u) + unzip3 u = (map fst u, map (fst . snd) u, map (snd . snd) u) + unzipWith f = unzip . map f + unzipWith3 f = unzip3 . map f + + -- To be moved appropriately as soon as we have some module like `Data.Pair`, like other prelude types have. +public export +Semigroup a => Zippable (Pair a) where + zipWith f (l, x) (r, y) = (l <+> r, f x y) + zipWith3 f (l, x) (m, y) (r, z) = (l <+> m <+> r, f x y z) + + unzipWith f (l, x) = bimap (l,) (l,) (f x) + unzipWith3 f (l, x) = bimap (l,) (bimap (l,) (l,)) (f x) diff --git a/libs/contrib/Data/IMaybe.idr b/libs/contrib/Data/IMaybe.idr index 4870f65d7d..a8bf7680ab 100644 --- a/libs/contrib/Data/IMaybe.idr +++ b/libs/contrib/Data/IMaybe.idr @@ -1,6 +1,8 @@ ||| Version of Maybe indexed by an `isJust' boolean module Data.IMaybe +import Data.Zippable + %default total public export @@ -21,3 +23,17 @@ public export Applicative (IMaybe True) where pure = Just Just f <*> Just x = Just (f x) + +public export +Zippable (IMaybe b) where + zipWith f Nothing Nothing = Nothing + zipWith f (Just x) (Just y) = Just $ f x y + + zipWith3 f Nothing Nothing Nothing = Nothing + zipWith3 f (Just x) (Just y) (Just z) = Just $ f x y z + + unzipWith f Nothing = (Nothing, Nothing) + unzipWith f (Just x) = let (a, b) = f x in (Just a, Just b) + + unzipWith3 f Nothing = (Nothing, Nothing, Nothing) + unzipWith3 f (Just x) = let (a, b, c) = f x in (Just a, Just b, Just c) From f64047b9ac748e2031dc7afe0a3d27aa3491589c Mon Sep 17 00:00:00 2001 From: Denis Buzdalov Date: Tue, 17 Oct 2023 14:45:52 +0300 Subject: [PATCH 38/43] [ safe ] Set deriving escape hatches to be marked as `%unsafe` --- libs/base/Deriving/Common.idr | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/libs/base/Deriving/Common.idr b/libs/base/Deriving/Common.idr index c002b62832..b305fcf8f0 100644 --- a/libs/base/Deriving/Common.idr +++ b/libs/base/Deriving/Common.idr @@ -18,7 +18,7 @@ data IsFreeOf : (x : Name) -> (ty : TTImp) -> Type where TrustMeFO : IsFreeOf a x ||| We may need to manufacture proofs and so we provide the `assert` escape hatch. -export +export -- %unsafe -- uncomment as soon as 0.7.0 is released assert_IsFreeOf : IsFreeOf x ty assert_IsFreeOf = TrustMeFO @@ -165,7 +165,7 @@ data HasImplementation : (intf : a -> Type) -> TTImp -> Type where TrustMeHI : HasImplementation intf t ||| We may need to manufacture proofs and so we provide the `assert` escape hatch. -export +export -- %unsafe -- uncomment as soon as 0.7.0 is released assert_hasImplementation : HasImplementation intf t assert_hasImplementation = TrustMeHI From 6c351570871d01ab46ed27afaeeecc24938316eb Mon Sep 17 00:00:00 2001 From: Denis Buzdalov Date: Tue, 17 Oct 2023 14:47:18 +0300 Subject: [PATCH 39/43] [ ux ] Make `isType` fail with positioned errors --- libs/base/Deriving/Common.idr | 16 ++++++++-------- tests/base/deriving_traversable/expected | 2 +- 2 files changed, 9 insertions(+), 9 deletions(-) diff --git a/libs/base/Deriving/Common.idr b/libs/base/Deriving/Common.idr index b305fcf8f0..c18834b85b 100644 --- a/libs/base/Deriving/Common.idr +++ b/libs/base/Deriving/Common.idr @@ -48,16 +48,16 @@ wording Func = "a function name" wording (DataCon tag arity) = "a data constructor" wording (TyCon tag arity) = "a type constructor" -isTypeCon : Elaboration m => Name -> m (List (Name, TTImp)) -isTypeCon ty = do +isTypeCon : Elaboration m => FC -> Name -> m (List (Name, TTImp)) +isTypeCon fc ty = do [(_, MkNameInfo (TyCon _ _))] <- getInfo ty - | [] => fail "\{show ty} out of scope" - | [(_, MkNameInfo nt)] => fail "\{show ty} is \{wording nt} rather than a type constructor" - | _ => fail "\{show ty} is ambiguous" + | [] => failAt fc "\{show ty} out of scope" + | [(_, MkNameInfo nt)] => failAt fc "\{show ty} is \{wording nt} rather than a type constructor" + | _ => failAt fc "\{show ty} is ambiguous" cs <- getCons ty for cs $ \ n => do [(_, ty)] <- getType n - | _ => fail "\{show n} is ambiguous" + | _ => failAt fc "\{show n} is ambiguous" pure (n, ty) export @@ -65,7 +65,7 @@ isType : Elaboration m => TTImp -> m IsType isType = go Z [] where go : Nat -> List (Argument Name, Nat) -> TTImp -> m IsType - go idx acc (IVar _ n) = MkIsType n (map (map (minus idx . S)) acc) <$> isTypeCon n + go idx acc (IVar fc n) = MkIsType n (map (map (minus idx . S)) acc) <$> isTypeCon fc n go idx acc (IApp _ t (IVar _ nm)) = case nm of -- Unqualified: that's a local variable UN (Basic _) => go (S idx) ((Arg emptyFC nm, idx) :: acc) t @@ -78,7 +78,7 @@ isType = go Z [] where -- Unqualified: that's a local variable UN (Basic _) => go (S idx) ((AutoArg emptyFC nm, idx) :: acc) t _ => go (S idx) acc t - go idx acc t = fail "Expected a type constructor, got: \{show t}" + go idx acc t = failAt (getFC t) "Expected a type constructor, got: \{show t}" ------------------------------------------------------------------------------ -- Being a (data) constructor with a parameter diff --git a/tests/base/deriving_traversable/expected b/tests/base/deriving_traversable/expected index a1219727a6..647d830cf9 100644 --- a/tests/base/deriving_traversable/expected +++ b/tests/base/deriving_traversable/expected @@ -65,7 +65,7 @@ LOG derive.traversable.clauses:1: traverseIVect : {0 m : _} -> {0 a, b : Type} -> {0 f : Type -> Type} -> Applicative f => (a -> f b) -> IVect {n = m} a -> f (IVect {n = m} b) traverseIVect f (MkIVect x2) = MkIVect <$> (traverse f x2) LOG derive.traversable.clauses:1: - traverseEqMap : {0 key, eq : _} -> {0 a, b : Type} -> {0 f : Type -> Type} -> Applicative f => (a -> f b) -> EqMap key {{conArg:13919} = eq} a -> f (EqMap key {{conArg:13919} = eq} b) + traverseEqMap : {0 key, eq : _} -> {0 a, b : Type} -> {0 f : Type -> Type} -> Applicative f => (a -> f b) -> EqMap key {{conArg:13933} = eq} a -> f (EqMap key {{conArg:13933} = eq} b) traverseEqMap f (MkEqMap x3) = MkEqMap <$> (traverse (traverse f) x3) LOG derive.traversable.clauses:1: traverseTree : {0 l : _} -> {0 a, b : Type} -> {0 f : Type -> Type} -> Applicative f => (a -> f b) -> Tree l a -> f (Tree l b) From f694e5e2f0ceb38e9487fc59157f2e545c08eed4 Mon Sep 17 00:00:00 2001 From: Raffi Sanna Date: Wed, 18 Oct 2023 18:39:15 -0400 Subject: [PATCH 40/43] Use `do` notation in `some` --- libs/contrib/Data/String/Parser.idr | 4 +++- 1 file changed, 3 insertions(+), 1 deletion(-) diff --git a/libs/contrib/Data/String/Parser.idr b/libs/contrib/Data/String/Parser.idr index fb536432df..4de6f12f4c 100644 --- a/libs/contrib/Data/String/Parser.idr +++ b/libs/contrib/Data/String/Parser.idr @@ -208,7 +208,9 @@ mutual export covering some : Monad m => ParseT m a -> ParseT m (List a) - some p = [| p :: many p |] + some p = do r <- p + rs <- many p + pure $ r :: rs ||| Always succeeds, will accumulate the results of `p` in a list until it fails. export From fcecb165b9e243117711899b86e0370759097d6d Mon Sep 17 00:00:00 2001 From: CodingCellist Date: Mon, 23 Oct 2023 12:40:18 +0200 Subject: [PATCH 41/43] [ doc ] Document `failing` blocks (#3114) --- docs/source/reference/failing.rst | 120 ++++++++++++++++++++++++++++++ docs/source/reference/index.rst | 1 + 2 files changed, 121 insertions(+) create mode 100644 docs/source/reference/failing.rst diff --git a/docs/source/reference/failing.rst b/docs/source/reference/failing.rst new file mode 100644 index 0000000000..60ba951321 --- /dev/null +++ b/docs/source/reference/failing.rst @@ -0,0 +1,120 @@ +Failing blocks +============== + +.. role:: idris(code) + :language: idris + +Failing blocks let users check that some code parses but is rejected during +elaboration. Their simplest form is using the keyword ``failing`` followed by +some indented Idris code: + +.. code-block:: idris + + failing + trueNotFalse : True === False + trueNotFalse = Refl + +Putting the code above in a file and asking Idris to check it will not produce +an error message since the code in the ``failing`` block is rejected, i.e. +fails. + +The ``failing`` keyword optionally takes a string argument. If present, the +string attached to the failing block is checked against the error message raised +to check that it appears in the error. This lets the user document the kind of +error expected in the block, while at the same time checking that the error +message is indeed what we expected. For example, in the following example, we +make sure that Idris rejects a proof that the character ``'a'`` is equal to +``'b'`` by throwing an error when unifying them: + +.. code-block:: idris + + failing "When unifying" + noteq : 'a' === 'b' + noteq = Refl + +Failing blocks can be helpful when demonstrating false paths or starts in a +program or example. Valid code is accepted in failing blocks, so intermediary +helper functions can be used as long as at least one statement in the block +fails (it does not have to be the first or last statement!): + +.. code-block:: idris + + failing "Mismatch between: Integer and Double" + record Point where + constructor MkPoint + x : Integer + y : Integer + + Origin : Point + Origin = MkPoint 0 0 + + dist : Point -> Point -> Integer + dist p1 p2 = sqrt $ (pow (p2.x - p1.x) 2) + (pow (p2.y - p1.y) 2) + +Invalid failing blocks +---------------------- + +Should the failing block not fail, i.e. the code inside is accepted during +elaboration, the compiler will report an error: + +.. code-block:: idris + + failing + validRefl : 1 === 1 + validRefl = Refl + +Checking the above gives: + +:: + + Error: Failing block did not fail. + +Similarly, if an expected error (sub)string is given but the error output does +not match, we get: + +:: + + Error: Failing block failed with the wrong error. + +Followed by the given expected error string and the actual error output, +allowing us to compare and contrast. + +One block per falsehood +----------------------- + +**Take care to use separate ``failing`` blocks per thing to check.** Using a +single block can lead to unexpected results. The following example passes, since +the second statement fails: + +.. code-block:: idris + + failing + stmt1 : True === True + stmt1 = Refl + + -- at least one failing statement, so the block is accepted + stmt2 : 'a' === 'b' + stmt2 = Refl + + stmt3 : 1 === 1 + stmt3 = Refl + +Instead, separate the statements out into separate failing blocks: + +.. code-block:: idris + + failing + stmt1 : True === True + stmt1 = Refl + + failing + stmt2 : 'a' === 'b' + stmt2 = Refl + + failing + stmt3 : 1 === 1 + stmt3 = Refl + +This causes two ``Error: Failing block did not fail`` messages to be emitted, as +one would expect. + diff --git a/docs/source/reference/index.rst b/docs/source/reference/index.rst index 054e553f70..d1217d09c0 100644 --- a/docs/source/reference/index.rst +++ b/docs/source/reference/index.rst @@ -26,3 +26,4 @@ This is a placeholder, to get set up with readthedocs. pragmas builtins debugging + failing From 4097e6c99381dc16818c86f62db8ae123a1e384c Mon Sep 17 00:00:00 2001 From: Raffi Sanna Date: Wed, 18 Oct 2023 13:14:02 -0400 Subject: [PATCH 42/43] Switch from 'fast' string functions to normal string functions --- libs/contrib/Text/Distance/Levenshtein.idr | 2 +- libs/contrib/Text/Lexer/Core.idr | 6 +++--- libs/contrib/Text/Lexer/Tokenizer.idr | 6 +++--- libs/contrib/Text/Literate.idr | 2 +- libs/contrib/Text/PrettyPrint/Prettyprinter/Render/HTML.idr | 2 +- 5 files changed, 9 insertions(+), 9 deletions(-) diff --git a/libs/contrib/Text/Distance/Levenshtein.idr b/libs/contrib/Text/Distance/Levenshtein.idr index 83102786e0..903bb51dcc 100644 --- a/libs/contrib/Text/Distance/Levenshtein.idr +++ b/libs/contrib/Text/Distance/Levenshtein.idr @@ -10,7 +10,7 @@ import Data.List.Extra ||| Self-evidently correct but O(3 ^ (min mn)) complexity spec : String -> String -> Nat -spec a b = loop (fastUnpack a) (fastUnpack b) where +spec a b = loop (unpack a) (unpack b) where loop : List Char -> List Char -> Nat loop [] ys = length ys -- deletions diff --git a/libs/contrib/Text/Lexer/Core.idr b/libs/contrib/Text/Lexer/Core.idr index be91d372c4..6fde0660b9 100644 --- a/libs/contrib/Text/Lexer/Core.idr +++ b/libs/contrib/Text/Lexer/Core.idr @@ -159,7 +159,7 @@ tokenise pred line col acc tmap str Just (tok, rest) => let line' = line + cast (countNLs tok) col' = getCols tok col in - Just (MkBounded (fn (fastPack (reverse tok))) False (MkBounds line col line' col'), + Just (MkBounded (fn (pack (reverse tok))) False (MkBounds line col line' col'), line', col', rest) Nothing => getFirstToken ts str @@ -171,11 +171,11 @@ export lex : TokenMap a -> String -> (List (WithBounds a), (Int, Int, String)) lex tmap str = let (ts, (l, c, str')) = tokenise (const False) 0 0 [] tmap (unpack str) in - (ts, (l, c, fastPack str')) + (ts, (l, c, pack str')) export lexTo : (a -> Bool) -> TokenMap a -> String -> (List (WithBounds a), (Int, Int, String)) lexTo pred tmap str = let (ts, (l, c, str')) = tokenise pred 0 0 [] tmap (unpack str) in - (ts, (l, c, fastPack str')) + (ts, (l, c, pack str')) diff --git a/libs/contrib/Text/Lexer/Tokenizer.idr b/libs/contrib/Text/Lexer/Tokenizer.idr index f9d9bc8169..5fb4bb287b 100644 --- a/libs/contrib/Text/Lexer/Tokenizer.idr +++ b/libs/contrib/Text/Lexer/Tokenizer.idr @@ -110,7 +110,7 @@ tokenise reject tokenizer line col acc str | _ => Nothing line' = line + cast (countNLs token) col' = getCols token col - tokenStr = fastPack $ reverse token + tokenStr = pack $ reverse token in pure (tokenStr, line', col', rest) getFirstMatch : Tokenizer a -> List Char -> @@ -149,8 +149,8 @@ lexTo : Lexer -> (List (WithBounds a), (StopReason, Int, Int, String)) lexTo reject tokenizer str = let (ts, reason, (l, c, str')) = - tokenise reject tokenizer 0 0 [] (fastUnpack str) in - (ts, reason, (l, c, fastPack str')) + tokenise reject tokenizer 0 0 [] (unpack str) in + (ts, reason, (l, c, pack str')) ||| Given a tokenizer and an input string, return a list of recognised tokens, ||| and the line, column, and remainder of the input at the first point in the string diff --git a/libs/contrib/Text/Literate.idr b/libs/contrib/Text/Literate.idr index ca52b9ca31..eecf5fb6ae 100644 --- a/libs/contrib/Text/Literate.idr +++ b/libs/contrib/Text/Literate.idr @@ -78,7 +78,7 @@ namespace Compat ||| Merge the tokens into a single source file. reduce : List (WithBounds Token) -> List String -> String -reduce [] acc = fastConcat (reverse acc) +reduce [] acc = concat (reverse acc) reduce (MkBounded (Any x) _ _ :: rest) acc = -- newline will always be tokenized as a single token if x == "\n" diff --git a/libs/contrib/Text/PrettyPrint/Prettyprinter/Render/HTML.idr b/libs/contrib/Text/PrettyPrint/Prettyprinter/Render/HTML.idr index e772b79e9e..c457474c5f 100644 --- a/libs/contrib/Text/PrettyPrint/Prettyprinter/Render/HTML.idr +++ b/libs/contrib/Text/PrettyPrint/Prettyprinter/Render/HTML.idr @@ -6,7 +6,7 @@ import Data.String export htmlEscape : String -> String -htmlEscape s = fastConcat $ reverse $ go [] s +htmlEscape s = concat $ reverse $ go [] s where isSafe : Char -> Bool isSafe '"' = False From 46be3b80828420f354d47808bbb80e2b79021d58 Mon Sep 17 00:00:00 2001 From: CodingCellist Date: Mon, 23 Oct 2023 16:26:20 +0200 Subject: [PATCH 43/43] [ ci ] Update deploy-action in ci-idris2-and-libs.yml (#3115) --- .github/workflows/ci-idris2-and-libs.yml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/.github/workflows/ci-idris2-and-libs.yml b/.github/workflows/ci-idris2-and-libs.yml index 9cc21f6268..95e5b19021 100644 --- a/.github/workflows/ci-idris2-and-libs.yml +++ b/.github/workflows/ci-idris2-and-libs.yml @@ -762,7 +762,7 @@ jobs: cd - cp -r www/html/* .github/scripts/html/ - name: Deploy HTML - uses: JamesIves/github-pages-deploy-action@4.1.3 + uses: JamesIves/github-pages-deploy-action@v4.4.3 if: ${{ success() && env.IDRIS2_DEPLOY }} with: