Merge pull request #206 from edwinb/coverage-fix

Calculate compile time references even in case
This commit is contained in:
Edwin Brady 2020-05-30 22:15:18 +01:00 committed by GitHub
commit bb240d4390
No known key found for this signature in database
GPG Key ID: 4AEE18F83AFDEB23
5 changed files with 30 additions and 3 deletions

View File

@ -768,9 +768,9 @@ processDef opts nest env fc n_in cs_in
put Ctxt (record { toCompileCase $= (n ::) } defs)
atotal <- toResolvedNames (NS ["Builtin"] (UN "assert_total"))
calcRefs False atotal (Resolved nidx)
when (not (InCase `elem` opts)) $
do calcRefs False atotal (Resolved nidx)
sc <- calculateSizeChange fc n
do sc <- calculateSizeChange fc n
setSizeChange fc n sc
checkIfGuarded fc n

View File

@ -43,7 +43,7 @@ idrisTests
-- Coverage checking
"coverage001", "coverage002", "coverage003", "coverage004",
"coverage005", "coverage006", "coverage007", "coverage008",
"coverage009",
"coverage009", "coverage010",
-- Error messages
"error001", "error002", "error003", "error004", "error005",
"error006", "error007", "error008", "error009", "error010",

View File

@ -0,0 +1,21 @@
module Main
import Data.Strings
import Data.Vect
import System
%default total
ints : Vect 4 Int
ints = [1, 2, 3, 4]
main : IO ()
main = do
[_, arg] <- getArgs
| _ => do putStrLn "One argument expected."
exitFailure
let n = stringToNatOrZ arg
case natToFin n (length ints + 1) of
Nothing => do putStrLn "Invalid number."
exitFailure
Just (FS i) => putStrLn $ "Value: " ++ show (index i ints)

View File

@ -0,0 +1,3 @@
1/1: Building casetot (casetot.idr)
casetot.idr:12:1--13:1:main is not covering:
Calls non covering function Main.case block in 1921(290)

3
tests/idris2/coverage010/run Executable file
View File

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