Add totality annotations to src and libs/{prelude, base}

This commit is contained in:
Fabián Heredia Montiel 2021-06-08 17:05:10 -05:00
parent 42d4c36833
commit dad4dcdaf8
71 changed files with 144 additions and 1 deletions

View File

@ -2,6 +2,8 @@ module Control.App
import Data.IORef
%default covering
||| `Error` is a type synonym for `Type`, specify for exception handling.
public export
Error : Type

View File

@ -2,6 +2,8 @@ module Control.App.Console
import public Control.App
%default total
public export
interface Console e where
putChar : Char -> App {l} e ()

View File

@ -6,6 +6,8 @@ import Data.List
import Data.Strings
import System.File
%default covering
toFileEx : FileError -> FileEx
toFileEx (GenericFileError i) = GenericFileEx i
toFileEx FileReadError = FileReadError

View File

@ -3,6 +3,8 @@ module Control.Applicative.Const
import Data.Contravariant
import Data.Bits
%default total
public export
record Const (a : Type) (b : Type) where
constructor MkConst

View File

@ -15,6 +15,8 @@ module Control.Monad.Error.Either
import Control.Monad.Trans
%default total
public export
data EitherT : (e : Type) -> (m : Type -> Type) -> (a : Type) -> Type where
MkEitherT : m (Either e a) -> EitherT e m a

View File

@ -9,6 +9,8 @@ import Control.Monad.State.State
import Control.Monad.Trans
import Control.Monad.Writer.CPS
%default total
||| The strategy of combining computations that can throw exceptions
||| by bypassing bound functions
||| from the point an exception is thrown to the point that it is handled.

View File

@ -2,6 +2,8 @@ module Control.Monad.Identity
import Data.Bits
%default total
public export
record Identity (a : Type) where
constructor Id

View File

@ -14,6 +14,8 @@ module Control.Monad.Maybe
import Control.Monad.Trans
import Data.Maybe
%default total
public export
data MaybeT : (m : Type -> Type) -> (a : Type) -> Type where
MkMaybeT : m (Maybe a) -> MaybeT m a

View File

@ -7,6 +7,8 @@ module Control.Monad.RWS.CPS
import Control.Monad.Identity
import Control.Monad.Trans
%default total
||| A monad transformer adding reading an environment of type `r`,
||| collecting an output of type `w` and updating a state of type `s`
||| to an inner monad `m`.

View File

@ -5,6 +5,8 @@ import Control.Monad.Reader.Interface
import Control.Monad.State.Interface
import Control.Monad.Writer.Interface
%default covering
public export
interface (MonadReader r m, MonadWriter w m, MonadState s m) =>
MonadRWS r w s m | m where

View File

@ -8,6 +8,8 @@ import Control.Monad.RWS.CPS
import Control.Monad.Trans
import Control.Monad.Writer.CPS
%default covering
||| A computation which runs in a static context and produces an output
public export
interface Monad m => MonadReader stateType m | m where

View File

@ -3,6 +3,8 @@ module Control.Monad.Reader.Reader
import Control.Monad.Identity
import Control.Monad.Trans
%default total
||| The transformer on which the Reader monad is based
public export
record ReaderT (stateType : Type) (m : Type -> Type) (a : Type) where

View File

@ -2,6 +2,8 @@ module Control.Monad.ST
import Data.IORef
%default covering
export
data STRef : Type -> Type -> Type where
MkSTRef : IORef a -> STRef s a

View File

@ -8,6 +8,8 @@ import Control.Monad.RWS.CPS
import Control.Monad.Trans
import Control.Monad.Writer.CPS
%default total
||| A computation which runs in a context and produces an output
public export
interface Monad m => MonadState stateType m | m where

View File

@ -3,6 +3,8 @@ module Control.Monad.State.State
import Control.Monad.Identity
import Control.Monad.Trans
%default total
||| The transformer on which the State monad is based
public export
record StateT (stateType : Type) (m : Type -> Type) (a : Type) where

View File

@ -1,5 +1,7 @@
module Control.Monad.Trans
%default total
public export
interface MonadTrans t where
lift : Monad m => m a -> t m a

View File

@ -7,6 +7,8 @@ module Control.Monad.Writer.CPS
import Control.Monad.Identity
import Control.Monad.Trans
%default total
||| A writer monad parameterized by:
|||
||| @w the output to accumulate.

View File

@ -8,6 +8,8 @@ import Control.Monad.RWS.CPS
import Control.Monad.Trans
import Control.Monad.Writer.CPS
%default total
||| MonadWriter interface
|||
||| tell is like tell on the MUD's it shouts to monad

View File

@ -3,6 +3,8 @@ module Control.WellFounded
import Data.Nat
import Data.List
%default total
public export
data Accessible : (rel : a -> a -> Type) -> (x : a) -> Type where
Access : (rec : (y : a) -> rel y x -> Accessible rel y) ->

View File

@ -1,6 +1,8 @@
||| Additional utility functions for the `Bifoldable` interface.
module Data.Bifoldable
%default total
||| Left associative monadic bifold over a structure.
public export
bifoldlM : (Bifoldable p, Monad m)

View File

@ -5,6 +5,8 @@ import System.File
import Data.List
%default total
-- Reading and writing binary buffers. Note that this primitives are unsafe,
-- in that they don't check that buffer locations are within bounds.
-- We really need a safe wrapper!
@ -224,11 +226,13 @@ getString buf loc len
= primIO (prim__getString buf loc len)
export
covering
bufferData : HasIO io => Buffer -> io (List Int)
bufferData buf
= do len <- rawSize buf
unpackTo [] len
where
covering
unpackTo : List Int -> Int -> io (List Int)
unpackTo acc 0 = pure acc
unpackTo acc loc

View File

@ -1,5 +1,7 @@
module Data.Contravariant
%default total
infixl 4 >$, >$<, >&<, $<
||| Contravariant functors

View File

@ -9,6 +9,8 @@ import Data.Nat.Order
import Decidable.Decidable
import Decidable.Order
%default total
using (k : Nat)
public export

View File

@ -3,6 +3,8 @@ module Data.IOArray
import Data.IOArray.Prims
import Data.List
%default total
export
record IOArray elem where
constructor MkIOArray

View File

@ -1,5 +1,7 @@
module Data.IOArray.Prims
%default total
export
data ArrayData : Type -> Type where [external]

View File

@ -1,5 +1,7 @@
module Data.IORef
%default total
-- Implemented externally
-- e.g., in Scheme, passed around as a box
data Mut : Type -> Type where [external]

View File

@ -1,5 +1,7 @@
module Data.Maybe
%default total
public export
isNothing : Maybe a -> Bool
isNothing Nothing = True

View File

@ -2,6 +2,8 @@ module Data.Morphisms
import Data.Contravariant
%default total
public export
record Morphism a b where
constructor Mor

View File

@ -8,6 +8,8 @@ import Decidable.Decidable
import Decidable.Equality
import Decidable.Order
%default total
public export
total LTEIsTransitive : (m : Nat) -> (n : Nat) -> (o : Nat) ->
LTE m n -> LTE n o ->

View File

@ -1,5 +1,7 @@
module Data.Primitives.Views
%default total
-- We need all the believe_mes and asserts throughout this file because we're
-- working with primitive here! We also have separate implementations per
-- primitive, rather than using interfaces, because we're only going to trust

View File

@ -3,6 +3,8 @@ module Data.Ref
import public Data.IORef
import public Control.Monad.ST
%default total
public export
interface Ref m r | m where
newRef : {0 a : Type} -> a -> m (r a)

View File

@ -2,6 +2,8 @@ module Data.Rel
import Data.Fun
%default total
||| Build an n-ary relation type from a Vect of Types
public export
Rel : Vect n Type -> Type

View File

@ -3,6 +3,8 @@ module Data.Vect.Elem
import Data.Vect
import Decidable.Equality
%default total
--------------------------------------------------------------------------------
-- Vector membership proof
--------------------------------------------------------------------------------

View File

@ -2,6 +2,8 @@ module Data.Vect.Quantifiers
import Data.Vect
%default total
||| A proof that some element of a vector satisfies some property
|||
||| @ p the property to be satsified

View File

@ -3,6 +3,8 @@ module Debug.Trace
import Prelude
import PrimIO
%default total
export
trace : (msg : String) -> (result : a) -> a
trace x val = unsafePerformIO (do putStrLn x; pure val)

View File

@ -3,6 +3,8 @@ module Decidable.Decidable
import Data.Rel
import Data.Fun
%default total
public export
isNo : Dec a -> Bool
isNo (Yes _) = False

View File

@ -6,6 +6,8 @@ import Data.Fin
import Data.Fun
import Data.Rel
%default total
%hide Prelude.Equal
--------------------------------------------------------------------------------

View File

@ -4,6 +4,8 @@ import public Data.So
import Data.List
import Data.Strings
%default total
support : String -> String
support fn = "C:" ++ fn ++ ", libidris2_support, idris_support.h"
@ -69,6 +71,7 @@ getEnv var
else pure (Just (prim__getString env))
export
covering
getEnvironment : HasIO io => io (List (String, String))
getEnvironment = getAllPairs 0 []
where

View File

@ -2,6 +2,8 @@ module System.Clock
import PrimIO
%default total
||| The various types of system clock available.
public export
data ClockType

View File

@ -1,5 +1,7 @@
module System.Concurrency
%default total
-- At the moment this is pretty fundamentally tied to the Scheme RTS.
-- Given that different back ends will have entirely different threading
-- models, it might be unavoidable, but we might want to think about possible

View File

@ -2,6 +2,8 @@ module System.Directory
import public System.File
%default total
public export
DirPtr : Type
DirPtr = AnyPtr

View File

@ -4,6 +4,8 @@
module System.FFI
%default total
export
data Struct : String -> -- C struct name
List (String, Type) -> -- field names and types

View File

@ -1,5 +1,7 @@
module System.Info
%default total
%extern prim__os : String
%extern prim__codegen : String

View File

@ -2,6 +2,8 @@ module System.REPL
import System.File
%default covering
||| A basic read-eval-print loop, maintaining a state
||| @ state the input state
||| @ prompt the prompt to show

View File

@ -1,5 +1,7 @@
module Builtin
%default total
-- The most primitive data types; things which are used by desugaring
-- Totality assertions

View File

@ -5,6 +5,8 @@ import Prelude.Basics
import Prelude.Num
import Prelude.Types
%default total
-----------
-- CASTS --
-----------

View File

@ -2,6 +2,8 @@ module PrimIO
import Builtin
%default total
public export
data IORes : Type -> Type where
MkIORes : (result : a) -> (1 x : %World) -> IORes a

View File

@ -10,6 +10,8 @@ import Libraries.Data.String.Extra
import Core.Directory
%default covering
%hide Data.Strings.lines
%hide Data.Strings.lines'
%hide Data.Strings.unlines

View File

@ -17,6 +17,7 @@ import public Core.Context
import Compiler.ES.RemoveUnused
import Compiler.ES.TailRec
%default covering
mutual
isNameUsed : Name -> NamedCExp -> Bool

View File

@ -10,6 +10,8 @@ import public Core.TT
import public Data.Vect
import public Data.List
%default covering
mutual
||| A tree representing an ES expression.
|||

View File

@ -17,6 +17,8 @@ import Data.Maybe
import Data.Strings
import Libraries.Data.String.Extra
%default covering
||| Compile a TT expression to Javascript
compileToJS : Ref Ctxt Defs ->
ClosedTerm -> Core String

View File

@ -17,6 +17,8 @@ import System.File
import Data.Maybe
%default covering
findNode : IO String
findNode = do
Nothing <- idrisGetEnv "NODE"

View File

@ -16,6 +16,8 @@ import Core.FC
import Debug.Trace
%default covering
mutual
usedNames : NamedCExp -> SortedSet Name
usedNames (NmLocal fc n) = empty

View File

@ -93,6 +93,8 @@ import Compiler.ES.ImperativeAst
import Debug.Trace
%default covering
data TailRecS : Type where
record TailSt where

View File

@ -9,6 +9,8 @@ import System
import Idris.Version
import Libraries.Utils.Path
%default total
findCC : IO String
findCC
= do Nothing <- getEnv "IDRIS2_CC"

View File

@ -23,6 +23,8 @@ import System.File
import Libraries.Utils.Hex
import Libraries.Utils.Path
%default covering
showcCleanStringChar : Char -> String -> String
showcCleanStringChar '+' = ("_plus" ++)
showcCleanStringChar '-' = ("__" ++)

View File

@ -16,6 +16,8 @@ import Data.List1
import Data.Vect
import Data.Maybe
%default covering
-- Compilation unit IDs are intended to be opaque,
-- just to be able to express dependencies via keys in a map and such.
export

View File

@ -17,6 +17,8 @@ import Idris.Package.Types
import Idris.Pretty
import Idris.Version
%default covering
getNS : Name -> String
getNS (NS ns _) = show ns
getNS _ = ""

View File

@ -31,6 +31,8 @@ import public Libraries.Text.PrettyPrint.Prettyprinter.Util
import Parser.Lexer.Source
%default covering
public export
data IdrisDocAnn
= TCon Name

View File

@ -7,6 +7,8 @@ import public Data.Maybe
import System
%default total
||| Environment variable used by Idris2 compiler
public export
record EnvDesc where

View File

@ -17,7 +17,7 @@ import Data.Maybe
import Libraries.Data.PosMap
%default covering
%default total
SExpable Decoration where
toSExp Typ = SExpList [ SymbolAtom "decor", SymbolAtom "type"]

View File

@ -11,6 +11,8 @@ import Libraries.Utils.Term
import System
%default total
getPageWidth : {auto o : Ref ROpts REPLOpts} -> Core PageWidth
getPageWidth = do
consoleWidth <- getConsoleWidth

View File

@ -53,6 +53,8 @@ import System
import System.File
import System.Directory
%default covering
export
fuzzySearch : {auto c : Ref Ctxt Defs}
-> {auto u : Ref UST UState}

View File

@ -1,6 +1,8 @@
||| Utilities functions for contitionally delaying values.
module Libraries.Control.Delayed
%default total
||| Type-level function for a conditionally infinite type.
public export
inf : Bool -> Type -> Type

View File

@ -2,6 +2,8 @@ module Libraries.Data.IOMatrix
import Data.IOArray.Prims
%default total
export
record IOMatrix a where
constructor MkIOMatrix

View File

@ -1,5 +1,7 @@
module Libraries.Data.SortedMap
%default total
-- TODO: write split
private

View File

@ -3,6 +3,8 @@ module Libraries.Data.SortedSet
import Data.Maybe
import Libraries.Data.SortedMap
%default total
export
data SortedSet k = SetWrapper (Data.SortedMap.SortedMap k ())

View File

@ -3,6 +3,8 @@ module Libraries.Text.PrettyPrint.Prettyprinter.Render.HTML
import Data.List
import Data.Strings
%default covering
export
htmlEscape : String -> String
htmlEscape s = fastAppend $ reverse $ go [] s

View File

@ -14,6 +14,8 @@ import Libraries.Text.Quantity
import System.Info
%default total
infixr 5 </>, />
infixr 7 <.>
@ -453,6 +455,7 @@ parent = map show . parent' . parse
||| parents "/etc/kernel" == ["/etc/kernel", "/etc", "/"]
||| ```
export
covering
parents : String -> List String
parents = map show . List.iterate parent' . parse

View File

@ -20,6 +20,8 @@ import TTImp.TTImp
import TTImp.Unelab
import TTImp.Utils
%default covering
export
elabScript : {vars : _} ->
{auto c : Ref Ctxt Defs} ->

View File

@ -18,6 +18,8 @@ import Core.UnifyState
import TTImp.TTImp
%default covering
showDefType : Def -> String
showDefType None = "undefined"
showDefType (PMDef {}) = "function"