Idris2/tests/base/system_file001/ReadFilePage.idr
Mathew Polzin f078d5f5dc
clean up some deprecations (#2057)
* deprecate Data.Nat.Order.decideLTE

* Add properties for LTE/GTE that produce the difference.

* remove deprecated function now that it is available in the base library.

* remove two deprecated lines.

* remove module deprecated since v0.4.0

* fix prelude reference to renamed primitive.

* finish removing Data.Num.Implementations

* remove deprecated dirEntry function.

* remove deprecated fastAppend. Update CHANGELOG.

* replace fastAppend in test case

* replace fastAppend uses in compiler.

* remove new properties that weren't actually very new.
2021-10-24 12:06:57 +01:00

45 lines
1.6 KiB
Idris

import System.File
import Data.String
putLines : List String -> IO ()
putLines = putStrLn . fastConcat
total
totalChecks : IO ()
totalChecks =
do Right (False, l1) <- readFilePage 0 (limit 0) "test.txt"
| Right (True, _) => putStrLn "Failed EOF check"
| Left err => putStrLn $ show err
putStrLn $ "empty: " ++ (show l1)
Right (False, l2) <- readFilePage 0 (limit 1) "test.txt"
| Right (True, _) => putStrLn "Failed EOF check"
| Left err => putStrLn $ show err
putLines l2
Right (False, l3) <- readFilePage 1 (limit 2) "test.txt"
| Right (True, _) => putStrLn "Failed EOF check"
| Left err => putStrLn $ show err
putLines l3
Right (True, l4) <- readFilePage 0 (limit 100) "test.txt"
| Right (False, _) => putStrLn "Failed EOF check"
| Left err => putStrLn $ show err
putLines l4
partial
main : IO ()
main = do totalChecks
Right (True, l5) <- readFilePage 0 forever "test.txt"
| Right (False, _) => putStrLn "Failed EOF check"
| Left err => putStrLn $ show err
putLines l5
Right (True, l6) <- readFilePage 2 forever "test.txt"
| Right (False, _) => putStrLn "Failed EOF check"
| Left err => putStrLn $ show err
putLines l6
Right (True, l7) <- readFilePage 100 forever "test.txt"
| Right (False, _) => putStrLn "Failed EOF check"
| Left err => putStrLn $ show err
putStrLn $ "empty: " ++ (show l7)
Right l8 <- readFile "test.txt"
| Left err => putStrLn $ show err
putStrLn l8