{-# LANGUAGE BangPatterns #-}
{-# LANGUAGE NondecreasingIndentation #-}
{-# LANGUAGE TypeFamilies #-}
module Agda.TypeChecking.Monad.Signature where
import Prelude hiding (null)
import qualified Control.Monad.Fail as Fail
import Control.Monad.State
import Control.Monad.Reader
import Control.Monad.Writer
import Control.Monad.Trans.Maybe
import Control.Monad.Trans.Identity
import qualified Data.List as List
import Data.Set (Set)
import qualified Data.Set as Set
import qualified Data.Map as Map
import qualified Data.HashMap.Strict as HMap
import Data.Maybe
import Agda.Syntax.Abstract.Name
import Agda.Syntax.Abstract (Ren, ScopeCopyInfo(..))
import Agda.Syntax.Common
import Agda.Syntax.Internal as I
import Agda.Syntax.Internal.Names
import Agda.Syntax.Position
import Agda.Syntax.Treeless (Compiled(..), TTerm)
import Agda.TypeChecking.Monad.Base
import Agda.TypeChecking.Monad.Builtin
import Agda.TypeChecking.Monad.Debug
import Agda.TypeChecking.Monad.Context
import Agda.TypeChecking.Monad.Env
import Agda.TypeChecking.Monad.Mutual
import Agda.TypeChecking.Monad.Open
import Agda.TypeChecking.Monad.State
import Agda.TypeChecking.Monad.Trace
import Agda.TypeChecking.DropArgs
import Agda.TypeChecking.Warnings
import Agda.TypeChecking.Positivity.Occurrence
import Agda.TypeChecking.Substitute
import Agda.TypeChecking.CompiledClause
import Agda.TypeChecking.Coverage.SplitTree
import {-# SOURCE #-} Agda.TypeChecking.CompiledClause.Compile
import {-# SOURCE #-} Agda.TypeChecking.Polarity
import {-# SOURCE #-} Agda.TypeChecking.Pretty
import {-# SOURCE #-} Agda.TypeChecking.ProjectionLike
import {-# SOURCE #-} Agda.Compiler.Treeless.Erase
import {-# SOURCE #-} Agda.Main
import Agda.Utils.Either
import Agda.Utils.Except ( ExceptT )
import Agda.Utils.Functor
import Agda.Utils.Lens
import Agda.Utils.List
import Agda.Utils.ListT
import Agda.Utils.Maybe
import Agda.Utils.Monad
import Agda.Utils.Null
import Agda.Utils.Size
import Agda.Utils.Update
import Agda.Utils.Impossible
addConstant :: QName -> Definition -> TCM ()
addConstant :: QName -> Definition -> TCM ()
addConstant QName
q Definition
d = do
VerboseKey -> VerboseLevel -> TCM Doc -> TCM ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> VerboseLevel -> TCM Doc -> m ()
reportSDoc VerboseKey
"tc.signature" VerboseLevel
20 (TCM Doc -> TCM ()) -> TCM Doc -> TCM ()
forall a b. (a -> b) -> a -> b
$ TCM Doc
"adding constant " TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> QName -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty QName
q TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> TCM Doc
" to signature"
Telescope
tel <- TCMT IO Telescope
forall (m :: * -> *). (Applicative m, MonadTCEnv m) => m Telescope
getContextTelescope
let tel' :: Telescope
tel' = VerboseKey -> Telescope -> Telescope
forall a. VerboseKey -> Tele a -> Tele a
replaceEmptyName VerboseKey
"r" (Telescope -> Telescope) -> Telescope -> Telescope
forall a b. (a -> b) -> a -> b
$ Telescope -> Telescope
forall a. KillRange a => KillRangeT a
killRange (Telescope -> Telescope) -> Telescope -> Telescope
forall a b. (a -> b) -> a -> b
$ case Definition -> Defn
theDef Definition
d of
Constructor{} -> (Dom Type -> Dom Type) -> Telescope -> Telescope
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
fmap Dom Type -> Dom Type
forall a. LensHiding a => a -> a
hideOrKeepInstance Telescope
tel
Function{ funProjection :: Defn -> Maybe Projection
funProjection = Just Projection{ projProper :: Projection -> Maybe QName
projProper = Just{}, projIndex :: Projection -> VerboseLevel
projIndex = VerboseLevel
n } } ->
let fallback :: Telescope
fallback = (Dom Type -> Dom Type) -> Telescope -> Telescope
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
fmap Dom Type -> Dom Type
forall a. LensHiding a => a -> a
hideOrKeepInstance Telescope
tel in
if VerboseLevel
n VerboseLevel -> VerboseLevel -> Bool
forall a. Ord a => a -> a -> Bool
> VerboseLevel
0 then Telescope
fallback else
case [Dom (VerboseKey, Type)]
-> Maybe ([Dom (VerboseKey, Type)], Dom (VerboseKey, Type))
forall a. [a] -> Maybe ([a], a)
initLast ([Dom (VerboseKey, Type)]
-> Maybe ([Dom (VerboseKey, Type)], Dom (VerboseKey, Type)))
-> [Dom (VerboseKey, Type)]
-> Maybe ([Dom (VerboseKey, Type)], Dom (VerboseKey, Type))
forall a b. (a -> b) -> a -> b
$ Telescope -> [Dom (VerboseKey, Type)]
forall t. Tele (Dom t) -> [Dom (VerboseKey, t)]
telToList Telescope
tel of
Maybe ([Dom (VerboseKey, Type)], Dom (VerboseKey, Type))
Nothing -> Telescope
fallback
Just ([Dom (VerboseKey, Type)]
doms, Dom (VerboseKey, Type)
dom) -> [Dom (VerboseKey, Type)] -> Telescope
telFromList ([Dom (VerboseKey, Type)] -> Telescope)
-> [Dom (VerboseKey, Type)] -> Telescope
forall a b. (a -> b) -> a -> b
$ (Dom (VerboseKey, Type) -> Dom (VerboseKey, Type))
-> [Dom (VerboseKey, Type)] -> [Dom (VerboseKey, Type)]
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
fmap Dom (VerboseKey, Type) -> Dom (VerboseKey, Type)
forall a. LensHiding a => a -> a
hideOrKeepInstance [Dom (VerboseKey, Type)]
doms [Dom (VerboseKey, Type)]
-> [Dom (VerboseKey, Type)] -> [Dom (VerboseKey, Type)]
forall a. [a] -> [a] -> [a]
++ [Dom (VerboseKey, Type)
dom]
Defn
_ -> Telescope
tel
let d' :: Definition
d' = Telescope -> Definition -> Definition
forall t. Abstract t => Telescope -> t -> t
abstract Telescope
tel' (Definition -> Definition) -> Definition -> Definition
forall a b. (a -> b) -> a -> b
$ Definition
d { defName :: QName
defName = QName
q }
VerboseKey -> VerboseLevel -> TCM Doc -> TCM ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> VerboseLevel -> TCM Doc -> m ()
reportSDoc VerboseKey
"tc.signature" VerboseLevel
60 (TCM Doc -> TCM ()) -> TCM Doc -> TCM ()
forall a b. (a -> b) -> a -> b
$ TCM Doc
"lambda-lifted definition =" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<?> Definition -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty Definition
d'
(Signature -> Signature) -> TCM ()
modifySignature ((Signature -> Signature) -> TCM ())
-> (Signature -> Signature) -> TCM ()
forall a b. (a -> b) -> a -> b
$ (Definitions -> Definitions) -> Signature -> Signature
updateDefinitions ((Definitions -> Definitions) -> Signature -> Signature)
-> (Definitions -> Definitions) -> Signature -> Signature
forall a b. (a -> b) -> a -> b
$ (Definition -> Definition -> Definition)
-> QName -> Definition -> Definitions -> Definitions
forall k v.
(Eq k, Hashable k) =>
(v -> v -> v) -> k -> v -> HashMap k v -> HashMap k v
HMap.insertWith Definition -> Definition -> Definition
(+++) QName
q Definition
d'
MutualId
i <- TCM MutualId
currentOrFreshMutualBlock
MutualId -> QName -> TCM ()
setMutualBlock MutualId
i QName
q
where
Definition
new +++ :: Definition -> Definition -> Definition
+++ Definition
old = Definition
new { defDisplay :: [LocalDisplayForm]
defDisplay = Definition -> [LocalDisplayForm]
defDisplay Definition
new [LocalDisplayForm] -> [LocalDisplayForm] -> [LocalDisplayForm]
forall a. [a] -> [a] -> [a]
++ Definition -> [LocalDisplayForm]
defDisplay Definition
old
, defInstance :: Maybe QName
defInstance = Definition -> Maybe QName
defInstance Definition
new Maybe QName -> Maybe QName -> Maybe QName
forall (m :: * -> *) a. MonadPlus m => m a -> m a -> m a
`mplus` Definition -> Maybe QName
defInstance Definition
old }
setTerminates :: QName -> Bool -> TCM ()
setTerminates :: QName -> Bool -> TCM ()
setTerminates QName
q Bool
b = (Signature -> Signature) -> TCM ()
modifySignature ((Signature -> Signature) -> TCM ())
-> (Signature -> Signature) -> TCM ()
forall a b. (a -> b) -> a -> b
$ QName -> (Definition -> Definition) -> Signature -> Signature
updateDefinition QName
q ((Definition -> Definition) -> Signature -> Signature)
-> (Definition -> Definition) -> Signature -> Signature
forall a b. (a -> b) -> a -> b
$ (Defn -> Defn) -> Definition -> Definition
updateTheDef ((Defn -> Defn) -> Definition -> Definition)
-> (Defn -> Defn) -> Definition -> Definition
forall a b. (a -> b) -> a -> b
$ \case
def :: Defn
def@Function{} -> Defn
def { funTerminates :: Maybe Bool
funTerminates = Bool -> Maybe Bool
forall a. a -> Maybe a
Just Bool
b }
Defn
def -> Defn
def
setCompiledClauses :: QName -> CompiledClauses -> TCM ()
setCompiledClauses :: QName -> CompiledClauses -> TCM ()
setCompiledClauses QName
q CompiledClauses
cc = (Signature -> Signature) -> TCM ()
modifySignature ((Signature -> Signature) -> TCM ())
-> (Signature -> Signature) -> TCM ()
forall a b. (a -> b) -> a -> b
$ QName -> (Definition -> Definition) -> Signature -> Signature
updateDefinition QName
q ((Definition -> Definition) -> Signature -> Signature)
-> (Definition -> Definition) -> Signature -> Signature
forall a b. (a -> b) -> a -> b
$ (Defn -> Defn) -> Definition -> Definition
updateTheDef ((Defn -> Defn) -> Definition -> Definition)
-> (Defn -> Defn) -> Definition -> Definition
forall a b. (a -> b) -> a -> b
$ Defn -> Defn
setT
where
setT :: Defn -> Defn
setT def :: Defn
def@Function{} = Defn
def { funCompiled :: Maybe CompiledClauses
funCompiled = CompiledClauses -> Maybe CompiledClauses
forall a. a -> Maybe a
Just CompiledClauses
cc }
setT Defn
def = Defn
def
setSplitTree :: QName -> SplitTree -> TCM ()
setSplitTree :: QName -> SplitTree -> TCM ()
setSplitTree QName
q SplitTree
st = (Signature -> Signature) -> TCM ()
modifySignature ((Signature -> Signature) -> TCM ())
-> (Signature -> Signature) -> TCM ()
forall a b. (a -> b) -> a -> b
$ QName -> (Definition -> Definition) -> Signature -> Signature
updateDefinition QName
q ((Definition -> Definition) -> Signature -> Signature)
-> (Definition -> Definition) -> Signature -> Signature
forall a b. (a -> b) -> a -> b
$ (Defn -> Defn) -> Definition -> Definition
updateTheDef ((Defn -> Defn) -> Definition -> Definition)
-> (Defn -> Defn) -> Definition -> Definition
forall a b. (a -> b) -> a -> b
$ Defn -> Defn
setT
where
setT :: Defn -> Defn
setT def :: Defn
def@Function{} = Defn
def { funSplitTree :: Maybe SplitTree
funSplitTree = SplitTree -> Maybe SplitTree
forall a. a -> Maybe a
Just SplitTree
st }
setT Defn
def = Defn
def
modifyFunClauses :: QName -> ([Clause] -> [Clause]) -> TCM ()
modifyFunClauses :: QName -> ([Clause] -> [Clause]) -> TCM ()
modifyFunClauses QName
q [Clause] -> [Clause]
f =
(Signature -> Signature) -> TCM ()
modifySignature ((Signature -> Signature) -> TCM ())
-> (Signature -> Signature) -> TCM ()
forall a b. (a -> b) -> a -> b
$ QName -> (Definition -> Definition) -> Signature -> Signature
updateDefinition QName
q ((Definition -> Definition) -> Signature -> Signature)
-> (Definition -> Definition) -> Signature -> Signature
forall a b. (a -> b) -> a -> b
$ (Defn -> Defn) -> Definition -> Definition
updateTheDef ((Defn -> Defn) -> Definition -> Definition)
-> (Defn -> Defn) -> Definition -> Definition
forall a b. (a -> b) -> a -> b
$ ([Clause] -> [Clause]) -> Defn -> Defn
updateFunClauses [Clause] -> [Clause]
f
addClauses :: QName -> [Clause] -> TCM ()
addClauses :: QName -> [Clause] -> TCM ()
addClauses QName
q [Clause]
cls = do
Telescope
tel <- TCMT IO Telescope
forall (m :: * -> *). (Applicative m, MonadTCEnv m) => m Telescope
getContextTelescope
(Signature -> Signature) -> TCM ()
modifySignature ((Signature -> Signature) -> TCM ())
-> (Signature -> Signature) -> TCM ()
forall a b. (a -> b) -> a -> b
$ QName -> (Definition -> Definition) -> Signature -> Signature
updateDefinition QName
q ((Definition -> Definition) -> Signature -> Signature)
-> (Definition -> Definition) -> Signature -> Signature
forall a b. (a -> b) -> a -> b
$
(Defn -> Defn) -> Definition -> Definition
updateTheDef (([Clause] -> [Clause]) -> Defn -> Defn
updateFunClauses ([Clause] -> [Clause] -> [Clause]
forall a. [a] -> [a] -> [a]
++ Telescope -> [Clause] -> [Clause]
forall t. Abstract t => Telescope -> t -> t
abstract Telescope
tel [Clause]
cls))
(Definition -> Definition)
-> (Definition -> Definition) -> Definition -> Definition
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (Bool -> Bool) -> Definition -> Definition
updateDefCopatternLHS (Bool -> Bool -> Bool
|| [Clause] -> Bool
isCopatternLHS [Clause]
cls)
mkPragma :: String -> TCM CompilerPragma
mkPragma :: VerboseKey -> TCM CompilerPragma
mkPragma VerboseKey
s = Range -> VerboseKey -> CompilerPragma
CompilerPragma (Range -> VerboseKey -> CompilerPragma)
-> TCMT IO Range -> TCMT IO (VerboseKey -> CompilerPragma)
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> TCMT IO Range
forall (m :: * -> *). MonadTCEnv m => m Range
getCurrentRange TCMT IO (VerboseKey -> CompilerPragma)
-> TCMT IO VerboseKey -> TCM CompilerPragma
forall (f :: * -> *) a b. Applicative f => f (a -> b) -> f a -> f b
<*> VerboseKey -> TCMT IO VerboseKey
forall (f :: * -> *) a. Applicative f => a -> f a
pure VerboseKey
s
addPragma :: BackendName -> QName -> String -> TCM ()
addPragma :: VerboseKey -> QName -> VerboseKey -> TCM ()
addPragma VerboseKey
b QName
q VerboseKey
s = TCMT IO Bool -> TCM () -> TCM () -> TCM ()
forall (m :: * -> *) a. Monad m => m Bool -> m a -> m a -> m a
ifM TCMT IO Bool
erased
(Warning -> TCM ()
forall (m :: * -> *). MonadWarning m => Warning -> m ()
warning (Warning -> TCM ()) -> Warning -> TCM ()
forall a b. (a -> b) -> a -> b
$ VerboseKey -> QName -> Warning
PragmaCompileErased VerboseKey
b QName
q)
(TCM () -> TCM ()) -> TCM () -> TCM ()
forall a b. (a -> b) -> a -> b
$ do
CompilerPragma
pragma <- VerboseKey -> TCM CompilerPragma
mkPragma VerboseKey
s
(Signature -> Signature) -> TCM ()
modifySignature ((Signature -> Signature) -> TCM ())
-> (Signature -> Signature) -> TCM ()
forall a b. (a -> b) -> a -> b
$ QName -> (Definition -> Definition) -> Signature -> Signature
updateDefinition QName
q ((Definition -> Definition) -> Signature -> Signature)
-> (Definition -> Definition) -> Signature -> Signature
forall a b. (a -> b) -> a -> b
$ VerboseKey -> CompilerPragma -> Definition -> Definition
addCompilerPragma VerboseKey
b CompilerPragma
pragma
where
erased :: TCM Bool
erased :: TCMT IO Bool
erased = do
Defn
def <- Definition -> Defn
theDef (Definition -> Defn) -> TCMT IO Definition -> TCMT IO Defn
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> QName -> TCMT IO Definition
forall (m :: * -> *). HasConstInfo m => QName -> m Definition
getConstInfo QName
q
case Defn
def of
Function{} ->
Lens' (Maybe VerboseKey) TCEnv
-> (Maybe VerboseKey -> Maybe VerboseKey)
-> TCMT IO Bool
-> TCMT IO Bool
forall (m :: * -> *) a b.
MonadTCEnv m =>
Lens' a TCEnv -> (a -> a) -> m b -> m b
locallyTC Lens' (Maybe VerboseKey) TCEnv
eActiveBackendName (Maybe VerboseKey -> Maybe VerboseKey -> Maybe VerboseKey
forall a b. a -> b -> a
const (Maybe VerboseKey -> Maybe VerboseKey -> Maybe VerboseKey)
-> Maybe VerboseKey -> Maybe VerboseKey -> Maybe VerboseKey
forall a b. (a -> b) -> a -> b
$ VerboseKey -> Maybe VerboseKey
forall a. a -> Maybe a
Just VerboseKey
b) (TCMT IO Bool -> TCMT IO Bool) -> TCMT IO Bool -> TCMT IO Bool
forall a b. (a -> b) -> a -> b
$
Lens' [Backend] TCState
-> ([Backend] -> [Backend]) -> TCMT IO Bool -> TCMT IO Bool
forall (m :: * -> *) a b.
ReadTCState m =>
Lens' a TCState -> (a -> a) -> m b -> m b
locallyTCState Lens' [Backend] TCState
stBackends ([Backend] -> [Backend] -> [Backend]
forall a b. a -> b -> a
const ([Backend] -> [Backend] -> [Backend])
-> [Backend] -> [Backend] -> [Backend]
forall a b. (a -> b) -> a -> b
$ [Backend]
builtinBackends) (TCMT IO Bool -> TCMT IO Bool) -> TCMT IO Bool -> TCMT IO Bool
forall a b. (a -> b) -> a -> b
$
QName -> TCMT IO Bool
isErasable QName
q
Defn
_ -> Bool -> TCMT IO Bool
forall (f :: * -> *) a. Applicative f => a -> f a
pure Bool
False
getUniqueCompilerPragma :: BackendName -> QName -> TCM (Maybe CompilerPragma)
getUniqueCompilerPragma :: VerboseKey -> QName -> TCM (Maybe CompilerPragma)
getUniqueCompilerPragma VerboseKey
backend QName
q = do
[CompilerPragma]
ps <- VerboseKey -> Definition -> [CompilerPragma]
defCompilerPragmas VerboseKey
backend (Definition -> [CompilerPragma])
-> TCMT IO Definition -> TCMT IO [CompilerPragma]
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> QName -> TCMT IO Definition
forall (m :: * -> *). HasConstInfo m => QName -> m Definition
getConstInfo QName
q
case [CompilerPragma]
ps of
[] -> Maybe CompilerPragma -> TCM (Maybe CompilerPragma)
forall (m :: * -> *) a. Monad m => a -> m a
return Maybe CompilerPragma
forall a. Maybe a
Nothing
[CompilerPragma
p] -> Maybe CompilerPragma -> TCM (Maybe CompilerPragma)
forall (m :: * -> *) a. Monad m => a -> m a
return (Maybe CompilerPragma -> TCM (Maybe CompilerPragma))
-> Maybe CompilerPragma -> TCM (Maybe CompilerPragma)
forall a b. (a -> b) -> a -> b
$ CompilerPragma -> Maybe CompilerPragma
forall a. a -> Maybe a
Just CompilerPragma
p
(CompilerPragma
_:CompilerPragma
p1:[CompilerPragma]
_) ->
CompilerPragma
-> TCM (Maybe CompilerPragma) -> TCM (Maybe CompilerPragma)
forall (tcm :: * -> *) x a.
(MonadTCM tcm, ReadTCState tcm, MonadDebug tcm, HasRange x) =>
x -> tcm a -> tcm a
setCurrentRange CompilerPragma
p1 (TCM (Maybe CompilerPragma) -> TCM (Maybe CompilerPragma))
-> TCM (Maybe CompilerPragma) -> TCM (Maybe CompilerPragma)
forall a b. (a -> b) -> a -> b
$
Doc -> TCM (Maybe CompilerPragma)
forall (m :: * -> *) a.
(MonadTCEnv m, ReadTCState m, MonadError TCErr m) =>
Doc -> m a
genericDocError (Doc -> TCM (Maybe CompilerPragma))
-> TCM Doc -> TCM (Maybe CompilerPragma)
forall (m :: * -> *) a b. Monad m => (a -> m b) -> m a -> m b
=<< do
TCM Doc -> VerboseLevel -> TCM Doc -> TCM Doc
forall (m :: * -> *).
Applicative m =>
m Doc -> VerboseLevel -> m Doc -> m Doc
hang (VerboseKey -> TCM Doc
forall (m :: * -> *). Monad m => VerboseKey -> m Doc
text (VerboseKey
"Conflicting " VerboseKey -> VerboseKey -> VerboseKey
forall a. [a] -> [a] -> [a]
++ VerboseKey
backend VerboseKey -> VerboseKey -> VerboseKey
forall a. [a] -> [a] -> [a]
++ VerboseKey
" pragmas for") TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> QName -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty QName
q TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> TCM Doc
"at") VerboseLevel
2 (TCM Doc -> TCM Doc) -> TCM Doc -> TCM Doc
forall a b. (a -> b) -> a -> b
$
[TCM Doc] -> TCM Doc
forall (m :: * -> *). Monad m => [m Doc] -> m Doc
vcat [ TCM Doc
"-" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Range -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty (CompilerPragma -> Range
forall t. HasRange t => t -> Range
getRange CompilerPragma
p) | CompilerPragma
p <- [CompilerPragma]
ps ]
setFunctionFlag :: FunctionFlag -> Bool -> QName -> TCM ()
setFunctionFlag :: FunctionFlag -> Bool -> QName -> TCM ()
setFunctionFlag FunctionFlag
flag Bool
val QName
q = QName -> (Definition -> Definition) -> TCM ()
modifyGlobalDefinition QName
q ((Definition -> Definition) -> TCM ())
-> (Definition -> Definition) -> TCM ()
forall a b. (a -> b) -> a -> b
$ Lens' Bool Definition -> LensSet Bool Definition
forall i o. Lens' i o -> LensSet i o
set ((Defn -> f Defn) -> Definition -> f Definition
Lens' Defn Definition
theDefLens ((Defn -> f Defn) -> Definition -> f Definition)
-> ((Bool -> f Bool) -> Defn -> f Defn)
-> (Bool -> f Bool)
-> Definition
-> f Definition
forall b c a. (b -> c) -> (a -> b) -> a -> c
. FunctionFlag -> Lens' Bool Defn
funFlag FunctionFlag
flag) Bool
val
markStatic :: QName -> TCM ()
markStatic :: QName -> TCM ()
markStatic = FunctionFlag -> Bool -> QName -> TCM ()
setFunctionFlag FunctionFlag
FunStatic Bool
True
markInline :: Bool -> QName -> TCM ()
markInline :: Bool -> QName -> TCM ()
markInline Bool
b = FunctionFlag -> Bool -> QName -> TCM ()
setFunctionFlag FunctionFlag
FunInline Bool
b
markInjective :: QName -> TCM ()
markInjective :: QName -> TCM ()
markInjective QName
q = QName -> (Definition -> Definition) -> TCM ()
modifyGlobalDefinition QName
q ((Definition -> Definition) -> TCM ())
-> (Definition -> Definition) -> TCM ()
forall a b. (a -> b) -> a -> b
$ \Definition
def -> Definition
def { defInjective :: Bool
defInjective = Bool
True }
unionSignatures :: [Signature] -> Signature
unionSignatures :: [Signature] -> Signature
unionSignatures [Signature]
ss = (Signature -> Signature -> Signature)
-> Signature -> [Signature] -> Signature
forall (t :: * -> *) a b.
Foldable t =>
(a -> b -> b) -> b -> t a -> b
foldr Signature -> Signature -> Signature
unionSignature Signature
emptySignature [Signature]
ss
where
unionSignature :: Signature -> Signature -> Signature
unionSignature (Sig Sections
a Definitions
b RewriteRuleMap
c) (Sig Sections
a' Definitions
b' RewriteRuleMap
c') =
Sections -> Definitions -> RewriteRuleMap -> Signature
Sig (Sections -> Sections -> Sections
forall k a. Ord k => Map k a -> Map k a -> Map k a
Map.union Sections
a Sections
a')
(Definitions -> Definitions -> Definitions
forall k v.
(Eq k, Hashable k) =>
HashMap k v -> HashMap k v -> HashMap k v
HMap.union Definitions
b Definitions
b')
((RewriteRules -> RewriteRules -> RewriteRules)
-> RewriteRuleMap -> RewriteRuleMap -> RewriteRuleMap
forall k v.
(Eq k, Hashable k) =>
(v -> v -> v) -> HashMap k v -> HashMap k v -> HashMap k v
HMap.unionWith RewriteRules -> RewriteRules -> RewriteRules
forall a. Monoid a => a -> a -> a
mappend RewriteRuleMap
c RewriteRuleMap
c')
addSection :: ModuleName -> TCM ()
addSection :: ModuleName -> TCM ()
addSection ModuleName
m = do
Telescope
tel <- TCMT IO Telescope
forall (m :: * -> *). (Applicative m, MonadTCEnv m) => m Telescope
getContextTelescope
let sec :: Section
sec = Telescope -> Section
Section Telescope
tel
TCMT IO (Maybe Section) -> (Section -> TCM ()) -> TCM ()
forall (m :: * -> *) a.
Monad m =>
m (Maybe a) -> (a -> m ()) -> m ()
whenJustM (ModuleName -> TCMT IO (Maybe Section)
forall (m :: * -> *).
(Functor m, ReadTCState m) =>
ModuleName -> m (Maybe Section)
getSection ModuleName
m) ((Section -> TCM ()) -> TCM ()) -> (Section -> TCM ()) -> TCM ()
forall a b. (a -> b) -> a -> b
$ \ Section
sec' -> do
if (Section
sec Section -> Section -> Bool
forall a. Eq a => a -> a -> Bool
== Section
sec') then do
VerboseKey -> VerboseLevel -> TCM Doc -> TCM ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> VerboseLevel -> TCM Doc -> m ()
reportSDoc VerboseKey
"tc.section" VerboseLevel
10 (TCM Doc -> TCM ()) -> TCM Doc -> TCM ()
forall a b. (a -> b) -> a -> b
$ TCM Doc
"warning: redundantly adding existing section" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> ModuleName -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty ModuleName
m
VerboseKey -> VerboseLevel -> TCM Doc -> TCM ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> VerboseLevel -> TCM Doc -> m ()
reportSDoc VerboseKey
"tc.section" VerboseLevel
60 (TCM Doc -> TCM ()) -> TCM Doc -> TCM ()
forall a b. (a -> b) -> a -> b
$ TCM Doc
"with content" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Section -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty Section
sec
else do
VerboseKey -> VerboseLevel -> TCM Doc -> TCM ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> VerboseLevel -> TCM Doc -> m ()
reportSDoc VerboseKey
"impossible" VerboseLevel
10 (TCM Doc -> TCM ()) -> TCM Doc -> TCM ()
forall a b. (a -> b) -> a -> b
$ TCM Doc
"overwriting existing section" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> ModuleName -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty ModuleName
m
VerboseKey -> VerboseLevel -> TCM Doc -> TCM ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> VerboseLevel -> TCM Doc -> m ()
reportSDoc VerboseKey
"impossible" VerboseLevel
60 (TCM Doc -> TCM ()) -> TCM Doc -> TCM ()
forall a b. (a -> b) -> a -> b
$ TCM Doc
"of content " TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Section -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty Section
sec'
VerboseKey -> VerboseLevel -> TCM Doc -> TCM ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> VerboseLevel -> TCM Doc -> m ()
reportSDoc VerboseKey
"impossible" VerboseLevel
60 (TCM Doc -> TCM ()) -> TCM Doc -> TCM ()
forall a b. (a -> b) -> a -> b
$ TCM Doc
"with content" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Section -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty Section
sec
TCM ()
forall a. HasCallStack => a
__IMPOSSIBLE__
ModuleName -> TCM ()
setModuleCheckpoint ModuleName
m
(Signature -> Signature) -> TCM ()
modifySignature ((Signature -> Signature) -> TCM ())
-> (Signature -> Signature) -> TCM ()
forall a b. (a -> b) -> a -> b
$ Lens' Sections Signature -> LensMap Sections Signature
forall i o. Lens' i o -> LensMap i o
over Lens' Sections Signature
sigSections LensMap Sections Signature -> LensMap Sections Signature
forall a b. (a -> b) -> a -> b
$ ModuleName -> Section -> Sections -> Sections
forall k a. Ord k => k -> a -> Map k a -> Map k a
Map.insert ModuleName
m Section
sec
setModuleCheckpoint :: ModuleName -> TCM ()
setModuleCheckpoint :: ModuleName -> TCM ()
setModuleCheckpoint ModuleName
m = do
CheckpointId
chkpt <- Lens' CheckpointId TCEnv -> TCMT IO CheckpointId
forall (m :: * -> *) a. MonadTCEnv m => Lens' a TCEnv -> m a
viewTC Lens' CheckpointId TCEnv
eCurrentCheckpoint
Lens' (Map ModuleName CheckpointId) TCState
stModuleCheckpoints Lens' (Map ModuleName CheckpointId) TCState
-> (Map ModuleName CheckpointId -> Map ModuleName CheckpointId)
-> TCM ()
forall (m :: * -> *) a.
MonadTCState m =>
Lens' a TCState -> (a -> a) -> m ()
`modifyTCLens` ModuleName
-> CheckpointId
-> Map ModuleName CheckpointId
-> Map ModuleName CheckpointId
forall k a. Ord k => k -> a -> Map k a -> Map k a
Map.insert ModuleName
m CheckpointId
chkpt
{-# SPECIALIZE getSection :: ModuleName -> TCM (Maybe Section) #-}
{-# SPECIALIZE getSection :: ModuleName -> ReduceM (Maybe Section) #-}
getSection :: (Functor m, ReadTCState m) => ModuleName -> m (Maybe Section)
getSection :: ModuleName -> m (Maybe Section)
getSection ModuleName
m = do
Sections
sig <- (TCState -> Lens' Sections TCState -> Sections
forall o i. o -> Lens' i o -> i
^. (Signature -> f Signature) -> TCState -> f TCState
Lens' Signature TCState
stSignature ((Signature -> f Signature) -> TCState -> f TCState)
-> ((Sections -> f Sections) -> Signature -> f Signature)
-> (Sections -> f Sections)
-> TCState
-> f TCState
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (Sections -> f Sections) -> Signature -> f Signature
Lens' Sections Signature
sigSections) (TCState -> Sections) -> m TCState -> m Sections
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> m TCState
forall (m :: * -> *). ReadTCState m => m TCState
getTCState
Sections
isig <- (TCState -> Lens' Sections TCState -> Sections
forall o i. o -> Lens' i o -> i
^. (Signature -> f Signature) -> TCState -> f TCState
Lens' Signature TCState
stImports ((Signature -> f Signature) -> TCState -> f TCState)
-> ((Sections -> f Sections) -> Signature -> f Signature)
-> (Sections -> f Sections)
-> TCState
-> f TCState
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (Sections -> f Sections) -> Signature -> f Signature
Lens' Sections Signature
sigSections) (TCState -> Sections) -> m TCState -> m Sections
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> m TCState
forall (m :: * -> *). ReadTCState m => m TCState
getTCState
Maybe Section -> m (Maybe Section)
forall (m :: * -> *) a. Monad m => a -> m a
return (Maybe Section -> m (Maybe Section))
-> Maybe Section -> m (Maybe Section)
forall a b. (a -> b) -> a -> b
$ ModuleName -> Sections -> Maybe Section
forall k a. Ord k => k -> Map k a -> Maybe a
Map.lookup ModuleName
m Sections
sig Maybe Section -> Maybe Section -> Maybe Section
forall (m :: * -> *) a. MonadPlus m => m a -> m a -> m a
`mplus` ModuleName -> Sections -> Maybe Section
forall k a. Ord k => k -> Map k a -> Maybe a
Map.lookup ModuleName
m Sections
isig
{-# SPECIALIZE lookupSection :: ModuleName -> TCM Telescope #-}
{-# SPECIALIZE lookupSection :: ModuleName -> ReduceM Telescope #-}
lookupSection :: (Functor m, ReadTCState m) => ModuleName -> m Telescope
lookupSection :: ModuleName -> m Telescope
lookupSection ModuleName
m = Telescope -> (Section -> Telescope) -> Maybe Section -> Telescope
forall b a. b -> (a -> b) -> Maybe a -> b
maybe Telescope
forall a. Tele a
EmptyTel (Section -> Lens' Telescope Section -> Telescope
forall o i. o -> Lens' i o -> i
^. Lens' Telescope Section
secTelescope) (Maybe Section -> Telescope) -> m (Maybe Section) -> m Telescope
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> ModuleName -> m (Maybe Section)
forall (m :: * -> *).
(Functor m, ReadTCState m) =>
ModuleName -> m (Maybe Section)
getSection ModuleName
m
addDisplayForms :: QName -> TCM ()
addDisplayForms :: QName -> TCM ()
addDisplayForms QName
x = do
Definition
def <- QName -> TCMT IO Definition
forall (m :: * -> *). HasConstInfo m => QName -> m Definition
getConstInfo QName
x
[Arg Term]
args <- VerboseLevel -> [Arg Term] -> [Arg Term]
forall a. VerboseLevel -> [a] -> [a]
drop (Defn -> VerboseLevel
projectionArgs (Defn -> VerboseLevel) -> Defn -> VerboseLevel
forall a b. (a -> b) -> a -> b
$ Definition -> Defn
theDef Definition
def) ([Arg Term] -> [Arg Term])
-> TCMT IO [Arg Term] -> TCMT IO [Arg Term]
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> TCMT IO [Arg Term]
forall (m :: * -> *). (Applicative m, MonadTCEnv m) => m [Arg Term]
getContextArgs
[Arg Term] -> QName -> QName -> Elims -> TCM ()
add [Arg Term]
args QName
x QName
x (Elims -> TCM ()) -> Elims -> TCM ()
forall a b. (a -> b) -> a -> b
$ (Arg Term -> Elim' Term) -> [Arg Term] -> Elims
forall a b. (a -> b) -> [a] -> [b]
map Arg Term -> Elim' Term
forall a. Arg a -> Elim' a
Apply ([Arg Term] -> Elims) -> [Arg Term] -> Elims
forall a b. (a -> b) -> a -> b
$ VerboseLevel -> [Arg Term] -> [Arg Term]
forall t a. Subst t a => VerboseLevel -> a -> a
raise VerboseLevel
1 [Arg Term]
args
where
add :: [Arg Term] -> QName -> QName -> Elims -> TCM ()
add [Arg Term]
args QName
top QName
x Elims
es0 = do
Definition
def <- QName -> TCMT IO Definition
forall (m :: * -> *). HasConstInfo m => QName -> m Definition
getConstInfo QName
x
let cs :: [Clause]
cs = Definition -> [Clause]
defClauses Definition
def
isCopy :: Bool
isCopy = Definition -> Bool
defCopy Definition
def
case [Clause]
cs of
[ Clause
cl ] -> do
if Bool -> Bool
not Bool
isCopy
then QName -> TCM Doc -> TCM ()
forall (m :: * -> *) a.
(MonadDebug m, Pretty a) =>
a -> TCM Doc -> m ()
noDispForm QName
x TCM Doc
"not a copy" else do
if Bool -> Bool
not (Bool -> Bool) -> Bool -> Bool
forall a b. (a -> b) -> a -> b
$ (NamedArg (Pattern' DBPatVar) -> Bool)
-> [NamedArg (Pattern' DBPatVar)] -> Bool
forall (t :: * -> *) a. Foldable t => (a -> Bool) -> t a -> Bool
all (Pattern' DBPatVar -> Bool
forall x. Pattern' x -> Bool
isVar (Pattern' DBPatVar -> Bool)
-> (NamedArg (Pattern' DBPatVar) -> Pattern' DBPatVar)
-> NamedArg (Pattern' DBPatVar)
-> Bool
forall b c a. (b -> c) -> (a -> b) -> a -> c
. NamedArg (Pattern' DBPatVar) -> Pattern' DBPatVar
forall a. NamedArg a -> a
namedArg) ([NamedArg (Pattern' DBPatVar)] -> Bool)
-> [NamedArg (Pattern' DBPatVar)] -> Bool
forall a b. (a -> b) -> a -> b
$ Clause -> [NamedArg (Pattern' DBPatVar)]
namedClausePats Clause
cl
then QName -> TCM Doc -> TCM ()
forall (m :: * -> *) a.
(MonadDebug m, Pretty a) =>
a -> TCM Doc -> m ()
noDispForm QName
x TCM Doc
"properly matching patterns" else do
let n :: VerboseLevel
n = [NamedArg (Pattern' DBPatVar)] -> VerboseLevel
forall a. Sized a => a -> VerboseLevel
size ([NamedArg (Pattern' DBPatVar)] -> VerboseLevel)
-> [NamedArg (Pattern' DBPatVar)] -> VerboseLevel
forall a b. (a -> b) -> a -> b
$ Clause -> [NamedArg (Pattern' DBPatVar)]
namedClausePats Clause
cl
(Elims
es1, Elims
es2) = VerboseLevel -> Elims -> (Elims, Elims)
forall a. VerboseLevel -> [a] -> ([a], [a])
splitAt VerboseLevel
n Elims
es0
m :: VerboseLevel
m = VerboseLevel
n VerboseLevel -> VerboseLevel -> VerboseLevel
forall a. Num a => a -> a -> a
- Elims -> VerboseLevel
forall a. Sized a => a -> VerboseLevel
size Elims
es1
vs1 :: [Term]
vs1 = (Arg Term -> Term) -> [Arg Term] -> [Term]
forall a b. (a -> b) -> [a] -> [b]
map Arg Term -> Term
forall e. Arg e -> e
unArg ([Arg Term] -> [Term]) -> [Arg Term] -> [Term]
forall a b. (a -> b) -> a -> b
$ [Arg Term] -> Maybe [Arg Term] -> [Arg Term]
forall a. a -> Maybe a -> a
fromMaybe [Arg Term]
forall a. HasCallStack => a
__IMPOSSIBLE__ (Maybe [Arg Term] -> [Arg Term]) -> Maybe [Arg Term] -> [Arg Term]
forall a b. (a -> b) -> a -> b
$ Elims -> Maybe [Arg Term]
forall a. [Elim' a] -> Maybe [Arg a]
allApplyElims Elims
es1
sub :: Substitution' Term
sub = [Term] -> Substitution' Term
forall a. DeBruijn a => [a] -> Substitution' a
parallelS ([Term] -> Substitution' Term) -> [Term] -> Substitution' Term
forall a b. (a -> b) -> a -> b
$ [Term] -> [Term]
forall a. [a] -> [a]
reverse ([Term] -> [Term]) -> [Term] -> [Term]
forall a b. (a -> b) -> a -> b
$ [Term]
vs1 [Term] -> [Term] -> [Term]
forall a. [a] -> [a] -> [a]
++ VerboseLevel -> Term -> [Term]
forall a. VerboseLevel -> a -> [a]
replicate VerboseLevel
m (VerboseLevel -> Term
var VerboseLevel
0)
body :: Maybe Term
body = Substitution' Term -> Maybe Term -> Maybe Term
forall t a. Subst t a => Substitution' t -> a -> a
applySubst Substitution' Term
sub (Clause -> Maybe Term
compiledClauseBody Clause
cl) Maybe Term -> Elims -> Maybe Term
forall t. Apply t => t -> Elims -> t
`applyE` Elims
es2
case Term -> Term
unSpine (Term -> Term) -> Maybe Term -> Maybe Term
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> Maybe Term
body of
Just (Def QName
y Elims
es) -> do
let df :: DisplayForm
df = VerboseLevel -> Elims -> DisplayTerm -> DisplayForm
Display VerboseLevel
m Elims
es (DisplayTerm -> DisplayForm) -> DisplayTerm -> DisplayForm
forall a b. (a -> b) -> a -> b
$ Term -> DisplayTerm
DTerm (Term -> DisplayTerm) -> Term -> DisplayTerm
forall a b. (a -> b) -> a -> b
$ QName -> Elims -> Term
Def QName
top (Elims -> Term) -> Elims -> Term
forall a b. (a -> b) -> a -> b
$ (Arg Term -> Elim' Term) -> [Arg Term] -> Elims
forall a b. (a -> b) -> [a] -> [b]
map Arg Term -> Elim' Term
forall a. Arg a -> Elim' a
Apply [Arg Term]
args
VerboseKey -> VerboseLevel -> TCM Doc -> TCM ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> VerboseLevel -> TCM Doc -> m ()
reportSDoc VerboseKey
"tc.display.section" VerboseLevel
20 (TCM Doc -> TCM ()) -> TCM Doc -> TCM ()
forall a b. (a -> b) -> a -> b
$ [TCM Doc] -> TCM Doc
forall (m :: * -> *). Monad m => [m Doc] -> m Doc
vcat
[ TCM Doc
"adding display form " TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> QName -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty QName
y TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> TCM Doc
" --> " TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> QName -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty QName
top
, VerboseKey -> TCM Doc
forall (m :: * -> *). Monad m => VerboseKey -> m Doc
text (DisplayForm -> VerboseKey
forall a. Show a => a -> VerboseKey
show DisplayForm
df)
]
QName -> DisplayForm -> TCM ()
addDisplayForm QName
y DisplayForm
df
[Arg Term] -> QName -> QName -> Elims -> TCM ()
add [Arg Term]
args QName
top QName
y Elims
es
Just Term
v -> QName -> TCM Doc -> TCM ()
forall (m :: * -> *) a.
(MonadDebug m, Pretty a) =>
a -> TCM Doc -> m ()
noDispForm QName
x (TCM Doc -> TCM ()) -> TCM Doc -> TCM ()
forall a b. (a -> b) -> a -> b
$ TCM Doc
"not a def body, but " TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Term -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty Term
v
Maybe Term
Nothing -> QName -> TCM Doc -> TCM ()
forall (m :: * -> *) a.
(MonadDebug m, Pretty a) =>
a -> TCM Doc -> m ()
noDispForm QName
x (TCM Doc -> TCM ()) -> TCM Doc -> TCM ()
forall a b. (a -> b) -> a -> b
$ TCM Doc
"bad body"
[] | Constructor{ conSrcCon :: Defn -> ConHead
conSrcCon = ConHead
h } <- Definition -> Defn
theDef Definition
def -> do
let y :: QName
y = ConHead -> QName
conName ConHead
h
df :: DisplayForm
df = VerboseLevel -> Elims -> DisplayTerm -> DisplayForm
Display VerboseLevel
0 [] (DisplayTerm -> DisplayForm) -> DisplayTerm -> DisplayForm
forall a b. (a -> b) -> a -> b
$ Term -> DisplayTerm
DTerm (Term -> DisplayTerm) -> Term -> DisplayTerm
forall a b. (a -> b) -> a -> b
$ ConHead -> ConInfo -> Elims -> Term
Con (ConHead
h {conName :: QName
conName = QName
top }) ConInfo
ConOSystem []
VerboseKey -> VerboseLevel -> TCM Doc -> TCM ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> VerboseLevel -> TCM Doc -> m ()
reportSDoc VerboseKey
"tc.display.section" VerboseLevel
20 (TCM Doc -> TCM ()) -> TCM Doc -> TCM ()
forall a b. (a -> b) -> a -> b
$ [TCM Doc] -> TCM Doc
forall (m :: * -> *). Monad m => [m Doc] -> m Doc
vcat
[ TCM Doc
"adding display form" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> QName -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty QName
y TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> TCM Doc
"-->" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> QName -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty QName
top
, VerboseKey -> TCM Doc
forall (m :: * -> *). Monad m => VerboseKey -> m Doc
text (DisplayForm -> VerboseKey
forall a. Show a => a -> VerboseKey
show DisplayForm
df)
]
QName -> DisplayForm -> TCM ()
addDisplayForm QName
y DisplayForm
df
[] -> QName -> TCM Doc -> TCM ()
forall (m :: * -> *) a.
(MonadDebug m, Pretty a) =>
a -> TCM Doc -> m ()
noDispForm QName
x TCM Doc
"no clauses"
(Clause
_:Clause
_:[Clause]
_) -> QName -> TCM Doc -> TCM ()
forall (m :: * -> *) a.
(MonadDebug m, Pretty a) =>
a -> TCM Doc -> m ()
noDispForm QName
x TCM Doc
"many clauses"
noDispForm :: a -> TCM Doc -> m ()
noDispForm a
x TCM Doc
reason = VerboseKey -> VerboseLevel -> TCM Doc -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> VerboseLevel -> TCM Doc -> m ()
reportSDoc VerboseKey
"tc.display.section" VerboseLevel
30 (TCM Doc -> m ()) -> TCM Doc -> m ()
forall a b. (a -> b) -> a -> b
$
TCM Doc
"no display form from" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> a -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty a
x TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> TCM Doc
"because" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> TCM Doc
reason
isVar :: Pattern' x -> Bool
isVar VarP{} = Bool
True
isVar Pattern' x
_ = Bool
False
applySection
:: ModuleName
-> Telescope
-> ModuleName
-> Args
-> ScopeCopyInfo
-> TCM ()
applySection :: ModuleName
-> Telescope -> ModuleName -> [Arg Term] -> ScopeCopyInfo -> TCM ()
applySection ModuleName
new Telescope
ptel ModuleName
old [Arg Term]
ts ScopeCopyInfo{ renModules :: ScopeCopyInfo -> Ren ModuleName
renModules = Ren ModuleName
rm, renNames :: ScopeCopyInfo -> Ren QName
renNames = Ren QName
rd } = do
Ren QName
rd <- Ren QName -> TCM (Ren QName)
closeConstructors Ren QName
rd
ModuleName
-> Telescope -> ModuleName -> [Arg Term] -> ScopeCopyInfo -> TCM ()
applySection' ModuleName
new Telescope
ptel ModuleName
old [Arg Term]
ts ScopeCopyInfo :: Ren ModuleName -> Ren QName -> ScopeCopyInfo
ScopeCopyInfo{ renModules :: Ren ModuleName
renModules = Ren ModuleName
rm, renNames :: Ren QName
renNames = Ren QName
rd }
where
closeConstructors :: Ren QName -> TCM (Ren QName)
closeConstructors :: Ren QName -> TCM (Ren QName)
closeConstructors Ren QName
rd = do
[QName]
ds <- (QName -> QName) -> [QName] -> [QName]
forall b a. Ord b => (a -> b) -> [a] -> [a]
nubOn QName -> QName
forall a. a -> a
id ([QName] -> [QName])
-> ([Maybe QName] -> [QName]) -> [Maybe QName] -> [QName]
forall b c a. (b -> c) -> (a -> b) -> a -> c
. [Maybe QName] -> [QName]
forall a. [Maybe a] -> [a]
catMaybes ([Maybe QName] -> [QName])
-> TCMT IO [Maybe QName] -> TCMT IO [QName]
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> ((QName, QName) -> TCMT IO (Maybe QName))
-> Ren QName -> TCMT IO [Maybe QName]
forall (t :: * -> *) (m :: * -> *) a b.
(Traversable t, Monad m) =>
(a -> m b) -> t a -> m (t b)
mapM (QName -> TCMT IO (Maybe QName)
constructorData (QName -> TCMT IO (Maybe QName))
-> ((QName, QName) -> QName)
-> (QName, QName)
-> TCMT IO (Maybe QName)
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (QName, QName) -> QName
forall a b. (a, b) -> a
fst) Ren QName
rd
[QName]
cs <- (QName -> QName) -> [QName] -> [QName]
forall b a. Ord b => (a -> b) -> [a] -> [a]
nubOn QName -> QName
forall a. a -> a
id ([QName] -> [QName])
-> ([[QName]] -> [QName]) -> [[QName]] -> [QName]
forall b c a. (b -> c) -> (a -> b) -> a -> c
. [[QName]] -> [QName]
forall (t :: * -> *) a. Foldable t => t [a] -> [a]
concat ([[QName]] -> [QName]) -> TCMT IO [[QName]] -> TCMT IO [QName]
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> ((QName, QName) -> TCMT IO [QName])
-> Ren QName -> TCMT IO [[QName]]
forall (t :: * -> *) (m :: * -> *) a b.
(Traversable t, Monad m) =>
(a -> m b) -> t a -> m (t b)
mapM (QName -> TCMT IO [QName]
dataConstructors (QName -> TCMT IO [QName])
-> ((QName, QName) -> QName) -> (QName, QName) -> TCMT IO [QName]
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (QName, QName) -> QName
forall a b. (a, b) -> a
fst) Ren QName
rd
Ren QName
new <- [Ren QName] -> Ren QName
forall (t :: * -> *) a. Foldable t => t [a] -> [a]
concat ([Ren QName] -> Ren QName)
-> TCMT IO [Ren QName] -> TCM (Ren QName)
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> (QName -> TCM (Ren QName)) -> [QName] -> TCMT IO [Ren QName]
forall (t :: * -> *) (m :: * -> *) a b.
(Traversable t, Monad m) =>
(a -> m b) -> t a -> m (t b)
mapM QName -> TCM (Ren QName)
rename ([QName]
ds [QName] -> [QName] -> [QName]
forall a. [a] -> [a] -> [a]
++ [QName]
cs)
VerboseKey -> VerboseLevel -> TCM Doc -> TCM ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> VerboseLevel -> TCM Doc -> m ()
reportSDoc VerboseKey
"tc.mod.apply.complete" VerboseLevel
30 (TCM Doc -> TCM ()) -> TCM Doc -> TCM ()
forall a b. (a -> b) -> a -> b
$
TCM Doc
"also copying: " TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Ren QName -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty Ren QName
new
Ren QName -> TCM (Ren QName)
forall (m :: * -> *) a. Monad m => a -> m a
return (Ren QName -> TCM (Ren QName)) -> Ren QName -> TCM (Ren QName)
forall a b. (a -> b) -> a -> b
$ Ren QName
new Ren QName -> Ren QName -> Ren QName
forall a. [a] -> [a] -> [a]
++ Ren QName
rd
where
rename :: QName -> TCM (Ren QName)
rename :: QName -> TCM (Ren QName)
rename QName
x =
case QName -> Ren QName -> Maybe QName
forall a b. Eq a => a -> [(a, b)] -> Maybe b
lookup QName
x Ren QName
rd of
Maybe QName
Nothing -> do Name
y <- VerboseKey -> TCMT IO Name
forall a (m :: * -> *).
(FreshName a, MonadFresh NameId m) =>
a -> m Name
freshName_ (Name -> VerboseKey
forall a. Show a => a -> VerboseKey
show (Name -> VerboseKey) -> Name -> VerboseKey
forall a b. (a -> b) -> a -> b
$ QName -> Name
qnameName QName
x)
Ren QName -> TCM (Ren QName)
forall (m :: * -> *) a. Monad m => a -> m a
return [(QName
x, [Name] -> QName
qnameFromList [Name
y])]
Just{} -> Ren QName -> TCM (Ren QName)
forall (m :: * -> *) a. Monad m => a -> m a
return []
constructorData :: QName -> TCM (Maybe QName)
constructorData :: QName -> TCMT IO (Maybe QName)
constructorData QName
x = do
(Definition -> Defn
theDef (Definition -> Defn) -> TCMT IO Definition -> TCMT IO Defn
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> QName -> TCMT IO Definition
forall (m :: * -> *). HasConstInfo m => QName -> m Definition
getConstInfo QName
x) TCMT IO Defn -> (Defn -> Maybe QName) -> TCMT IO (Maybe QName)
forall (m :: * -> *) a b. Functor m => m a -> (a -> b) -> m b
<&> \case
Constructor{ conData :: Defn -> QName
conData = QName
d } -> QName -> Maybe QName
forall a. a -> Maybe a
Just QName
d
Defn
_ -> Maybe QName
forall a. Maybe a
Nothing
dataConstructors :: QName -> TCM [QName]
dataConstructors :: QName -> TCMT IO [QName]
dataConstructors QName
x = do
(Definition -> Defn
theDef (Definition -> Defn) -> TCMT IO Definition -> TCMT IO Defn
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> QName -> TCMT IO Definition
forall (m :: * -> *). HasConstInfo m => QName -> m Definition
getConstInfo QName
x) TCMT IO Defn -> (Defn -> [QName]) -> TCMT IO [QName]
forall (m :: * -> *) a b. Functor m => m a -> (a -> b) -> m b
<&> \case
Datatype{ dataCons :: Defn -> [QName]
dataCons = [QName]
cs } -> [QName]
cs
Record{ recConHead :: Defn -> ConHead
recConHead = ConHead
h } -> [ConHead -> QName
conName ConHead
h]
Defn
_ -> []
applySection' :: ModuleName -> Telescope -> ModuleName -> Args -> ScopeCopyInfo -> TCM ()
applySection' :: ModuleName
-> Telescope -> ModuleName -> [Arg Term] -> ScopeCopyInfo -> TCM ()
applySection' ModuleName
new Telescope
ptel ModuleName
old [Arg Term]
ts ScopeCopyInfo{ renNames :: ScopeCopyInfo -> Ren QName
renNames = Ren QName
rd, renModules :: ScopeCopyInfo -> Ren ModuleName
renModules = Ren ModuleName
rm } = do
do
[QName]
noCopyList <- [Maybe QName] -> [QName]
forall a. [Maybe a] -> [a]
catMaybes ([Maybe QName] -> [QName])
-> TCMT IO [Maybe QName] -> TCMT IO [QName]
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> (VerboseKey -> TCMT IO (Maybe QName))
-> [VerboseKey] -> TCMT IO [Maybe QName]
forall (t :: * -> *) (m :: * -> *) a b.
(Traversable t, Monad m) =>
(a -> m b) -> t a -> m (t b)
mapM VerboseKey -> TCMT IO (Maybe QName)
forall (m :: * -> *).
HasBuiltins m =>
VerboseKey -> m (Maybe QName)
getName' [VerboseKey]
constrainedPrims
[QName] -> (QName -> TCM ()) -> TCM ()
forall (t :: * -> *) (m :: * -> *) a b.
(Foldable t, Monad m) =>
t a -> (a -> m b) -> m ()
forM_ (((QName, QName) -> QName) -> Ren QName -> [QName]
forall a b. (a -> b) -> [a] -> [b]
map (QName, QName) -> QName
forall a b. (a, b) -> a
fst Ren QName
rd) ((QName -> TCM ()) -> TCM ()) -> (QName -> TCM ()) -> TCM ()
forall a b. (a -> b) -> a -> b
$ \ QName
q -> do
Bool -> TCM () -> TCM ()
forall (f :: * -> *). Applicative f => Bool -> f () -> f ()
when (QName
q QName -> [QName] -> Bool
forall (t :: * -> *) a. (Foldable t, Eq a) => a -> t a -> Bool
`elem` [QName]
noCopyList) (TCM () -> TCM ()) -> TCM () -> TCM ()
forall a b. (a -> b) -> a -> b
$ TypeError -> TCM ()
forall (m :: * -> *) a.
(MonadTCEnv m, ReadTCState m, MonadError TCErr m) =>
TypeError -> m a
typeError (QName -> TypeError
TriedToCopyConstrainedPrim QName
q)
VerboseKey -> VerboseLevel -> TCM Doc -> TCM ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> VerboseLevel -> TCM Doc -> m ()
reportSDoc VerboseKey
"tc.mod.apply" VerboseLevel
10 (TCM Doc -> TCM ()) -> TCM Doc -> TCM ()
forall a b. (a -> b) -> a -> b
$ [TCM Doc] -> TCM Doc
forall (m :: * -> *). Monad m => [m Doc] -> m Doc
vcat
[ TCM Doc
"applySection"
, TCM Doc
"new =" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> ModuleName -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty ModuleName
new
, TCM Doc
"ptel =" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Telescope -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty Telescope
ptel
, TCM Doc
"old =" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> ModuleName -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty ModuleName
old
, TCM Doc
"ts =" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> [Arg Term] -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty [Arg Term]
ts
]
((QName, QName) -> TCM ()) -> Ren QName -> TCM ()
forall (t :: * -> *) (m :: * -> *) a b.
(Foldable t, Monad m) =>
(a -> m b) -> t a -> m ()
mapM_ ([Arg Term] -> (QName, QName) -> TCM ()
copyDef [Arg Term]
ts) Ren QName
rd
((ModuleName, ModuleName) -> TCM ()) -> Ren ModuleName -> TCM ()
forall (t :: * -> *) (m :: * -> *) a b.
(Foldable t, Monad m) =>
(a -> m b) -> t a -> m ()
mapM_ ([Arg Term] -> (ModuleName, ModuleName) -> TCM ()
copySec [Arg Term]
ts) Ren ModuleName
rm
[QName] -> TCM ()
computePolarity (((QName, QName) -> QName) -> Ren QName -> [QName]
forall a b. (a -> b) -> [a] -> [b]
map (QName, QName) -> QName
forall a b. (a, b) -> b
snd Ren QName
rd)
where
copyName :: QName -> QName
copyName QName
x = QName -> Maybe QName -> QName
forall a. a -> Maybe a -> a
fromMaybe QName
x (Maybe QName -> QName) -> Maybe QName -> QName
forall a b. (a -> b) -> a -> b
$ QName -> Ren QName -> Maybe QName
forall a b. Eq a => a -> [(a, b)] -> Maybe b
lookup QName
x Ren QName
rd
argsToUse :: ModuleName -> TCMT IO VerboseLevel
argsToUse ModuleName
x = do
let m :: ModuleName
m = ModuleName -> ModuleName -> ModuleName
commonParentModule ModuleName
old ModuleName
x
VerboseKey -> VerboseLevel -> TCM Doc -> TCM ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> VerboseLevel -> TCM Doc -> m ()
reportSDoc VerboseKey
"tc.mod.apply" VerboseLevel
80 (TCM Doc -> TCM ()) -> TCM Doc -> TCM ()
forall a b. (a -> b) -> a -> b
$ TCM Doc
"Common prefix: " TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> ModuleName -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty ModuleName
m
Telescope -> VerboseLevel
forall a. Sized a => a -> VerboseLevel
size (Telescope -> VerboseLevel)
-> TCMT IO Telescope -> TCMT IO VerboseLevel
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> ModuleName -> TCMT IO Telescope
forall (m :: * -> *).
(Functor m, ReadTCState m) =>
ModuleName -> m Telescope
lookupSection ModuleName
m
copyDef :: Args -> (QName, QName) -> TCM ()
copyDef :: [Arg Term] -> (QName, QName) -> TCM ()
copyDef [Arg Term]
ts (QName
x, QName
y) = do
Definition
def <- QName -> TCMT IO Definition
forall (m :: * -> *). HasConstInfo m => QName -> m Definition
getConstInfo QName
x
VerboseLevel
np <- ModuleName -> TCMT IO VerboseLevel
argsToUse (QName -> ModuleName
qnameModule QName
x)
[Hiding]
hidings <- (Dom (VerboseKey, Type) -> Hiding)
-> [Dom (VerboseKey, Type)] -> [Hiding]
forall a b. (a -> b) -> [a] -> [b]
map Dom (VerboseKey, Type) -> Hiding
forall a. LensHiding a => a -> Hiding
getHiding ([Dom (VerboseKey, Type)] -> [Hiding])
-> (Telescope -> [Dom (VerboseKey, Type)]) -> Telescope -> [Hiding]
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Telescope -> [Dom (VerboseKey, Type)]
forall t. Tele (Dom t) -> [Dom (VerboseKey, t)]
telToList (Telescope -> [Hiding]) -> TCMT IO Telescope -> TCMT IO [Hiding]
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> ModuleName -> TCMT IO Telescope
forall (m :: * -> *).
(Functor m, ReadTCState m) =>
ModuleName -> m Telescope
lookupSection (QName -> ModuleName
qnameModule QName
x)
let ts' :: [Arg Term]
ts' = (Hiding -> Arg Term -> Arg Term)
-> [Hiding] -> [Arg Term] -> [Arg Term]
forall a b c. (a -> b -> c) -> [a] -> [b] -> [c]
zipWith Hiding -> Arg Term -> Arg Term
forall a. LensHiding a => Hiding -> a -> a
setHiding [Hiding]
hidings [Arg Term]
ts
Telescope
commonTel <- ModuleName -> TCMT IO Telescope
forall (m :: * -> *).
(Functor m, ReadTCState m) =>
ModuleName -> m Telescope
lookupSection (ModuleName -> ModuleName -> ModuleName
commonParentModule ModuleName
old (ModuleName -> ModuleName) -> ModuleName -> ModuleName
forall a b. (a -> b) -> a -> b
$ QName -> ModuleName
qnameModule QName
x)
VerboseKey -> VerboseLevel -> TCM Doc -> TCM ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> VerboseLevel -> TCM Doc -> m ()
reportSDoc VerboseKey
"tc.mod.apply" VerboseLevel
80 (TCM Doc -> TCM ()) -> TCM Doc -> TCM ()
forall a b. (a -> b) -> a -> b
$ [TCM Doc] -> TCM Doc
forall (m :: * -> *). Monad m => [m Doc] -> m Doc
vcat
[ TCM Doc
"copyDef" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> QName -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty QName
x TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> TCM Doc
"->" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> QName -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty QName
y
, TCM Doc
"ts' = " TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> [Arg Term] -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty [Arg Term]
ts' ]
[Arg Term] -> VerboseLevel -> Definition -> TCM ()
copyDef' [Arg Term]
ts' VerboseLevel
np Definition
def
where
copyDef' :: [Arg Term] -> VerboseLevel -> Definition -> TCM ()
copyDef' [Arg Term]
ts VerboseLevel
np Definition
d = do
VerboseKey -> VerboseLevel -> TCM Doc -> TCM ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> VerboseLevel -> TCM Doc -> m ()
reportSDoc VerboseKey
"tc.mod.apply" VerboseLevel
60 (TCM Doc -> TCM ()) -> TCM Doc -> TCM ()
forall a b. (a -> b) -> a -> b
$ TCM Doc
"making new def for" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> QName -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty QName
y TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> TCM Doc
"from" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> QName -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty QName
x TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> TCM Doc
"with" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> VerboseKey -> TCM Doc
forall (m :: * -> *). Monad m => VerboseKey -> m Doc
text (VerboseLevel -> VerboseKey
forall a. Show a => a -> VerboseKey
show VerboseLevel
np) TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> TCM Doc
"args" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> VerboseKey -> TCM Doc
forall (m :: * -> *). Monad m => VerboseKey -> m Doc
text (IsAbstract -> VerboseKey
forall a. Show a => a -> VerboseKey
show (IsAbstract -> VerboseKey) -> IsAbstract -> VerboseKey
forall a b. (a -> b) -> a -> b
$ Definition -> IsAbstract
defAbstract Definition
d)
VerboseKey -> VerboseLevel -> TCM Doc -> TCM ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> VerboseLevel -> TCM Doc -> m ()
reportSDoc VerboseKey
"tc.mod.apply" VerboseLevel
80 (TCM Doc -> TCM ()) -> TCM Doc -> TCM ()
forall a b. (a -> b) -> a -> b
$ [TCM Doc] -> TCM Doc
forall (m :: * -> *). Monad m => [m Doc] -> m Doc
vcat
[ TCM Doc
"args = " TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> VerboseKey -> TCM Doc
forall (m :: * -> *). Monad m => VerboseKey -> m Doc
text ([Arg Term] -> VerboseKey
forall a. Show a => a -> VerboseKey
show [Arg Term]
ts')
, TCM Doc
"old type = " TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Type -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty (Definition -> Type
defType Definition
d) ]
VerboseKey -> VerboseLevel -> TCM Doc -> TCM ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> VerboseLevel -> TCM Doc -> m ()
reportSDoc VerboseKey
"tc.mod.apply" VerboseLevel
80 (TCM Doc -> TCM ()) -> TCM Doc -> TCM ()
forall a b. (a -> b) -> a -> b
$
TCM Doc
"new type = " TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Type -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty Type
t
QName -> Definition -> TCM ()
addConstant QName
y (Definition -> TCM ()) -> TCMT IO Definition -> TCM ()
forall (m :: * -> *) a b. Monad m => (a -> m b) -> m a -> m b
=<< QName -> TCMT IO Definition
nd QName
y
QName -> TCM ()
makeProjection QName
y
Bool -> TCM () -> TCM ()
forall (f :: * -> *). Applicative f => Bool -> f () -> f ()
unless Bool
isCon (TCM () -> TCM ()) -> TCM () -> TCM ()
forall a b. (a -> b) -> a -> b
$ Maybe QName -> (QName -> TCM ()) -> TCM ()
forall (m :: * -> *) a. Monad m => Maybe a -> (a -> m ()) -> m ()
whenJust Maybe QName
inst ((QName -> TCM ()) -> TCM ()) -> (QName -> TCM ()) -> TCM ()
forall a b. (a -> b) -> a -> b
$ \ QName
c -> QName -> QName -> TCM ()
addNamedInstance QName
y QName
c
Bool -> TCM () -> TCM ()
forall (f :: * -> *). Applicative f => Bool -> f () -> f ()
when (Telescope -> VerboseLevel
forall a. Sized a => a -> VerboseLevel
size Telescope
ptel VerboseLevel -> VerboseLevel -> Bool
forall a. Eq a => a -> a -> Bool
== VerboseLevel
0) (TCM () -> TCM ()) -> TCM () -> TCM ()
forall a b. (a -> b) -> a -> b
$ do
QName -> TCM ()
addDisplayForms QName
y
where
ts' :: [Arg Term]
ts' = VerboseLevel -> [Arg Term] -> [Arg Term]
forall a. VerboseLevel -> [a] -> [a]
take VerboseLevel
np [Arg Term]
ts
t :: Type
t = Definition -> Type
defType Definition
d Type -> [Arg Term] -> Type
`piApply` [Arg Term]
ts'
pol :: [Polarity]
pol = Definition -> [Polarity]
defPolarity Definition
d [Polarity] -> [Arg Term] -> [Polarity]
forall t. Apply t => t -> [Arg Term] -> t
`apply` [Arg Term]
ts'
occ :: [Occurrence]
occ = Definition -> [Occurrence]
defArgOccurrences Definition
d [Occurrence] -> [Arg Term] -> [Occurrence]
forall t. Apply t => t -> [Arg Term] -> t
`apply` [Arg Term]
ts'
gen :: NumGeneralizableArgs
gen = Definition -> NumGeneralizableArgs
defArgGeneralizable Definition
d NumGeneralizableArgs -> [Arg Term] -> NumGeneralizableArgs
forall t. Apply t => t -> [Arg Term] -> t
`apply` [Arg Term]
ts'
inst :: Maybe QName
inst = Definition -> Maybe QName
defInstance Definition
d
nd :: QName -> TCM Definition
nd :: QName -> TCMT IO Definition
nd QName
y = TCMT IO Defn -> (Defn -> Definition) -> TCMT IO Definition
forall (m :: * -> *) a b. Functor m => m a -> (a -> b) -> m b
for TCMT IO Defn
def ((Defn -> Definition) -> TCMT IO Definition)
-> (Defn -> Definition) -> TCMT IO Definition
forall a b. (a -> b) -> a -> b
$ \ Defn
df -> Defn :: ArgInfo
-> QName
-> Type
-> [Polarity]
-> [Occurrence]
-> NumGeneralizableArgs
-> [Maybe Name]
-> [LocalDisplayForm]
-> MutualId
-> CompiledRepresentation
-> Maybe QName
-> Bool
-> Set QName
-> Bool
-> Bool
-> Bool
-> Blocked_
-> Defn
-> Definition
Defn
{ defArgInfo :: ArgInfo
defArgInfo = Definition -> ArgInfo
defArgInfo Definition
d
, defName :: QName
defName = QName
y
, defType :: Type
defType = Type
t
, defPolarity :: [Polarity]
defPolarity = [Polarity]
pol
, defArgOccurrences :: [Occurrence]
defArgOccurrences = [Occurrence]
occ
, defArgGeneralizable :: NumGeneralizableArgs
defArgGeneralizable = NumGeneralizableArgs
gen
, defGeneralizedParams :: [Maybe Name]
defGeneralizedParams = []
, defDisplay :: [LocalDisplayForm]
defDisplay = []
, defMutual :: MutualId
defMutual = -MutualId
1
, defCompiledRep :: CompiledRepresentation
defCompiledRep = CompiledRepresentation
noCompiledRep
, defInstance :: Maybe QName
defInstance = Maybe QName
inst
, defCopy :: Bool
defCopy = Bool
True
, defMatchable :: Set QName
defMatchable = Set QName
forall a. Set a
Set.empty
, defNoCompilation :: Bool
defNoCompilation = Definition -> Bool
defNoCompilation Definition
d
, defInjective :: Bool
defInjective = Bool
False
, defCopatternLHS :: Bool
defCopatternLHS = [Clause] -> Bool
isCopatternLHS [Clause
cl]
, defBlocked :: Blocked_
defBlocked = Definition -> Blocked_
defBlocked Definition
d
, theDef :: Defn
theDef = Defn
df }
oldDef :: Defn
oldDef = Definition -> Defn
theDef Definition
d
isCon :: Bool
isCon = case Defn
oldDef of { Constructor{} -> Bool
True ; Defn
_ -> Bool
False }
mutual :: Maybe [QName]
mutual = case Defn
oldDef of { Function{funMutual :: Defn -> Maybe [QName]
funMutual = Maybe [QName]
m} -> Maybe [QName]
m ; Defn
_ -> Maybe [QName]
forall a. Maybe a
Nothing }
extlam :: Maybe ExtLamInfo
extlam = case Defn
oldDef of { Function{funExtLam :: Defn -> Maybe ExtLamInfo
funExtLam = Maybe ExtLamInfo
e} -> Maybe ExtLamInfo
e ; Defn
_ -> Maybe ExtLamInfo
forall a. Maybe a
Nothing }
with :: Maybe QName
with = case Defn
oldDef of { Function{funWith :: Defn -> Maybe QName
funWith = Maybe QName
w} -> QName -> QName
copyName (QName -> QName) -> Maybe QName -> Maybe QName
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> Maybe QName
w ; Defn
_ -> Maybe QName
forall a. Maybe a
Nothing }
isVar0 :: Arg Term -> Bool
isVar0 Arg Term
t = case Arg Term -> Term
forall e. Arg e -> e
unArg Arg Term
t of Var VerboseLevel
0 [] -> Bool
True; Term
_ -> Bool
False
proj :: Maybe Projection
proj = case Defn
oldDef of
Function{funProjection :: Defn -> Maybe Projection
funProjection = Just p :: Projection
p@Projection{projIndex :: Projection -> VerboseLevel
projIndex = VerboseLevel
n}}
| [Arg Term] -> VerboseLevel
forall a. Sized a => a -> VerboseLevel
size [Arg Term]
ts' VerboseLevel -> VerboseLevel -> Bool
forall a. Ord a => a -> a -> Bool
< VerboseLevel
n Bool -> Bool -> Bool
|| ([Arg Term] -> VerboseLevel
forall a. Sized a => a -> VerboseLevel
size [Arg Term]
ts' VerboseLevel -> VerboseLevel -> Bool
forall a. Eq a => a -> a -> Bool
== VerboseLevel
n Bool -> Bool -> Bool
&& Bool -> (Arg Term -> Bool) -> Maybe (Arg Term) -> Bool
forall b a. b -> (a -> b) -> Maybe a -> b
maybe Bool
True Arg Term -> Bool
isVar0 ([Arg Term] -> Maybe (Arg Term)
forall a. [a] -> Maybe a
lastMaybe [Arg Term]
ts'))
-> Projection -> Maybe Projection
forall a. a -> Maybe a
Just (Projection -> Maybe Projection) -> Projection -> Maybe Projection
forall a b. (a -> b) -> a -> b
$ Projection
p { projIndex :: VerboseLevel
projIndex = VerboseLevel
n VerboseLevel -> VerboseLevel -> VerboseLevel
forall a. Num a => a -> a -> a
- [Arg Term] -> VerboseLevel
forall a. Sized a => a -> VerboseLevel
size [Arg Term]
ts'
, projLams :: ProjLams
projLams = Projection -> ProjLams
projLams Projection
p ProjLams -> [Arg Term] -> ProjLams
forall t. Apply t => t -> [Arg Term] -> t
`apply` [Arg Term]
ts'
, projProper :: Maybe QName
projProper= (QName -> QName) -> Maybe QName -> Maybe QName
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
fmap QName -> QName
copyName (Maybe QName -> Maybe QName) -> Maybe QName -> Maybe QName
forall a b. (a -> b) -> a -> b
$ Projection -> Maybe QName
projProper Projection
p
}
Defn
_ -> Maybe Projection
forall a. Maybe a
Nothing
def :: TCMT IO Defn
def =
case Defn
oldDef of
Constructor{ conPars :: Defn -> VerboseLevel
conPars = VerboseLevel
np, conData :: Defn -> QName
conData = QName
d } -> Defn -> TCMT IO Defn
forall (m :: * -> *) a. Monad m => a -> m a
return (Defn -> TCMT IO Defn) -> Defn -> TCMT IO Defn
forall a b. (a -> b) -> a -> b
$
Defn
oldDef { conPars :: VerboseLevel
conPars = VerboseLevel
np VerboseLevel -> VerboseLevel -> VerboseLevel
forall a. Num a => a -> a -> a
- [Arg Term] -> VerboseLevel
forall a. Sized a => a -> VerboseLevel
size [Arg Term]
ts'
, conData :: QName
conData = QName -> QName
copyName QName
d
}
Datatype{ dataPars :: Defn -> VerboseLevel
dataPars = VerboseLevel
np, dataCons :: Defn -> [QName]
dataCons = [QName]
cs } -> Defn -> TCMT IO Defn
forall (m :: * -> *) a. Monad m => a -> m a
return (Defn -> TCMT IO Defn) -> Defn -> TCMT IO Defn
forall a b. (a -> b) -> a -> b
$
Defn
oldDef { dataPars :: VerboseLevel
dataPars = VerboseLevel
np VerboseLevel -> VerboseLevel -> VerboseLevel
forall a. Num a => a -> a -> a
- [Arg Term] -> VerboseLevel
forall a. Sized a => a -> VerboseLevel
size [Arg Term]
ts'
, dataClause :: Maybe Clause
dataClause = Clause -> Maybe Clause
forall a. a -> Maybe a
Just Clause
cl
, dataCons :: [QName]
dataCons = (QName -> QName) -> [QName] -> [QName]
forall a b. (a -> b) -> [a] -> [b]
map QName -> QName
copyName [QName]
cs
}
Record{ recPars :: Defn -> VerboseLevel
recPars = VerboseLevel
np, recTel :: Defn -> Telescope
recTel = Telescope
tel } -> Defn -> TCMT IO Defn
forall (m :: * -> *) a. Monad m => a -> m a
return (Defn -> TCMT IO Defn) -> Defn -> TCMT IO Defn
forall a b. (a -> b) -> a -> b
$
Defn
oldDef { recPars :: VerboseLevel
recPars = VerboseLevel
np VerboseLevel -> VerboseLevel -> VerboseLevel
forall a. Num a => a -> a -> a
- [Arg Term] -> VerboseLevel
forall a. Sized a => a -> VerboseLevel
size [Arg Term]
ts'
, recClause :: Maybe Clause
recClause = Clause -> Maybe Clause
forall a. a -> Maybe a
Just Clause
cl
, recTel :: Telescope
recTel = Telescope -> [Arg Term] -> Telescope
forall t. Apply t => t -> [Arg Term] -> t
apply Telescope
tel [Arg Term]
ts'
}
Defn
GeneralizableVar -> Defn -> TCMT IO Defn
forall (m :: * -> *) a. Monad m => a -> m a
return Defn
GeneralizableVar
Defn
_ -> do
(Maybe SplitTree
mst, Bool
_, CompiledClauses
cc) <- Maybe (QName, Type)
-> [Clause] -> TCM (Maybe SplitTree, Bool, CompiledClauses)
compileClauses Maybe (QName, Type)
forall a. Maybe a
Nothing [Clause
cl]
let newDef :: Defn
newDef =
Lens' Bool Defn -> LensSet Bool Defn
forall i o. Lens' i o -> LensSet i o
set Lens' Bool Defn
funMacro (Defn
oldDef Defn -> Lens' Bool Defn -> Bool
forall o i. o -> Lens' i o -> i
^. Lens' Bool Defn
funMacro) (Defn -> Defn) -> Defn -> Defn
forall a b. (a -> b) -> a -> b
$
Lens' Bool Defn -> LensSet Bool Defn
forall i o. Lens' i o -> LensSet i o
set Lens' Bool Defn
funStatic (Defn
oldDef Defn -> Lens' Bool Defn -> Bool
forall o i. o -> Lens' i o -> i
^. Lens' Bool Defn
funStatic) (Defn -> Defn) -> Defn -> Defn
forall a b. (a -> b) -> a -> b
$
Lens' Bool Defn -> LensSet Bool Defn
forall i o. Lens' i o -> LensSet i o
set Lens' Bool Defn
funInline Bool
True (Defn -> Defn) -> Defn -> Defn
forall a b. (a -> b) -> a -> b
$
Defn
emptyFunction
{ funClauses :: [Clause]
funClauses = [Clause
cl]
, funCompiled :: Maybe CompiledClauses
funCompiled = CompiledClauses -> Maybe CompiledClauses
forall a. a -> Maybe a
Just CompiledClauses
cc
, funSplitTree :: Maybe SplitTree
funSplitTree = Maybe SplitTree
mst
, funMutual :: Maybe [QName]
funMutual = Maybe [QName]
mutual
, funProjection :: Maybe Projection
funProjection = Maybe Projection
proj
, funTerminates :: Maybe Bool
funTerminates = Bool -> Maybe Bool
forall a. a -> Maybe a
Just Bool
True
, funExtLam :: Maybe ExtLamInfo
funExtLam = Maybe ExtLamInfo
extlam
, funWith :: Maybe QName
funWith = Maybe QName
with
}
VerboseKey -> VerboseLevel -> TCM Doc -> TCM ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> VerboseLevel -> TCM Doc -> m ()
reportSDoc VerboseKey
"tc.mod.apply" VerboseLevel
80 (TCM Doc -> TCM ()) -> TCM Doc -> TCM ()
forall a b. (a -> b) -> a -> b
$ (TCM Doc
"new def for" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> QName -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty QName
x) TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<?> Defn -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty Defn
newDef
Defn -> TCMT IO Defn
forall (m :: * -> *) a. Monad m => a -> m a
return Defn
newDef
cl :: Clause
cl = Clause :: Range
-> Range
-> Telescope
-> [NamedArg (Pattern' DBPatVar)]
-> Maybe Term
-> Maybe (Arg Type)
-> Bool
-> Maybe Bool
-> Maybe Bool
-> ExpandedEllipsis
-> Clause
Clause { clauseLHSRange :: Range
clauseLHSRange = [Clause] -> Range
forall t. HasRange t => t -> Range
getRange ([Clause] -> Range) -> [Clause] -> Range
forall a b. (a -> b) -> a -> b
$ Definition -> [Clause]
defClauses Definition
d
, clauseFullRange :: Range
clauseFullRange = [Clause] -> Range
forall t. HasRange t => t -> Range
getRange ([Clause] -> Range) -> [Clause] -> Range
forall a b. (a -> b) -> a -> b
$ Definition -> [Clause]
defClauses Definition
d
, clauseTel :: Telescope
clauseTel = Telescope
forall a. Tele a
EmptyTel
, namedClausePats :: [NamedArg (Pattern' DBPatVar)]
namedClausePats = []
, clauseBody :: Maybe Term
clauseBody = Term -> Maybe Term
forall a. a -> Maybe a
Just (Term -> Maybe Term) -> Term -> Maybe Term
forall a b. (a -> b) -> a -> b
$ VerboseLevel -> Term -> Term
forall a. DropArgs a => VerboseLevel -> a -> a
dropArgs VerboseLevel
pars (Term -> Term) -> Term -> Term
forall a b. (a -> b) -> a -> b
$ case Defn
oldDef of
Function{funProjection :: Defn -> Maybe Projection
funProjection = Just Projection
p} -> Projection -> ProjOrigin -> Relevance -> [Arg Term] -> Term
projDropParsApply Projection
p ProjOrigin
ProjSystem Relevance
rel [Arg Term]
ts'
Defn
_ -> QName -> Elims -> Term
Def QName
x (Elims -> Term) -> Elims -> Term
forall a b. (a -> b) -> a -> b
$ (Arg Term -> Elim' Term) -> [Arg Term] -> Elims
forall a b. (a -> b) -> [a] -> [b]
map Arg Term -> Elim' Term
forall a. Arg a -> Elim' a
Apply [Arg Term]
ts'
, clauseType :: Maybe (Arg Type)
clauseType = Arg Type -> Maybe (Arg Type)
forall a. a -> Maybe a
Just (Arg Type -> Maybe (Arg Type)) -> Arg Type -> Maybe (Arg Type)
forall a b. (a -> b) -> a -> b
$ Type -> Arg Type
forall a. a -> Arg a
defaultArg Type
t
, clauseCatchall :: Bool
clauseCatchall = Bool
False
, clauseRecursive :: Maybe Bool
clauseRecursive = Bool -> Maybe Bool
forall a. a -> Maybe a
Just Bool
False
, clauseUnreachable :: Maybe Bool
clauseUnreachable = Bool -> Maybe Bool
forall a. a -> Maybe a
Just Bool
False
, clauseEllipsis :: ExpandedEllipsis
clauseEllipsis = ExpandedEllipsis
NoEllipsis
}
where
pars :: VerboseLevel
pars = VerboseLevel -> VerboseLevel -> VerboseLevel
forall a. Ord a => a -> a -> a
max VerboseLevel
0 (VerboseLevel -> VerboseLevel) -> VerboseLevel -> VerboseLevel
forall a b. (a -> b) -> a -> b
$ VerboseLevel
-> (Projection -> VerboseLevel) -> Maybe Projection -> VerboseLevel
forall b a. b -> (a -> b) -> Maybe a -> b
maybe VerboseLevel
0 (VerboseLevel -> VerboseLevel
forall a. Enum a => a -> a
pred (VerboseLevel -> VerboseLevel)
-> (Projection -> VerboseLevel) -> Projection -> VerboseLevel
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Projection -> VerboseLevel
projIndex) Maybe Projection
proj
rel :: Relevance
rel = ArgInfo -> Relevance
forall a. LensRelevance a => a -> Relevance
getRelevance (ArgInfo -> Relevance) -> ArgInfo -> Relevance
forall a b. (a -> b) -> a -> b
$ Definition -> ArgInfo
defArgInfo Definition
d
copySec :: Args -> (ModuleName, ModuleName) -> TCM ()
copySec :: [Arg Term] -> (ModuleName, ModuleName) -> TCM ()
copySec [Arg Term]
ts (ModuleName
x, ModuleName
y) = do
VerboseLevel
totalArgs <- ModuleName -> TCMT IO VerboseLevel
argsToUse ModuleName
x
Telescope
tel <- ModuleName -> TCMT IO Telescope
forall (m :: * -> *).
(Functor m, ReadTCState m) =>
ModuleName -> m Telescope
lookupSection ModuleName
x
let sectionTel :: Telescope
sectionTel = Telescope -> [Arg Term] -> Telescope
forall t. Apply t => t -> [Arg Term] -> t
apply Telescope
tel ([Arg Term] -> Telescope) -> [Arg Term] -> Telescope
forall a b. (a -> b) -> a -> b
$ VerboseLevel -> [Arg Term] -> [Arg Term]
forall a. VerboseLevel -> [a] -> [a]
take VerboseLevel
totalArgs [Arg Term]
ts
VerboseKey -> VerboseLevel -> TCM Doc -> TCM ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> VerboseLevel -> TCM Doc -> m ()
reportSDoc VerboseKey
"tc.mod.apply" VerboseLevel
80 (TCM Doc -> TCM ()) -> TCM Doc -> TCM ()
forall a b. (a -> b) -> a -> b
$ TCM Doc
"Copying section" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> ModuleName -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty ModuleName
x TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> TCM Doc
"to" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> ModuleName -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty ModuleName
y
VerboseKey -> VerboseLevel -> TCM Doc -> TCM ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> VerboseLevel -> TCM Doc -> m ()
reportSDoc VerboseKey
"tc.mod.apply" VerboseLevel
80 (TCM Doc -> TCM ()) -> TCM Doc -> TCM ()
forall a b. (a -> b) -> a -> b
$ TCM Doc
" ts = " TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> [TCM Doc] -> TCM Doc
forall a. Monoid a => [a] -> a
mconcat (TCM Doc -> [TCM Doc] -> [TCM Doc]
forall a. a -> [a] -> [a]
List.intersperse TCM Doc
"; " ((Arg Term -> TCM Doc) -> [Arg Term] -> [TCM Doc]
forall a b. (a -> b) -> [a] -> [b]
map Arg Term -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty [Arg Term]
ts))
VerboseKey -> VerboseLevel -> TCM Doc -> TCM ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> VerboseLevel -> TCM Doc -> m ()
reportSDoc VerboseKey
"tc.mod.apply" VerboseLevel
80 (TCM Doc -> TCM ()) -> TCM Doc -> TCM ()
forall a b. (a -> b) -> a -> b
$ TCM Doc
" totalArgs = " TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> VerboseKey -> TCM Doc
forall (m :: * -> *). Monad m => VerboseKey -> m Doc
text (VerboseLevel -> VerboseKey
forall a. Show a => a -> VerboseKey
show VerboseLevel
totalArgs)
VerboseKey -> VerboseLevel -> TCM Doc -> TCM ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> VerboseLevel -> TCM Doc -> m ()
reportSDoc VerboseKey
"tc.mod.apply" VerboseLevel
80 (TCM Doc -> TCM ()) -> TCM Doc -> TCM ()
forall a b. (a -> b) -> a -> b
$ TCM Doc
" tel = " TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> VerboseKey -> TCM Doc
forall (m :: * -> *). Monad m => VerboseKey -> m Doc
text (VerboseKey -> [VerboseKey] -> VerboseKey
forall a. [a] -> [[a]] -> [a]
List.intercalate VerboseKey
" " ((Dom (VerboseKey, Type) -> VerboseKey)
-> [Dom (VerboseKey, Type)] -> [VerboseKey]
forall a b. (a -> b) -> [a] -> [b]
map ((VerboseKey, Type) -> VerboseKey
forall a b. (a, b) -> a
fst ((VerboseKey, Type) -> VerboseKey)
-> (Dom (VerboseKey, Type) -> (VerboseKey, Type))
-> Dom (VerboseKey, Type)
-> VerboseKey
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Dom (VerboseKey, Type) -> (VerboseKey, Type)
forall t e. Dom' t e -> e
unDom) ([Dom (VerboseKey, Type)] -> [VerboseKey])
-> [Dom (VerboseKey, Type)] -> [VerboseKey]
forall a b. (a -> b) -> a -> b
$ Telescope -> [Dom (VerboseKey, Type)]
forall t. Tele (Dom t) -> [Dom (VerboseKey, t)]
telToList Telescope
tel))
VerboseKey -> VerboseLevel -> TCM Doc -> TCM ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> VerboseLevel -> TCM Doc -> m ()
reportSDoc VerboseKey
"tc.mod.apply" VerboseLevel
80 (TCM Doc -> TCM ()) -> TCM Doc -> TCM ()
forall a b. (a -> b) -> a -> b
$ TCM Doc
" sectionTel = " TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> VerboseKey -> TCM Doc
forall (m :: * -> *). Monad m => VerboseKey -> m Doc
text (VerboseKey -> [VerboseKey] -> VerboseKey
forall a. [a] -> [[a]] -> [a]
List.intercalate VerboseKey
" " ((Dom (VerboseKey, Type) -> VerboseKey)
-> [Dom (VerboseKey, Type)] -> [VerboseKey]
forall a b. (a -> b) -> [a] -> [b]
map ((VerboseKey, Type) -> VerboseKey
forall a b. (a, b) -> a
fst ((VerboseKey, Type) -> VerboseKey)
-> (Dom (VerboseKey, Type) -> (VerboseKey, Type))
-> Dom (VerboseKey, Type)
-> VerboseKey
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Dom (VerboseKey, Type) -> (VerboseKey, Type)
forall t e. Dom' t e -> e
unDom) ([Dom (VerboseKey, Type)] -> [VerboseKey])
-> [Dom (VerboseKey, Type)] -> [VerboseKey]
forall a b. (a -> b) -> a -> b
$ Telescope -> [Dom (VerboseKey, Type)]
forall t. Tele (Dom t) -> [Dom (VerboseKey, t)]
telToList Telescope
ptel))
Telescope -> TCM () -> TCM ()
forall b (m :: * -> *) a.
(AddContext b, MonadAddContext m) =>
b -> m a -> m a
addContext Telescope
sectionTel (TCM () -> TCM ()) -> TCM () -> TCM ()
forall a b. (a -> b) -> a -> b
$ ModuleName -> TCM ()
addSection ModuleName
y
addDisplayForm :: QName -> DisplayForm -> TCM ()
addDisplayForm :: QName -> DisplayForm -> TCM ()
addDisplayForm QName
x DisplayForm
df = do
LocalDisplayForm
d <- DisplayForm -> TCMT IO LocalDisplayForm
forall (m :: * -> *) a. MonadTCEnv m => a -> m (Open a)
makeOpen DisplayForm
df
let add :: Signature -> Signature
add = QName -> (Definition -> Definition) -> Signature -> Signature
updateDefinition QName
x ((Definition -> Definition) -> Signature -> Signature)
-> (Definition -> Definition) -> Signature -> Signature
forall a b. (a -> b) -> a -> b
$ \ Definition
def -> Definition
def{ defDisplay :: [LocalDisplayForm]
defDisplay = LocalDisplayForm
d LocalDisplayForm -> [LocalDisplayForm] -> [LocalDisplayForm]
forall a. a -> [a] -> [a]
: Definition -> [LocalDisplayForm]
defDisplay Definition
def }
TCMT IO Bool -> TCM () -> TCM () -> TCM ()
forall (m :: * -> *) a. Monad m => m Bool -> m a -> m a -> m a
ifM (QName -> TCMT IO Bool
forall (m :: * -> *). ReadTCState m => QName -> m Bool
isLocal QName
x)
((Signature -> Signature) -> TCM ()
modifySignature Signature -> Signature
add)
(Lens' DisplayForms TCState
stImportsDisplayForms Lens' DisplayForms TCState
-> (DisplayForms -> DisplayForms) -> TCM ()
forall (m :: * -> *) a.
MonadTCState m =>
Lens' a TCState -> (a -> a) -> m ()
`modifyTCLens` ([LocalDisplayForm] -> [LocalDisplayForm] -> [LocalDisplayForm])
-> QName -> [LocalDisplayForm] -> DisplayForms -> DisplayForms
forall k v.
(Eq k, Hashable k) =>
(v -> v -> v) -> k -> v -> HashMap k v -> HashMap k v
HMap.insertWith [LocalDisplayForm] -> [LocalDisplayForm] -> [LocalDisplayForm]
forall a. [a] -> [a] -> [a]
(++) QName
x [LocalDisplayForm
d])
TCMT IO Bool -> TCM () -> TCM ()
forall (m :: * -> *). Monad m => m Bool -> m () -> m ()
whenM (QName -> TCMT IO Bool
hasLoopingDisplayForm QName
x) (TCM () -> TCM ()) -> TCM () -> TCM ()
forall a b. (a -> b) -> a -> b
$
TypeError -> TCM ()
forall (m :: * -> *) a.
(MonadTCEnv m, ReadTCState m, MonadError TCErr m) =>
TypeError -> m a
typeError (TypeError -> TCM ()) -> (Doc -> TypeError) -> Doc -> TCM ()
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Doc -> TypeError
GenericDocError (Doc -> TCM ()) -> TCM Doc -> TCM ()
forall (m :: * -> *) a b. Monad m => (a -> m b) -> m a -> m b
=<< do TCM Doc
"Cannot add recursive display form for" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> QName -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty QName
x
isLocal :: ReadTCState m => QName -> m Bool
isLocal :: QName -> m Bool
isLocal QName
x = QName -> Definitions -> Bool
forall k a. (Eq k, Hashable k) => k -> HashMap k a -> Bool
HMap.member QName
x (Definitions -> Bool) -> m Definitions -> m Bool
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> Lens' Definitions TCState -> m Definitions
forall (m :: * -> *) a. ReadTCState m => Lens' a TCState -> m a
useR ((Signature -> f Signature) -> TCState -> f TCState
Lens' Signature TCState
stSignature ((Signature -> f Signature) -> TCState -> f TCState)
-> ((Definitions -> f Definitions) -> Signature -> f Signature)
-> (Definitions -> f Definitions)
-> TCState
-> f TCState
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (Definitions -> f Definitions) -> Signature -> f Signature
Lens' Definitions Signature
sigDefinitions)
getDisplayForms :: (HasConstInfo m, ReadTCState m) => QName -> m [LocalDisplayForm]
getDisplayForms :: QName -> m [LocalDisplayForm]
getDisplayForms QName
q = do
[LocalDisplayForm]
ds <- (SigError -> [LocalDisplayForm])
-> (Definition -> [LocalDisplayForm])
-> Either SigError Definition
-> [LocalDisplayForm]
forall a c b. (a -> c) -> (b -> c) -> Either a b -> c
either ([LocalDisplayForm] -> SigError -> [LocalDisplayForm]
forall a b. a -> b -> a
const []) Definition -> [LocalDisplayForm]
defDisplay (Either SigError Definition -> [LocalDisplayForm])
-> m (Either SigError Definition) -> m [LocalDisplayForm]
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> QName -> m (Either SigError Definition)
forall (m :: * -> *).
HasConstInfo m =>
QName -> m (Either SigError Definition)
getConstInfo' QName
q
[LocalDisplayForm]
ds1 <- [LocalDisplayForm] -> QName -> DisplayForms -> [LocalDisplayForm]
forall k v. (Eq k, Hashable k) => v -> k -> HashMap k v -> v
HMap.lookupDefault [] QName
q (DisplayForms -> [LocalDisplayForm])
-> m DisplayForms -> m [LocalDisplayForm]
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> Lens' DisplayForms TCState -> m DisplayForms
forall (m :: * -> *) a. ReadTCState m => Lens' a TCState -> m a
useR Lens' DisplayForms TCState
stImportsDisplayForms
[LocalDisplayForm]
ds2 <- [LocalDisplayForm] -> QName -> DisplayForms -> [LocalDisplayForm]
forall k v. (Eq k, Hashable k) => v -> k -> HashMap k v -> v
HMap.lookupDefault [] QName
q (DisplayForms -> [LocalDisplayForm])
-> m DisplayForms -> m [LocalDisplayForm]
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> Lens' DisplayForms TCState -> m DisplayForms
forall (m :: * -> *) a. ReadTCState m => Lens' a TCState -> m a
useR Lens' DisplayForms TCState
stImportedDisplayForms
m Bool
-> m [LocalDisplayForm]
-> m [LocalDisplayForm]
-> m [LocalDisplayForm]
forall (m :: * -> *) a. Monad m => m Bool -> m a -> m a -> m a
ifM (QName -> m Bool
forall (m :: * -> *). ReadTCState m => QName -> m Bool
isLocal QName
q) ([LocalDisplayForm] -> m [LocalDisplayForm]
forall (m :: * -> *) a. Monad m => a -> m a
return ([LocalDisplayForm] -> m [LocalDisplayForm])
-> [LocalDisplayForm] -> m [LocalDisplayForm]
forall a b. (a -> b) -> a -> b
$ [LocalDisplayForm]
ds [LocalDisplayForm] -> [LocalDisplayForm] -> [LocalDisplayForm]
forall a. [a] -> [a] -> [a]
++ [LocalDisplayForm]
ds1 [LocalDisplayForm] -> [LocalDisplayForm] -> [LocalDisplayForm]
forall a. [a] -> [a] -> [a]
++ [LocalDisplayForm]
ds2)
([LocalDisplayForm] -> m [LocalDisplayForm]
forall (m :: * -> *) a. Monad m => a -> m a
return ([LocalDisplayForm] -> m [LocalDisplayForm])
-> [LocalDisplayForm] -> m [LocalDisplayForm]
forall a b. (a -> b) -> a -> b
$ [LocalDisplayForm]
ds1 [LocalDisplayForm] -> [LocalDisplayForm] -> [LocalDisplayForm]
forall a. [a] -> [a] -> [a]
++ [LocalDisplayForm]
ds [LocalDisplayForm] -> [LocalDisplayForm] -> [LocalDisplayForm]
forall a. [a] -> [a] -> [a]
++ [LocalDisplayForm]
ds2)
chaseDisplayForms :: QName -> TCM (Set QName)
chaseDisplayForms :: QName -> TCM (Set QName)
chaseDisplayForms QName
q = Set QName -> [QName] -> TCM (Set QName)
go Set QName
forall a. Set a
Set.empty [QName
q]
where
go :: Set QName -> [QName] -> TCM (Set QName)
go Set QName
used [] = Set QName -> TCM (Set QName)
forall (f :: * -> *) a. Applicative f => a -> f a
pure Set QName
used
go Set QName
used (QName
q : [QName]
qs) = do
let rhs :: DisplayForm -> DisplayTerm
rhs (Display VerboseLevel
_ Elims
_ DisplayTerm
e) = DisplayTerm
e
Set QName
ds <- (Set QName -> Set QName -> Set QName
forall a. Ord a => Set a -> Set a -> Set a
`Set.difference` Set QName
used) (Set QName -> Set QName)
-> ([LocalDisplayForm] -> Set QName)
-> [LocalDisplayForm]
-> Set QName
forall b c a. (b -> c) -> (a -> b) -> a -> c
. [Set QName] -> Set QName
forall (f :: * -> *) a. (Foldable f, Ord a) => f (Set a) -> Set a
Set.unions ([Set QName] -> Set QName)
-> ([LocalDisplayForm] -> [Set QName])
-> [LocalDisplayForm]
-> Set QName
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (LocalDisplayForm -> Set QName)
-> [LocalDisplayForm] -> [Set QName]
forall a b. (a -> b) -> [a] -> [b]
map (DisplayTerm -> Set QName
forall a. NamesIn a => a -> Set QName
namesIn (DisplayTerm -> Set QName)
-> (LocalDisplayForm -> DisplayTerm)
-> LocalDisplayForm
-> Set QName
forall b c a. (b -> c) -> (a -> b) -> a -> c
. DisplayForm -> DisplayTerm
rhs (DisplayForm -> DisplayTerm)
-> (LocalDisplayForm -> DisplayForm)
-> LocalDisplayForm
-> DisplayTerm
forall b c a. (b -> c) -> (a -> b) -> a -> c
. LocalDisplayForm -> DisplayForm
forall (t :: * -> *) a. Decoration t => t a -> a
dget)
([LocalDisplayForm] -> Set QName)
-> TCMT IO [LocalDisplayForm] -> TCM (Set QName)
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> (QName -> TCMT IO [LocalDisplayForm]
forall (m :: * -> *).
(HasConstInfo m, ReadTCState m) =>
QName -> m [LocalDisplayForm]
getDisplayForms QName
q TCMT IO [LocalDisplayForm]
-> (TCErr -> TCMT IO [LocalDisplayForm])
-> TCMT IO [LocalDisplayForm]
forall a. TCM a -> (TCErr -> TCM a) -> TCM a
`catchError_` \ TCErr
_ -> [LocalDisplayForm] -> TCMT IO [LocalDisplayForm]
forall (f :: * -> *) a. Applicative f => a -> f a
pure [])
Set QName -> [QName] -> TCM (Set QName)
go (Set QName -> Set QName -> Set QName
forall a. Ord a => Set a -> Set a -> Set a
Set.union Set QName
ds Set QName
used) (Set QName -> [QName]
forall a. Set a -> [a]
Set.toList Set QName
ds [QName] -> [QName] -> [QName]
forall a. [a] -> [a] -> [a]
++ [QName]
qs)
hasLoopingDisplayForm :: QName -> TCM Bool
hasLoopingDisplayForm :: QName -> TCMT IO Bool
hasLoopingDisplayForm QName
q = QName -> Set QName -> Bool
forall a. Ord a => a -> Set a -> Bool
Set.member QName
q (Set QName -> Bool) -> TCM (Set QName) -> TCMT IO Bool
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> QName -> TCM (Set QName)
chaseDisplayForms QName
q
canonicalName :: HasConstInfo m => QName -> m QName
canonicalName :: QName -> m QName
canonicalName QName
x = do
Defn
def <- Definition -> Defn
theDef (Definition -> Defn) -> m Definition -> m Defn
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> QName -> m Definition
forall (m :: * -> *). HasConstInfo m => QName -> m Definition
getConstInfo QName
x
case Defn
def of
Constructor{conSrcCon :: Defn -> ConHead
conSrcCon = ConHead
c} -> QName -> m QName
forall (m :: * -> *) a. Monad m => a -> m a
return (QName -> m QName) -> QName -> m QName
forall a b. (a -> b) -> a -> b
$ ConHead -> QName
conName ConHead
c
Record{recClause :: Defn -> Maybe Clause
recClause = Just (Clause{ clauseBody :: Clause -> Maybe Term
clauseBody = Maybe Term
body })} -> Maybe Term -> m QName
forall (m :: * -> *). HasConstInfo m => Maybe Term -> m QName
can Maybe Term
body
Datatype{dataClause :: Defn -> Maybe Clause
dataClause = Just (Clause{ clauseBody :: Clause -> Maybe Term
clauseBody = Maybe Term
body })} -> Maybe Term -> m QName
forall (m :: * -> *). HasConstInfo m => Maybe Term -> m QName
can Maybe Term
body
Defn
_ -> QName -> m QName
forall (m :: * -> *) a. Monad m => a -> m a
return QName
x
where
can :: Maybe Term -> m QName
can Maybe Term
body = QName -> m QName
forall (m :: * -> *). HasConstInfo m => QName -> m QName
canonicalName (QName -> m QName) -> QName -> m QName
forall a b. (a -> b) -> a -> b
$ Term -> QName
extract (Term -> QName) -> Term -> QName
forall a b. (a -> b) -> a -> b
$ Term -> Maybe Term -> Term
forall a. a -> Maybe a -> a
fromMaybe Term
forall a. HasCallStack => a
__IMPOSSIBLE__ Maybe Term
body
extract :: Term -> QName
extract (Def QName
x Elims
_) = QName
x
extract Term
_ = QName
forall a. HasCallStack => a
__IMPOSSIBLE__
sameDef :: HasConstInfo m => QName -> QName -> m (Maybe QName)
sameDef :: QName -> QName -> m (Maybe QName)
sameDef QName
d1 QName
d2 = do
QName
c1 <- QName -> m QName
forall (m :: * -> *). HasConstInfo m => QName -> m QName
canonicalName QName
d1
QName
c2 <- QName -> m QName
forall (m :: * -> *). HasConstInfo m => QName -> m QName
canonicalName QName
d2
if (QName
c1 QName -> QName -> Bool
forall a. Eq a => a -> a -> Bool
== QName
c2) then Maybe QName -> m (Maybe QName)
forall (m :: * -> *) a. Monad m => a -> m a
return (Maybe QName -> m (Maybe QName)) -> Maybe QName -> m (Maybe QName)
forall a b. (a -> b) -> a -> b
$ QName -> Maybe QName
forall a. a -> Maybe a
Just QName
c1 else Maybe QName -> m (Maybe QName)
forall (m :: * -> *) a. Monad m => a -> m a
return Maybe QName
forall a. Maybe a
Nothing
whatInduction :: MonadTCM tcm => QName -> tcm Induction
whatInduction :: QName -> tcm Induction
whatInduction QName
c = TCM Induction -> tcm Induction
forall (tcm :: * -> *) a. MonadTCM tcm => TCM a -> tcm a
liftTCM (TCM Induction -> tcm Induction) -> TCM Induction -> tcm Induction
forall a b. (a -> b) -> a -> b
$ do
Defn
def <- Definition -> Defn
theDef (Definition -> Defn) -> TCMT IO Definition -> TCMT IO Defn
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> QName -> TCMT IO Definition
forall (m :: * -> *). HasConstInfo m => QName -> m Definition
getConstInfo QName
c
Maybe QName
mz <- VerboseKey -> TCMT IO (Maybe QName)
forall (m :: * -> *).
HasBuiltins m =>
VerboseKey -> m (Maybe QName)
getBuiltinName' VerboseKey
builtinIZero
Maybe QName
mo <- VerboseKey -> TCMT IO (Maybe QName)
forall (m :: * -> *).
HasBuiltins m =>
VerboseKey -> m (Maybe QName)
getBuiltinName' VerboseKey
builtinIOne
case Defn
def of
Datatype{} -> Induction -> TCM Induction
forall (m :: * -> *) a. Monad m => a -> m a
return Induction
Inductive
Record{} | Bool -> Bool
not (Defn -> Bool
recRecursive Defn
def) -> Induction -> TCM Induction
forall (m :: * -> *) a. Monad m => a -> m a
return Induction
Inductive
Record{ recInduction :: Defn -> Maybe Induction
recInduction = Maybe Induction
i } -> Induction -> TCM Induction
forall (m :: * -> *) a. Monad m => a -> m a
return (Induction -> TCM Induction) -> Induction -> TCM Induction
forall a b. (a -> b) -> a -> b
$ Induction -> Maybe Induction -> Induction
forall a. a -> Maybe a -> a
fromMaybe Induction
Inductive Maybe Induction
i
Constructor{ conInd :: Defn -> Induction
conInd = Induction
i } -> Induction -> TCM Induction
forall (m :: * -> *) a. Monad m => a -> m a
return Induction
i
Defn
_ | QName -> Maybe QName
forall a. a -> Maybe a
Just QName
c Maybe QName -> Maybe QName -> Bool
forall a. Eq a => a -> a -> Bool
== Maybe QName
mz Bool -> Bool -> Bool
|| QName -> Maybe QName
forall a. a -> Maybe a
Just QName
c Maybe QName -> Maybe QName -> Bool
forall a. Eq a => a -> a -> Bool
== Maybe QName
mo
-> Induction -> TCM Induction
forall (m :: * -> *) a. Monad m => a -> m a
return Induction
Inductive
Defn
_ -> TCM Induction
forall a. HasCallStack => a
__IMPOSSIBLE__
singleConstructorType :: QName -> TCM Bool
singleConstructorType :: QName -> TCMT IO Bool
singleConstructorType QName
q = do
Defn
d <- Definition -> Defn
theDef (Definition -> Defn) -> TCMT IO Definition -> TCMT IO Defn
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> QName -> TCMT IO Definition
forall (m :: * -> *). HasConstInfo m => QName -> m Definition
getConstInfo QName
q
case Defn
d of
Record {} -> Bool -> TCMT IO Bool
forall (m :: * -> *) a. Monad m => a -> m a
return Bool
True
Constructor { conData :: Defn -> QName
conData = QName
d } -> do
Defn
di <- Definition -> Defn
theDef (Definition -> Defn) -> TCMT IO Definition -> TCMT IO Defn
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> QName -> TCMT IO Definition
forall (m :: * -> *). HasConstInfo m => QName -> m Definition
getConstInfo QName
d
Bool -> TCMT IO Bool
forall (m :: * -> *) a. Monad m => a -> m a
return (Bool -> TCMT IO Bool) -> Bool -> TCMT IO Bool
forall a b. (a -> b) -> a -> b
$ case Defn
di of
Record {} -> Bool
True
Datatype { dataCons :: Defn -> [QName]
dataCons = [QName]
cs } -> [QName] -> VerboseLevel
forall (t :: * -> *) a. Foldable t => t a -> VerboseLevel
length [QName]
cs VerboseLevel -> VerboseLevel -> Bool
forall a. Eq a => a -> a -> Bool
== VerboseLevel
1
Defn
_ -> Bool
forall a. HasCallStack => a
__IMPOSSIBLE__
Defn
_ -> TCMT IO Bool
forall a. HasCallStack => a
__IMPOSSIBLE__
data SigError
= SigUnknown String
| SigAbstract
sigError :: (String -> a) -> a -> SigError -> a
sigError :: (VerboseKey -> a) -> a -> SigError -> a
sigError VerboseKey -> a
f a
a = \case
SigUnknown VerboseKey
s -> VerboseKey -> a
f VerboseKey
s
SigError
SigAbstract -> a
a
class ( Functor m
, Applicative m
, Fail.MonadFail m
, HasOptions m
, MonadDebug m
, MonadTCEnv m
) => HasConstInfo m where
getConstInfo :: QName -> m Definition
getConstInfo QName
q = QName -> m (Either SigError Definition)
forall (m :: * -> *).
HasConstInfo m =>
QName -> m (Either SigError Definition)
getConstInfo' QName
q m (Either SigError Definition)
-> (Either SigError Definition -> m Definition) -> m Definition
forall (m :: * -> *) a b. Monad m => m a -> (a -> m b) -> m b
>>= \case
Right Definition
d -> Definition -> m Definition
forall (m :: * -> *) a. Monad m => a -> m a
return Definition
d
Left (SigUnknown VerboseKey
err) -> VerboseKey -> m Definition
forall (m :: * -> *) a.
(HasCallStack, MonadDebug m) =>
VerboseKey -> m a
__IMPOSSIBLE_VERBOSE__ VerboseKey
err
Left SigError
SigAbstract -> VerboseKey -> m Definition
forall (m :: * -> *) a.
(HasCallStack, MonadDebug m) =>
VerboseKey -> m a
__IMPOSSIBLE_VERBOSE__ (VerboseKey -> m Definition) -> VerboseKey -> m Definition
forall a b. (a -> b) -> a -> b
$
VerboseKey
"Abstract, thus, not in scope: " VerboseKey -> VerboseKey -> VerboseKey
forall a. [a] -> [a] -> [a]
++ QName -> VerboseKey
forall a. Show a => a -> VerboseKey
show QName
q
getConstInfo' :: QName -> m (Either SigError Definition)
getRewriteRulesFor :: QName -> m RewriteRules
default getConstInfo'
:: (HasConstInfo n, MonadTrans t, m ~ t n)
=> QName -> m (Either SigError Definition)
getConstInfo' = n (Either SigError Definition) -> t n (Either SigError Definition)
forall (t :: (* -> *) -> * -> *) (m :: * -> *) a.
(MonadTrans t, Monad m) =>
m a -> t m a
lift (n (Either SigError Definition)
-> t n (Either SigError Definition))
-> (QName -> n (Either SigError Definition))
-> QName
-> t n (Either SigError Definition)
forall b c a. (b -> c) -> (a -> b) -> a -> c
. QName -> n (Either SigError Definition)
forall (m :: * -> *).
HasConstInfo m =>
QName -> m (Either SigError Definition)
getConstInfo'
default getRewriteRulesFor
:: (HasConstInfo n, MonadTrans t, m ~ t n)
=> QName -> m RewriteRules
getRewriteRulesFor = n RewriteRules -> t n RewriteRules
forall (t :: (* -> *) -> * -> *) (m :: * -> *) a.
(MonadTrans t, Monad m) =>
m a -> t m a
lift (n RewriteRules -> t n RewriteRules)
-> (QName -> n RewriteRules) -> QName -> t n RewriteRules
forall b c a. (b -> c) -> (a -> b) -> a -> c
. QName -> n RewriteRules
forall (m :: * -> *). HasConstInfo m => QName -> m RewriteRules
getRewriteRulesFor
{-# SPECIALIZE getConstInfo :: QName -> TCM Definition #-}
defaultGetRewriteRulesFor :: (Monad m) => m TCState -> QName -> m RewriteRules
defaultGetRewriteRulesFor :: m TCState -> QName -> m RewriteRules
defaultGetRewriteRulesFor m TCState
getTCState QName
q = do
TCState
st <- m TCState
getTCState
let sig :: Signature
sig = TCState
stTCState -> Lens' Signature TCState -> Signature
forall o i. o -> Lens' i o -> i
^.Lens' Signature TCState
stSignature
imp :: Signature
imp = TCState
stTCState -> Lens' Signature TCState -> Signature
forall o i. o -> Lens' i o -> i
^.Lens' Signature TCState
stImports
look :: Signature -> Maybe RewriteRules
look Signature
s = QName -> RewriteRuleMap -> Maybe RewriteRules
forall k v. (Eq k, Hashable k) => k -> HashMap k v -> Maybe v
HMap.lookup QName
q (RewriteRuleMap -> Maybe RewriteRules)
-> RewriteRuleMap -> Maybe RewriteRules
forall a b. (a -> b) -> a -> b
$ Signature
s Signature -> Lens' RewriteRuleMap Signature -> RewriteRuleMap
forall o i. o -> Lens' i o -> i
^. Lens' RewriteRuleMap Signature
sigRewriteRules
RewriteRules -> m RewriteRules
forall (m :: * -> *) a. Monad m => a -> m a
return (RewriteRules -> m RewriteRules) -> RewriteRules -> m RewriteRules
forall a b. (a -> b) -> a -> b
$ [RewriteRules] -> RewriteRules
forall a. Monoid a => [a] -> a
mconcat ([RewriteRules] -> RewriteRules) -> [RewriteRules] -> RewriteRules
forall a b. (a -> b) -> a -> b
$ [Maybe RewriteRules] -> [RewriteRules]
forall a. [Maybe a] -> [a]
catMaybes [Signature -> Maybe RewriteRules
look Signature
sig, Signature -> Maybe RewriteRules
look Signature
imp]
getOriginalProjection :: HasConstInfo m => QName -> m QName
getOriginalProjection :: QName -> m QName
getOriginalProjection QName
q = Projection -> QName
projOrig (Projection -> QName)
-> (Maybe Projection -> Projection) -> Maybe Projection -> QName
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Projection -> Maybe Projection -> Projection
forall a. a -> Maybe a -> a
fromMaybe Projection
forall a. HasCallStack => a
__IMPOSSIBLE__ (Maybe Projection -> QName) -> m (Maybe Projection) -> m QName
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> QName -> m (Maybe Projection)
forall (m :: * -> *).
HasConstInfo m =>
QName -> m (Maybe Projection)
isProjection QName
q
instance HasConstInfo (TCMT IO) where
getRewriteRulesFor :: QName -> TCMT IO RewriteRules
getRewriteRulesFor = TCMT IO TCState -> QName -> TCMT IO RewriteRules
forall (m :: * -> *).
Monad m =>
m TCState -> QName -> m RewriteRules
defaultGetRewriteRulesFor TCMT IO TCState
forall (m :: * -> *). MonadTCState m => m TCState
getTC
getConstInfo' :: QName -> TCMT IO (Either SigError Definition)
getConstInfo' QName
q = do
TCState
st <- TCMT IO TCState
forall (m :: * -> *). MonadTCState m => m TCState
getTC
TCEnv
env <- TCMT IO TCEnv
forall (m :: * -> *). MonadTCEnv m => m TCEnv
askTC
TCState -> TCEnv -> QName -> TCMT IO (Either SigError Definition)
forall (m :: * -> *).
(HasOptions m, MonadDebug m, MonadTCEnv m) =>
TCState -> TCEnv -> QName -> m (Either SigError Definition)
defaultGetConstInfo TCState
st TCEnv
env QName
q
getConstInfo :: QName -> TCMT IO Definition
getConstInfo QName
q = QName -> TCMT IO (Either SigError Definition)
forall (m :: * -> *).
HasConstInfo m =>
QName -> m (Either SigError Definition)
getConstInfo' QName
q TCMT IO (Either SigError Definition)
-> (Either SigError Definition -> TCMT IO Definition)
-> TCMT IO Definition
forall (m :: * -> *) a b. Monad m => m a -> (a -> m b) -> m b
>>= \case
Right Definition
d -> Definition -> TCMT IO Definition
forall (m :: * -> *) a. Monad m => a -> m a
return Definition
d
Left (SigUnknown VerboseKey
err) -> VerboseKey -> TCMT IO Definition
forall (m :: * -> *) a. MonadFail m => VerboseKey -> m a
fail VerboseKey
err
Left SigError
SigAbstract -> QName -> TCMT IO Definition
forall a. QName -> TCM a
notInScopeError (QName -> TCMT IO Definition) -> QName -> TCMT IO Definition
forall a b. (a -> b) -> a -> b
$ QName -> QName
qnameToConcrete QName
q
defaultGetConstInfo
:: (HasOptions m, MonadDebug m, MonadTCEnv m)
=> TCState -> TCEnv -> QName -> m (Either SigError Definition)
defaultGetConstInfo :: TCState -> TCEnv -> QName -> m (Either SigError Definition)
defaultGetConstInfo TCState
st TCEnv
env QName
q = do
let defs :: Definitions
defs = TCState
stTCState -> Lens' Definitions TCState -> Definitions
forall o i. o -> Lens' i o -> i
^.((Signature -> f Signature) -> TCState -> f TCState
Lens' Signature TCState
stSignature ((Signature -> f Signature) -> TCState -> f TCState)
-> ((Definitions -> f Definitions) -> Signature -> f Signature)
-> (Definitions -> f Definitions)
-> TCState
-> f TCState
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (Definitions -> f Definitions) -> Signature -> f Signature
Lens' Definitions Signature
sigDefinitions)
idefs :: Definitions
idefs = TCState
stTCState -> Lens' Definitions TCState -> Definitions
forall o i. o -> Lens' i o -> i
^.((Signature -> f Signature) -> TCState -> f TCState
Lens' Signature TCState
stImports ((Signature -> f Signature) -> TCState -> f TCState)
-> ((Definitions -> f Definitions) -> Signature -> f Signature)
-> (Definitions -> f Definitions)
-> TCState
-> f TCState
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (Definitions -> f Definitions) -> Signature -> f Signature
Lens' Definitions Signature
sigDefinitions)
case [Maybe Definition] -> [Definition]
forall a. [Maybe a] -> [a]
catMaybes [QName -> Definitions -> Maybe Definition
forall k v. (Eq k, Hashable k) => k -> HashMap k v -> Maybe v
HMap.lookup QName
q Definitions
defs, QName -> Definitions -> Maybe Definition
forall k v. (Eq k, Hashable k) => k -> HashMap k v -> Maybe v
HMap.lookup QName
q Definitions
idefs] of
[] -> Either SigError Definition -> m (Either SigError Definition)
forall (m :: * -> *) a. Monad m => a -> m a
return (Either SigError Definition -> m (Either SigError Definition))
-> Either SigError Definition -> m (Either SigError Definition)
forall a b. (a -> b) -> a -> b
$ SigError -> Either SigError Definition
forall a b. a -> Either a b
Left (SigError -> Either SigError Definition)
-> SigError -> Either SigError Definition
forall a b. (a -> b) -> a -> b
$ VerboseKey -> SigError
SigUnknown (VerboseKey -> SigError) -> VerboseKey -> SigError
forall a b. (a -> b) -> a -> b
$ VerboseKey
"Unbound name: " VerboseKey -> VerboseKey -> VerboseKey
forall a. [a] -> [a] -> [a]
++ QName -> VerboseKey
forall a. Show a => a -> VerboseKey
show QName
q VerboseKey -> VerboseKey -> VerboseKey
forall a. [a] -> [a] -> [a]
++ QName -> VerboseKey
showQNameId QName
q
[Definition
d] -> TCEnv -> Definition -> m (Either SigError Definition)
mkAbs TCEnv
env Definition
d
[Definition]
ds -> VerboseKey -> m (Either SigError Definition)
forall (m :: * -> *) a.
(HasCallStack, MonadDebug m) =>
VerboseKey -> m a
__IMPOSSIBLE_VERBOSE__ (VerboseKey -> m (Either SigError Definition))
-> VerboseKey -> m (Either SigError Definition)
forall a b. (a -> b) -> a -> b
$ VerboseKey
"Ambiguous name: " VerboseKey -> VerboseKey -> VerboseKey
forall a. [a] -> [a] -> [a]
++ QName -> VerboseKey
forall a. Show a => a -> VerboseKey
show QName
q
where
mkAbs :: TCEnv -> Definition -> m (Either SigError Definition)
mkAbs TCEnv
env Definition
d
| QName -> TCEnv -> Bool
treatAbstractly' QName
q' TCEnv
env =
case Definition -> Maybe Definition
makeAbstract Definition
d of
Just Definition
d -> Either SigError Definition -> m (Either SigError Definition)
forall (m :: * -> *) a. Monad m => a -> m a
return (Either SigError Definition -> m (Either SigError Definition))
-> Either SigError Definition -> m (Either SigError Definition)
forall a b. (a -> b) -> a -> b
$ Definition -> Either SigError Definition
forall a b. b -> Either a b
Right Definition
d
Maybe Definition
Nothing -> Either SigError Definition -> m (Either SigError Definition)
forall (m :: * -> *) a. Monad m => a -> m a
return (Either SigError Definition -> m (Either SigError Definition))
-> Either SigError Definition -> m (Either SigError Definition)
forall a b. (a -> b) -> a -> b
$ SigError -> Either SigError Definition
forall a b. a -> Either a b
Left SigError
SigAbstract
| Bool
otherwise = Either SigError Definition -> m (Either SigError Definition)
forall (m :: * -> *) a. Monad m => a -> m a
return (Either SigError Definition -> m (Either SigError Definition))
-> Either SigError Definition -> m (Either SigError Definition)
forall a b. (a -> b) -> a -> b
$ Definition -> Either SigError Definition
forall a b. b -> Either a b
Right Definition
d
where
q' :: QName
q' = case Definition -> Defn
theDef Definition
d of
Constructor{} -> QName -> QName
dropLastModule QName
q
Defn
_ -> QName
q
dropLastModule :: QName -> QName
dropLastModule q :: QName
q@QName{ qnameModule :: QName -> ModuleName
qnameModule = ModuleName
m } =
QName
q{ qnameModule :: ModuleName
qnameModule = [Name] -> ModuleName
mnameFromList ([Name] -> ModuleName) -> [Name] -> ModuleName
forall a b. (a -> b) -> a -> b
$ [Name] -> [Name] -> ([Name] -> [Name]) -> [Name]
forall a b. Null a => a -> b -> (a -> b) -> b
ifNull (ModuleName -> [Name]
mnameToList ModuleName
m) [Name]
forall a. HasCallStack => a
__IMPOSSIBLE__ [Name] -> [Name]
forall a. [a] -> [a]
init }
instance HasConstInfo m => HasConstInfo (ChangeT m)
instance HasConstInfo m => HasConstInfo (ExceptT err m)
instance HasConstInfo m => HasConstInfo (IdentityT m)
instance HasConstInfo m => HasConstInfo (ListT m)
instance HasConstInfo m => HasConstInfo (MaybeT m)
instance HasConstInfo m => HasConstInfo (ReaderT r m)
instance HasConstInfo m => HasConstInfo (StateT s m)
instance (Monoid w, HasConstInfo m) => HasConstInfo (WriterT w m)
{-# INLINE getConInfo #-}
getConInfo :: HasConstInfo m => ConHead -> m Definition
getConInfo :: ConHead -> m Definition
getConInfo = QName -> m Definition
forall (m :: * -> *). HasConstInfo m => QName -> m Definition
getConstInfo (QName -> m Definition)
-> (ConHead -> QName) -> ConHead -> m Definition
forall b c a. (b -> c) -> (a -> b) -> a -> c
. ConHead -> QName
conName
getPolarity :: HasConstInfo m => QName -> m [Polarity]
getPolarity :: QName -> m [Polarity]
getPolarity QName
q = Definition -> [Polarity]
defPolarity (Definition -> [Polarity]) -> m Definition -> m [Polarity]
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> QName -> m Definition
forall (m :: * -> *). HasConstInfo m => QName -> m Definition
getConstInfo QName
q
getPolarity' :: HasConstInfo m => Comparison -> QName -> m [Polarity]
getPolarity' :: Comparison -> QName -> m [Polarity]
getPolarity' Comparison
CmpEq QName
q = (Polarity -> Polarity) -> [Polarity] -> [Polarity]
forall a b. (a -> b) -> [a] -> [b]
map (Polarity -> Polarity -> Polarity
composePol Polarity
Invariant) ([Polarity] -> [Polarity]) -> m [Polarity] -> m [Polarity]
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> QName -> m [Polarity]
forall (m :: * -> *). HasConstInfo m => QName -> m [Polarity]
getPolarity QName
q
getPolarity' Comparison
CmpLeq QName
q = QName -> m [Polarity]
forall (m :: * -> *). HasConstInfo m => QName -> m [Polarity]
getPolarity QName
q
setPolarity :: QName -> [Polarity] -> TCM ()
setPolarity :: QName -> [Polarity] -> TCM ()
setPolarity QName
q [Polarity]
pol = do
VerboseKey -> VerboseLevel -> TCM Doc -> TCM ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> VerboseLevel -> TCM Doc -> m ()
reportSDoc VerboseKey
"tc.polarity.set" VerboseLevel
20 (TCM Doc -> TCM ()) -> TCM Doc -> TCM ()
forall a b. (a -> b) -> a -> b
$
TCM Doc
"Setting polarity of" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> QName -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty QName
q TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> TCM Doc
"to" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> [Polarity] -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty [Polarity]
pol TCM Doc -> TCM Doc -> TCM Doc
forall a. Semigroup a => a -> a -> a
<> TCM Doc
"."
(Signature -> Signature) -> TCM ()
modifySignature ((Signature -> Signature) -> TCM ())
-> (Signature -> Signature) -> TCM ()
forall a b. (a -> b) -> a -> b
$ QName -> (Definition -> Definition) -> Signature -> Signature
updateDefinition QName
q ((Definition -> Definition) -> Signature -> Signature)
-> (Definition -> Definition) -> Signature -> Signature
forall a b. (a -> b) -> a -> b
$ ([Polarity] -> [Polarity]) -> Definition -> Definition
updateDefPolarity (([Polarity] -> [Polarity]) -> Definition -> Definition)
-> ([Polarity] -> [Polarity]) -> Definition -> Definition
forall a b. (a -> b) -> a -> b
$ [Polarity] -> [Polarity] -> [Polarity]
forall a b. a -> b -> a
const [Polarity]
pol
getForcedArgs :: HasConstInfo m => QName -> m [IsForced]
getForcedArgs :: QName -> m [IsForced]
getForcedArgs QName
q = Definition -> [IsForced]
defForced (Definition -> [IsForced]) -> m Definition -> m [IsForced]
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> QName -> m Definition
forall (m :: * -> *). HasConstInfo m => QName -> m Definition
getConstInfo QName
q
getArgOccurrence :: QName -> Nat -> TCM Occurrence
getArgOccurrence :: QName -> VerboseLevel -> TCM Occurrence
getArgOccurrence QName
d VerboseLevel
i = do
Definition
def <- QName -> TCMT IO Definition
forall (m :: * -> *). HasConstInfo m => QName -> m Definition
getConstInfo QName
d
Occurrence -> TCM Occurrence
forall (m :: * -> *) a. Monad m => a -> m a
return (Occurrence -> TCM Occurrence) -> Occurrence -> TCM Occurrence
forall a b. (a -> b) -> a -> b
$! case Definition -> Defn
theDef Definition
def of
Constructor{} -> Occurrence
StrictPos
Defn
_ -> Occurrence -> Maybe Occurrence -> Occurrence
forall a. a -> Maybe a -> a
fromMaybe Occurrence
Mixed (Maybe Occurrence -> Occurrence) -> Maybe Occurrence -> Occurrence
forall a b. (a -> b) -> a -> b
$ Definition -> [Occurrence]
defArgOccurrences Definition
def [Occurrence] -> VerboseLevel -> Maybe Occurrence
forall a. [a] -> VerboseLevel -> Maybe a
!!! VerboseLevel
i
setArgOccurrences :: QName -> [Occurrence] -> TCM ()
setArgOccurrences :: QName -> [Occurrence] -> TCM ()
setArgOccurrences QName
d [Occurrence]
os = QName -> ([Occurrence] -> [Occurrence]) -> TCM ()
modifyArgOccurrences QName
d (([Occurrence] -> [Occurrence]) -> TCM ())
-> ([Occurrence] -> [Occurrence]) -> TCM ()
forall a b. (a -> b) -> a -> b
$ [Occurrence] -> [Occurrence] -> [Occurrence]
forall a b. a -> b -> a
const [Occurrence]
os
modifyArgOccurrences :: QName -> ([Occurrence] -> [Occurrence]) -> TCM ()
modifyArgOccurrences :: QName -> ([Occurrence] -> [Occurrence]) -> TCM ()
modifyArgOccurrences QName
d [Occurrence] -> [Occurrence]
f =
(Signature -> Signature) -> TCM ()
modifySignature ((Signature -> Signature) -> TCM ())
-> (Signature -> Signature) -> TCM ()
forall a b. (a -> b) -> a -> b
$ QName -> (Definition -> Definition) -> Signature -> Signature
updateDefinition QName
d ((Definition -> Definition) -> Signature -> Signature)
-> (Definition -> Definition) -> Signature -> Signature
forall a b. (a -> b) -> a -> b
$ ([Occurrence] -> [Occurrence]) -> Definition -> Definition
updateDefArgOccurrences [Occurrence] -> [Occurrence]
f
setTreeless :: QName -> TTerm -> TCM ()
setTreeless :: QName -> TTerm -> TCM ()
setTreeless QName
q TTerm
t =
QName -> (Definition -> Definition) -> TCM ()
modifyGlobalDefinition QName
q ((Definition -> Definition) -> TCM ())
-> (Definition -> Definition) -> TCM ()
forall a b. (a -> b) -> a -> b
$ (Defn -> Defn) -> Definition -> Definition
updateTheDef ((Defn -> Defn) -> Definition -> Definition)
-> (Defn -> Defn) -> Definition -> Definition
forall a b. (a -> b) -> a -> b
$ \case
fun :: Defn
fun@Function{} -> Defn
fun{ funTreeless :: Maybe Compiled
funTreeless = Compiled -> Maybe Compiled
forall a. a -> Maybe a
Just (Compiled -> Maybe Compiled) -> Compiled -> Maybe Compiled
forall a b. (a -> b) -> a -> b
$ TTerm -> [Bool] -> Compiled
Compiled TTerm
t [] }
Defn
_ -> Defn
forall a. HasCallStack => a
__IMPOSSIBLE__
setCompiledArgUse :: QName -> [Bool] -> TCM ()
setCompiledArgUse :: QName -> [Bool] -> TCM ()
setCompiledArgUse QName
q [Bool]
use =
QName -> (Definition -> Definition) -> TCM ()
modifyGlobalDefinition QName
q ((Definition -> Definition) -> TCM ())
-> (Definition -> Definition) -> TCM ()
forall a b. (a -> b) -> a -> b
$ (Defn -> Defn) -> Definition -> Definition
updateTheDef ((Defn -> Defn) -> Definition -> Definition)
-> (Defn -> Defn) -> Definition -> Definition
forall a b. (a -> b) -> a -> b
$ \case
fun :: Defn
fun@Function{} ->
Defn
fun{ funTreeless :: Maybe Compiled
funTreeless = Maybe Compiled -> (Compiled -> Compiled) -> Maybe Compiled
forall (m :: * -> *) a b. Functor m => m a -> (a -> b) -> m b
for (Defn -> Maybe Compiled
funTreeless Defn
fun) ((Compiled -> Compiled) -> Maybe Compiled)
-> (Compiled -> Compiled) -> Maybe Compiled
forall a b. (a -> b) -> a -> b
$ \ Compiled
c -> Compiled
c { cArgUsage :: [Bool]
cArgUsage = [Bool]
use } }
Defn
_ -> Defn
forall a. HasCallStack => a
__IMPOSSIBLE__
getCompiled :: QName -> TCM (Maybe Compiled)
getCompiled :: QName -> TCM (Maybe Compiled)
getCompiled QName
q = do
(Definition -> Defn
theDef (Definition -> Defn) -> TCMT IO Definition -> TCMT IO Defn
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> QName -> TCMT IO Definition
forall (m :: * -> *). HasConstInfo m => QName -> m Definition
getConstInfo QName
q) TCMT IO Defn -> (Defn -> Maybe Compiled) -> TCM (Maybe Compiled)
forall (m :: * -> *) a b. Functor m => m a -> (a -> b) -> m b
<&> \case
Function{ funTreeless :: Defn -> Maybe Compiled
funTreeless = Maybe Compiled
t } -> Maybe Compiled
t
Defn
_ -> Maybe Compiled
forall a. Maybe a
Nothing
getErasedConArgs :: QName -> TCM [Bool]
getErasedConArgs :: QName -> TCM [Bool]
getErasedConArgs QName
q = do
Definition
def <- QName -> TCMT IO Definition
forall (m :: * -> *). HasConstInfo m => QName -> m Definition
getConstInfo QName
q
case Definition -> Defn
theDef Definition
def of
Constructor{ VerboseLevel
conArity :: Defn -> VerboseLevel
conArity :: VerboseLevel
conArity, Maybe [Bool]
conErased :: Defn -> Maybe [Bool]
conErased :: Maybe [Bool]
conErased } -> [Bool] -> TCM [Bool]
forall (m :: * -> *) a. Monad m => a -> m a
return ([Bool] -> TCM [Bool]) -> [Bool] -> TCM [Bool]
forall a b. (a -> b) -> a -> b
$
[Bool] -> Maybe [Bool] -> [Bool]
forall a. a -> Maybe a -> a
fromMaybe (VerboseLevel -> Bool -> [Bool]
forall a. VerboseLevel -> a -> [a]
replicate VerboseLevel
conArity Bool
False) Maybe [Bool]
conErased
Defn
_ -> TCM [Bool]
forall a. HasCallStack => a
__IMPOSSIBLE__
setErasedConArgs :: QName -> [Bool] -> TCM ()
setErasedConArgs :: QName -> [Bool] -> TCM ()
setErasedConArgs QName
q [Bool]
args = QName -> (Definition -> Definition) -> TCM ()
modifyGlobalDefinition QName
q ((Definition -> Definition) -> TCM ())
-> (Definition -> Definition) -> TCM ()
forall a b. (a -> b) -> a -> b
$ (Defn -> Defn) -> Definition -> Definition
updateTheDef ((Defn -> Defn) -> Definition -> Definition)
-> (Defn -> Defn) -> Definition -> Definition
forall a b. (a -> b) -> a -> b
$ \case
def :: Defn
def@Constructor{ VerboseLevel
conArity :: VerboseLevel
conArity :: Defn -> VerboseLevel
conArity }
| [Bool] -> VerboseLevel
forall (t :: * -> *) a. Foldable t => t a -> VerboseLevel
length [Bool]
args VerboseLevel -> VerboseLevel -> Bool
forall a. Eq a => a -> a -> Bool
== VerboseLevel
conArity -> Defn
def{ conErased :: Maybe [Bool]
conErased = [Bool] -> Maybe [Bool]
forall a. a -> Maybe a
Just [Bool]
args }
| Bool
otherwise -> Defn
forall a. HasCallStack => a
__IMPOSSIBLE__
Defn
def -> Defn
def
getTreeless :: QName -> TCM (Maybe TTerm)
getTreeless :: QName -> TCM (Maybe TTerm)
getTreeless QName
q = (Compiled -> TTerm) -> Maybe Compiled -> Maybe TTerm
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
fmap Compiled -> TTerm
cTreeless (Maybe Compiled -> Maybe TTerm)
-> TCM (Maybe Compiled) -> TCM (Maybe TTerm)
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> QName -> TCM (Maybe Compiled)
getCompiled QName
q
getCompiledArgUse :: QName -> TCM [Bool]
getCompiledArgUse :: QName -> TCM [Bool]
getCompiledArgUse QName
q = [Bool] -> (Compiled -> [Bool]) -> Maybe Compiled -> [Bool]
forall b a. b -> (a -> b) -> Maybe a -> b
maybe [] Compiled -> [Bool]
cArgUsage (Maybe Compiled -> [Bool]) -> TCM (Maybe Compiled) -> TCM [Bool]
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> QName -> TCM (Maybe Compiled)
getCompiled QName
q
addDataCons :: QName -> [QName] -> TCM ()
addDataCons :: QName -> [QName] -> TCM ()
addDataCons QName
d [QName]
cs = (Signature -> Signature) -> TCM ()
modifySignature ((Signature -> Signature) -> TCM ())
-> (Signature -> Signature) -> TCM ()
forall a b. (a -> b) -> a -> b
$ QName -> (Definition -> Definition) -> Signature -> Signature
updateDefinition QName
d ((Definition -> Definition) -> Signature -> Signature)
-> (Definition -> Definition) -> Signature -> Signature
forall a b. (a -> b) -> a -> b
$ (Defn -> Defn) -> Definition -> Definition
updateTheDef ((Defn -> Defn) -> Definition -> Definition)
-> (Defn -> Defn) -> Definition -> Definition
forall a b. (a -> b) -> a -> b
$ \ Defn
def ->
let !cs' :: [QName]
cs' = [QName]
cs [QName] -> [QName] -> [QName]
forall a. [a] -> [a] -> [a]
++ Defn -> [QName]
dataCons Defn
def in
case Defn
def of
Datatype{} -> Defn
def {dataCons :: [QName]
dataCons = [QName]
cs' }
Defn
_ -> Defn
forall a. HasCallStack => a
__IMPOSSIBLE__
getMutual :: QName -> TCM (Maybe [QName])
getMutual :: QName -> TCM (Maybe [QName])
getMutual QName
d = Defn -> Maybe [QName]
getMutual_ (Defn -> Maybe [QName])
-> (Definition -> Defn) -> Definition -> Maybe [QName]
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Definition -> Defn
theDef (Definition -> Maybe [QName])
-> TCMT IO Definition -> TCM (Maybe [QName])
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> QName -> TCMT IO Definition
forall (m :: * -> *). HasConstInfo m => QName -> m Definition
getConstInfo QName
d
getMutual_ :: Defn -> Maybe [QName]
getMutual_ :: Defn -> Maybe [QName]
getMutual_ = \case
Function { funMutual :: Defn -> Maybe [QName]
funMutual = Maybe [QName]
m } -> Maybe [QName]
m
Datatype { dataMutual :: Defn -> Maybe [QName]
dataMutual = Maybe [QName]
m } -> Maybe [QName]
m
Record { recMutual :: Defn -> Maybe [QName]
recMutual = Maybe [QName]
m } -> Maybe [QName]
m
Defn
_ -> Maybe [QName]
forall a. Maybe a
Nothing
setMutual :: QName -> [QName] -> TCM ()
setMutual :: QName -> [QName] -> TCM ()
setMutual QName
d [QName]
m = (Signature -> Signature) -> TCM ()
modifySignature ((Signature -> Signature) -> TCM ())
-> (Signature -> Signature) -> TCM ()
forall a b. (a -> b) -> a -> b
$ QName -> (Definition -> Definition) -> Signature -> Signature
updateDefinition QName
d ((Definition -> Definition) -> Signature -> Signature)
-> (Definition -> Definition) -> Signature -> Signature
forall a b. (a -> b) -> a -> b
$ (Defn -> Defn) -> Definition -> Definition
updateTheDef ((Defn -> Defn) -> Definition -> Definition)
-> (Defn -> Defn) -> Definition -> Definition
forall a b. (a -> b) -> a -> b
$ \ Defn
def ->
case Defn
def of
Function{} -> Defn
def { funMutual :: Maybe [QName]
funMutual = [QName] -> Maybe [QName]
forall a. a -> Maybe a
Just [QName]
m }
Datatype{} -> Defn
def {dataMutual :: Maybe [QName]
dataMutual = [QName] -> Maybe [QName]
forall a. a -> Maybe a
Just [QName]
m }
Record{} -> Defn
def { recMutual :: Maybe [QName]
recMutual = [QName] -> Maybe [QName]
forall a. a -> Maybe a
Just [QName]
m }
Defn
_ -> if [QName] -> Bool
forall a. Null a => a -> Bool
null [QName]
m then Defn
def else Defn
forall a. HasCallStack => a
__IMPOSSIBLE__
mutuallyRecursive :: QName -> QName -> TCM Bool
mutuallyRecursive :: QName -> QName -> TCMT IO Bool
mutuallyRecursive QName
d QName
d1 = (QName
d QName -> [QName] -> Bool
forall (t :: * -> *) a. (Foldable t, Eq a) => a -> t a -> Bool
`elem`) ([QName] -> Bool)
-> (Maybe [QName] -> [QName]) -> Maybe [QName] -> Bool
forall b c a. (b -> c) -> (a -> b) -> a -> c
. [QName] -> Maybe [QName] -> [QName]
forall a. a -> Maybe a -> a
fromMaybe [QName]
forall a. HasCallStack => a
__IMPOSSIBLE__ (Maybe [QName] -> Bool) -> TCM (Maybe [QName]) -> TCMT IO Bool
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> QName -> TCM (Maybe [QName])
getMutual QName
d1
definitelyNonRecursive_ :: Defn -> Bool
definitelyNonRecursive_ :: Defn -> Bool
definitelyNonRecursive_ = Bool -> ([QName] -> Bool) -> Maybe [QName] -> Bool
forall b a. b -> (a -> b) -> Maybe a -> b
maybe Bool
False [QName] -> Bool
forall a. Null a => a -> Bool
null (Maybe [QName] -> Bool) -> (Defn -> Maybe [QName]) -> Defn -> Bool
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Defn -> Maybe [QName]
getMutual_
getCurrentModuleFreeVars :: TCM Nat
getCurrentModuleFreeVars :: TCMT IO VerboseLevel
getCurrentModuleFreeVars = Telescope -> VerboseLevel
forall a. Sized a => a -> VerboseLevel
size (Telescope -> VerboseLevel)
-> TCMT IO Telescope -> TCMT IO VerboseLevel
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> (ModuleName -> TCMT IO Telescope
forall (m :: * -> *).
(Functor m, ReadTCState m) =>
ModuleName -> m Telescope
lookupSection (ModuleName -> TCMT IO Telescope)
-> TCMT IO ModuleName -> TCMT IO Telescope
forall (m :: * -> *) a b. Monad m => (a -> m b) -> m a -> m b
=<< TCMT IO ModuleName
forall (m :: * -> *). MonadTCEnv m => m ModuleName
currentModule)
getDefModule :: HasConstInfo m => QName -> m (Either SigError ModuleName)
getDefModule :: QName -> m (Either SigError ModuleName)
getDefModule QName
f = (Definition -> ModuleName)
-> Either SigError Definition -> Either SigError ModuleName
forall b d a. (b -> d) -> Either a b -> Either a d
mapRight Definition -> ModuleName
modName (Either SigError Definition -> Either SigError ModuleName)
-> m (Either SigError Definition) -> m (Either SigError ModuleName)
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> QName -> m (Either SigError Definition)
forall (m :: * -> *).
HasConstInfo m =>
QName -> m (Either SigError Definition)
getConstInfo' QName
f
where
modName :: Definition -> ModuleName
modName Definition
def = case Definition -> Defn
theDef Definition
def of
Function{ funExtLam :: Defn -> Maybe ExtLamInfo
funExtLam = Just (ExtLamInfo ModuleName
m Maybe System
_) } -> ModuleName
m
Defn
_ -> QName -> ModuleName
qnameModule QName
f
getDefFreeVars :: (Functor m, Applicative m, ReadTCState m, MonadTCEnv m) => QName -> m Nat
getDefFreeVars :: QName -> m VerboseLevel
getDefFreeVars = ModuleName -> m VerboseLevel
forall (m :: * -> *).
(Functor m, Applicative m, MonadTCEnv m, ReadTCState m) =>
ModuleName -> m VerboseLevel
getModuleFreeVars (ModuleName -> m VerboseLevel)
-> (QName -> ModuleName) -> QName -> m VerboseLevel
forall b c a. (b -> c) -> (a -> b) -> a -> c
. QName -> ModuleName
qnameModule
freeVarsToApply :: (Functor m, HasConstInfo m, HasOptions m,
ReadTCState m, MonadTCEnv m, MonadDebug m)
=> QName -> m Args
freeVarsToApply :: QName -> m [Arg Term]
freeVarsToApply QName
q = do
[Arg Term]
vs <- ModuleName -> m [Arg Term]
forall (m :: * -> *).
(Functor m, Applicative m, HasOptions m, MonadTCEnv m,
ReadTCState m, MonadDebug m) =>
ModuleName -> m [Arg Term]
moduleParamsToApply (ModuleName -> m [Arg Term]) -> ModuleName -> m [Arg Term]
forall a b. (a -> b) -> a -> b
$ QName -> ModuleName
qnameModule QName
q
Type
t <- Definition -> Type
defType (Definition -> Type) -> m Definition -> m Type
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> QName -> m Definition
forall (m :: * -> *). HasConstInfo m => QName -> m Definition
getConstInfo QName
q
let TelV Telescope
tel Type
_ = VerboseLevel -> Type -> TelV Type
telView'UpTo ([Arg Term] -> VerboseLevel
forall a. Sized a => a -> VerboseLevel
size [Arg Term]
vs) Type
t
Bool -> m () -> m ()
forall (f :: * -> *). Applicative f => Bool -> f () -> f ()
unless (Telescope -> VerboseLevel
forall a. Sized a => a -> VerboseLevel
size Telescope
tel VerboseLevel -> VerboseLevel -> Bool
forall a. Eq a => a -> a -> Bool
== [Arg Term] -> VerboseLevel
forall a. Sized a => a -> VerboseLevel
size [Arg Term]
vs) m ()
forall a. HasCallStack => a
__IMPOSSIBLE__
[Arg Term] -> m [Arg Term]
forall (m :: * -> *) a. Monad m => a -> m a
return ([Arg Term] -> m [Arg Term]) -> [Arg Term] -> m [Arg Term]
forall a b. (a -> b) -> a -> b
$ (Arg Term -> Dom (VerboseKey, Type) -> Arg Term)
-> [Arg Term] -> [Dom (VerboseKey, Type)] -> [Arg Term]
forall a b c. (a -> b -> c) -> [a] -> [b] -> [c]
zipWith (\ Arg Term
arg Dom (VerboseKey, Type)
dom -> Arg Term -> Term
forall e. Arg e -> e
unArg Arg Term
arg Term -> Arg (VerboseKey, Type) -> Arg Term
forall (f :: * -> *) a b. Functor f => a -> f b -> f a
<$ Dom (VerboseKey, Type) -> Arg (VerboseKey, Type)
forall t a. Dom' t a -> Arg a
argFromDom Dom (VerboseKey, Type)
dom) [Arg Term]
vs ([Dom (VerboseKey, Type)] -> [Arg Term])
-> [Dom (VerboseKey, Type)] -> [Arg Term]
forall a b. (a -> b) -> a -> b
$ Telescope -> [Dom (VerboseKey, Type)]
forall t. Tele (Dom t) -> [Dom (VerboseKey, t)]
telToList Telescope
tel
{-# SPECIALIZE getModuleFreeVars :: ModuleName -> TCM Nat #-}
{-# SPECIALIZE getModuleFreeVars :: ModuleName -> ReduceM Nat #-}
getModuleFreeVars :: (Functor m, Applicative m, MonadTCEnv m, ReadTCState m)
=> ModuleName -> m Nat
getModuleFreeVars :: ModuleName -> m VerboseLevel
getModuleFreeVars ModuleName
m = do
ModuleName
m0 <- ModuleName -> ModuleName -> ModuleName
commonParentModule ModuleName
m (ModuleName -> ModuleName) -> m ModuleName -> m ModuleName
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> m ModuleName
forall (m :: * -> *). MonadTCEnv m => m ModuleName
currentModule
VerboseLevel -> VerboseLevel -> VerboseLevel
forall a. Num a => a -> a -> a
(+) (VerboseLevel -> VerboseLevel -> VerboseLevel)
-> m VerboseLevel -> m (VerboseLevel -> VerboseLevel)
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> ModuleName -> m VerboseLevel
forall (m :: * -> *). MonadTCEnv m => ModuleName -> m VerboseLevel
getAnonymousVariables ModuleName
m m (VerboseLevel -> VerboseLevel)
-> m VerboseLevel -> m VerboseLevel
forall (f :: * -> *) a b. Applicative f => f (a -> b) -> f a -> f b
<*> (Telescope -> VerboseLevel
forall a. Sized a => a -> VerboseLevel
size (Telescope -> VerboseLevel) -> m Telescope -> m VerboseLevel
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> ModuleName -> m Telescope
forall (m :: * -> *).
(Functor m, ReadTCState m) =>
ModuleName -> m Telescope
lookupSection ModuleName
m0)
moduleParamsToApply :: (Functor m, Applicative m, HasOptions m,
MonadTCEnv m, ReadTCState m, MonadDebug m)
=> ModuleName -> m Args
moduleParamsToApply :: ModuleName -> m [Arg Term]
moduleParamsToApply ModuleName
m = do
VerboseKey
-> VerboseLevel -> TCM Doc -> m [Arg Term] -> m [Arg Term]
forall (m :: * -> *) a.
MonadDebug m =>
VerboseKey -> VerboseLevel -> TCM Doc -> m a -> m a
traceSDoc VerboseKey
"tc.sig.param" VerboseLevel
90 (TCM Doc
"computing module parameters of " TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> ModuleName -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty ModuleName
m) (m [Arg Term] -> m [Arg Term]) -> m [Arg Term] -> m [Arg Term]
forall a b. (a -> b) -> a -> b
$ do
m (Maybe (Substitution' Term))
-> m [Arg Term]
-> (Substitution' Term -> m [Arg Term])
-> m [Arg Term]
forall (m :: * -> *) a b.
Monad m =>
m (Maybe a) -> m b -> (a -> m b) -> m b
caseMaybeM (ModuleName -> m (Maybe (Substitution' Term))
forall (m :: * -> *).
(MonadTCEnv m, ReadTCState m) =>
ModuleName -> m (Maybe (Substitution' Term))
getModuleParameterSub ModuleName
m) ([Arg Term] -> m [Arg Term]
forall (m :: * -> *) a. Monad m => a -> m a
return []) ((Substitution' Term -> m [Arg Term]) -> m [Arg Term])
-> (Substitution' Term -> m [Arg Term]) -> m [Arg Term]
forall a b. (a -> b) -> a -> b
$ \Substitution' Term
sub -> do
VerboseKey
-> VerboseLevel -> TCM Doc -> m [Arg Term] -> m [Arg Term]
forall (m :: * -> *) a.
MonadDebug m =>
VerboseKey -> VerboseLevel -> TCM Doc -> m a -> m a
traceSDoc VerboseKey
"tc.sig.param" VerboseLevel
60 (do
[Dom (Name, Type)]
cxt <- TCMT IO [Dom (Name, Type)]
forall (m :: * -> *). MonadTCEnv m => m [Dom (Name, Type)]
getContext
VerboseLevel -> TCM Doc -> TCM Doc
forall (m :: * -> *). Functor m => VerboseLevel -> m Doc -> m Doc
nest VerboseLevel
2 (TCM Doc -> TCM Doc) -> TCM Doc -> TCM Doc
forall a b. (a -> b) -> a -> b
$ [TCM Doc] -> TCM Doc
forall (m :: * -> *). Monad m => [m Doc] -> m Doc
vcat
[ TCM Doc
"cxt = " TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> PrettyContext -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM ([Dom (Name, Type)] -> PrettyContext
PrettyContext [Dom (Name, Type)]
cxt)
, TCM Doc
"sub = " TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Substitution' Term -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty Substitution' Term
sub
]) (m [Arg Term] -> m [Arg Term]) -> m [Arg Term] -> m [Arg Term]
forall a b. (a -> b) -> a -> b
$ do
VerboseLevel
n <- ModuleName -> m VerboseLevel
forall (m :: * -> *).
(Functor m, Applicative m, MonadTCEnv m, ReadTCState m) =>
ModuleName -> m VerboseLevel
getModuleFreeVars ModuleName
m
VerboseKey
-> VerboseLevel -> TCM Doc -> m [Arg Term] -> m [Arg Term]
forall (m :: * -> *) a.
MonadDebug m =>
VerboseKey -> VerboseLevel -> TCM Doc -> m a -> m a
traceSDoc VerboseKey
"tc.sig.param" VerboseLevel
60 (VerboseLevel -> TCM Doc -> TCM Doc
forall (m :: * -> *). Functor m => VerboseLevel -> m Doc -> m Doc
nest VerboseLevel
2 (TCM Doc -> TCM Doc) -> TCM Doc -> TCM Doc
forall a b. (a -> b) -> a -> b
$ TCM Doc
"n = " TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> VerboseKey -> TCM Doc
forall (m :: * -> *). Monad m => VerboseKey -> m Doc
text (VerboseLevel -> VerboseKey
forall a. Show a => a -> VerboseKey
show VerboseLevel
n)) (m [Arg Term] -> m [Arg Term]) -> m [Arg Term] -> m [Arg Term]
forall a b. (a -> b) -> a -> b
$ do
[Dom (VerboseKey, Type)]
tel <- VerboseLevel
-> [Dom (VerboseKey, Type)] -> [Dom (VerboseKey, Type)]
forall a. VerboseLevel -> [a] -> [a]
take VerboseLevel
n ([Dom (VerboseKey, Type)] -> [Dom (VerboseKey, Type)])
-> (Telescope -> [Dom (VerboseKey, Type)])
-> Telescope
-> [Dom (VerboseKey, Type)]
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Telescope -> [Dom (VerboseKey, Type)]
forall t. Tele (Dom t) -> [Dom (VerboseKey, t)]
telToList (Telescope -> [Dom (VerboseKey, Type)])
-> m Telescope -> m [Dom (VerboseKey, Type)]
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> ModuleName -> m Telescope
forall (m :: * -> *).
(Functor m, ReadTCState m) =>
ModuleName -> m Telescope
lookupSection ModuleName
m
VerboseKey
-> VerboseLevel -> TCM Doc -> m [Arg Term] -> m [Arg Term]
forall (m :: * -> *) a.
MonadDebug m =>
VerboseKey -> VerboseLevel -> TCM Doc -> m a -> m a
traceSDoc VerboseKey
"tc.sig.param" VerboseLevel
60 (VerboseLevel -> TCM Doc -> TCM Doc
forall (m :: * -> *). Functor m => VerboseLevel -> m Doc -> m Doc
nest VerboseLevel
2 (TCM Doc -> TCM Doc) -> TCM Doc -> TCM Doc
forall a b. (a -> b) -> a -> b
$ TCM Doc
"tel = " TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> [Dom (VerboseKey, Type)] -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty [Dom (VerboseKey, Type)]
tel) (m [Arg Term] -> m [Arg Term]) -> m [Arg Term] -> m [Arg Term]
forall a b. (a -> b) -> a -> b
$ do
Bool -> m () -> m ()
forall (f :: * -> *). Applicative f => Bool -> f () -> f ()
unless ([Dom (VerboseKey, Type)] -> VerboseLevel
forall a. Sized a => a -> VerboseLevel
size [Dom (VerboseKey, Type)]
tel VerboseLevel -> VerboseLevel -> Bool
forall a. Eq a => a -> a -> Bool
== VerboseLevel
n) m ()
forall a. HasCallStack => a
__IMPOSSIBLE__
let args :: [Arg Term]
args = Substitution' Term -> [Arg Term] -> [Arg Term]
forall t a. Subst t a => Substitution' t -> a -> a
applySubst Substitution' Term
sub ([Arg Term] -> [Arg Term]) -> [Arg Term] -> [Arg Term]
forall a b. (a -> b) -> a -> b
$ (VerboseLevel -> Dom (VerboseKey, Type) -> Arg Term)
-> [VerboseLevel] -> [Dom (VerboseKey, Type)] -> [Arg Term]
forall a b c. (a -> b -> c) -> [a] -> [b] -> [c]
zipWith (\ VerboseLevel
i Dom (VerboseKey, Type)
a -> VerboseLevel -> Term
var VerboseLevel
i Term -> Arg (VerboseKey, Type) -> Arg Term
forall (f :: * -> *) a b. Functor f => a -> f b -> f a
<$ Dom (VerboseKey, Type) -> Arg (VerboseKey, Type)
forall t a. Dom' t a -> Arg a
argFromDom Dom (VerboseKey, Type)
a) (VerboseLevel -> [VerboseLevel]
forall a. Integral a => a -> [a]
downFrom VerboseLevel
n) [Dom (VerboseKey, Type)]
tel
VerboseKey
-> VerboseLevel -> TCM Doc -> m [Arg Term] -> m [Arg Term]
forall (m :: * -> *) a.
MonadDebug m =>
VerboseKey -> VerboseLevel -> TCM Doc -> m a -> m a
traceSDoc VerboseKey
"tc.sig.param" VerboseLevel
60 (VerboseLevel -> TCM Doc -> TCM Doc
forall (m :: * -> *). Functor m => VerboseLevel -> m Doc -> m Doc
nest VerboseLevel
2 (TCM Doc -> TCM Doc) -> TCM Doc -> TCM Doc
forall a b. (a -> b) -> a -> b
$ TCM Doc
"args = " TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> [TCM Doc] -> TCM Doc
forall (m :: * -> *).
(Monad m, Semigroup (m Doc)) =>
[m Doc] -> m Doc
prettyList_ ((Arg Term -> TCM Doc) -> [Arg Term] -> [TCM Doc]
forall a b. (a -> b) -> [a] -> [b]
map Arg Term -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty [Arg Term]
args)) (m [Arg Term] -> m [Arg Term]) -> m [Arg Term] -> m [Arg Term]
forall a b. (a -> b) -> a -> b
$ do
ModuleName -> m (Maybe Section)
forall (m :: * -> *).
(Functor m, ReadTCState m) =>
ModuleName -> m (Maybe Section)
getSection ModuleName
m m (Maybe Section)
-> (Maybe Section -> m [Arg Term]) -> m [Arg Term]
forall (m :: * -> *) a b. Monad m => m a -> (a -> m b) -> m b
>>= \case
Maybe Section
Nothing -> do
[Arg Term] -> m [Arg Term]
forall (m :: * -> *) a. Monad m => a -> m a
return [Arg Term]
args
Just (Section Telescope
stel) -> do
Bool -> m () -> m ()
forall (f :: * -> *). Applicative f => Bool -> f () -> f ()
when (Telescope -> VerboseLevel
forall a. Sized a => a -> VerboseLevel
size Telescope
stel VerboseLevel -> VerboseLevel -> Bool
forall a. Ord a => a -> a -> Bool
< [Arg Term] -> VerboseLevel
forall a. Sized a => a -> VerboseLevel
size [Arg Term]
args) m ()
forall a. HasCallStack => a
__IMPOSSIBLE__
[Arg Term] -> m [Arg Term]
forall (m :: * -> *) a. Monad m => a -> m a
return ([Arg Term] -> m [Arg Term]) -> [Arg Term] -> m [Arg Term]
forall a b. (a -> b) -> a -> b
$ (Dom (VerboseKey, Type) -> Arg Term -> Arg Term)
-> [Dom (VerboseKey, Type)] -> [Arg Term] -> [Arg Term]
forall a b c. (a -> b -> c) -> [a] -> [b] -> [c]
zipWith (\ !Dom (VerboseKey, Type)
dom (Arg ArgInfo
_ Term
v) -> Term
v Term -> Arg (VerboseKey, Type) -> Arg Term
forall (f :: * -> *) a b. Functor f => a -> f b -> f a
<$ Dom (VerboseKey, Type) -> Arg (VerboseKey, Type)
forall t a. Dom' t a -> Arg a
argFromDom Dom (VerboseKey, Type)
dom) (Telescope -> [Dom (VerboseKey, Type)]
forall t. Tele (Dom t) -> [Dom (VerboseKey, t)]
telToList Telescope
stel) [Arg Term]
args
inFreshModuleIfFreeParams :: TCM a -> TCM a
inFreshModuleIfFreeParams :: TCM a -> TCM a
inFreshModuleIfFreeParams TCM a
k = do
Maybe (Substitution' Term)
msub <- ModuleName -> TCMT IO (Maybe (Substitution' Term))
forall (m :: * -> *).
(MonadTCEnv m, ReadTCState m) =>
ModuleName -> m (Maybe (Substitution' Term))
getModuleParameterSub (ModuleName -> TCMT IO (Maybe (Substitution' Term)))
-> TCMT IO ModuleName -> TCMT IO (Maybe (Substitution' Term))
forall (m :: * -> *) a b. Monad m => (a -> m b) -> m a -> m b
=<< TCMT IO ModuleName
forall (m :: * -> *). MonadTCEnv m => m ModuleName
currentModule
if Maybe (Substitution' Term)
msub Maybe (Substitution' Term) -> Maybe (Substitution' Term) -> Bool
forall a. Eq a => a -> a -> Bool
== Maybe (Substitution' Term)
forall a. Maybe a
Nothing Bool -> Bool -> Bool
|| Maybe (Substitution' Term)
msub Maybe (Substitution' Term) -> Maybe (Substitution' Term) -> Bool
forall a. Eq a => a -> a -> Bool
== Substitution' Term -> Maybe (Substitution' Term)
forall a. a -> Maybe a
Just Substitution' Term
forall a. Substitution' a
IdS then TCM a
k else do
ModuleName
m <- TCMT IO ModuleName
forall (m :: * -> *). MonadTCEnv m => m ModuleName
currentModule
ModuleName
m' <- ModuleName -> ModuleName -> ModuleName
qualifyM ModuleName
m (ModuleName -> ModuleName)
-> (Name -> ModuleName) -> Name -> ModuleName
forall b c a. (b -> c) -> (a -> b) -> a -> c
. [Name] -> ModuleName
mnameFromList ([Name] -> ModuleName) -> (Name -> [Name]) -> Name -> ModuleName
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (Name -> [Name] -> [Name]
forall a. a -> [a] -> [a]
:[]) (Name -> ModuleName) -> TCMT IO Name -> TCMT IO ModuleName
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$>
VerboseKey -> TCMT IO Name
forall a (m :: * -> *).
(FreshName a, MonadFresh NameId m) =>
a -> m Name
freshName_ (VerboseKey
"_" :: String)
ModuleName -> TCM ()
addSection ModuleName
m'
ModuleName -> TCM a -> TCM a
forall (m :: * -> *) a. MonadTCEnv m => ModuleName -> m a -> m a
withCurrentModule ModuleName
m' TCM a
k
{-# SPECIALIZE instantiateDef :: Definition -> TCM Definition #-}
instantiateDef
:: ( Functor m, HasConstInfo m, HasOptions m
, ReadTCState m, MonadTCEnv m, MonadDebug m )
=> Definition -> m Definition
instantiateDef :: Definition -> m Definition
instantiateDef Definition
d = do
[Arg Term]
vs <- QName -> m [Arg Term]
forall (m :: * -> *).
(Functor m, HasConstInfo m, HasOptions m, ReadTCState m,
MonadTCEnv m, MonadDebug m) =>
QName -> m [Arg Term]
freeVarsToApply (QName -> m [Arg Term]) -> QName -> m [Arg Term]
forall a b. (a -> b) -> a -> b
$ Definition -> QName
defName Definition
d
VerboseKey -> VerboseLevel -> m () -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> VerboseLevel -> m () -> m ()
verboseS VerboseKey
"tc.sig.inst" VerboseLevel
30 (m () -> m ()) -> m () -> m ()
forall a b. (a -> b) -> a -> b
$ do
[Dom (Name, Type)]
ctx <- m [Dom (Name, Type)]
forall (m :: * -> *). MonadTCEnv m => m [Dom (Name, Type)]
getContext
ModuleName
m <- m ModuleName
forall (m :: * -> *). MonadTCEnv m => m ModuleName
currentModule
VerboseKey -> VerboseLevel -> TCM Doc -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> VerboseLevel -> TCM Doc -> m ()
reportSDoc VerboseKey
"tc.sig.inst" VerboseLevel
30 (TCM Doc -> m ()) -> TCM Doc -> m ()
forall a b. (a -> b) -> a -> b
$
TCM Doc
"instDef in" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> ModuleName -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty ModuleName
m TCM Doc -> TCM Doc -> TCM Doc
forall a. Semigroup a => a -> a -> a
<> TCM Doc
":" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> QName -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty (Definition -> QName
defName Definition
d) TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+>
VerboseKey -> TCM Doc
forall (m :: * -> *). Monad m => VerboseKey -> m Doc
text ([VerboseKey] -> VerboseKey
unwords ([VerboseKey] -> VerboseKey) -> [VerboseKey] -> VerboseKey
forall a b. (a -> b) -> a -> b
$ (Arg Name -> VerboseKey) -> [Arg Name] -> [VerboseKey]
forall a b. (a -> b) -> [a] -> [b]
map Arg Name -> VerboseKey
forall a. Show a => a -> VerboseKey
show ([Arg Name] -> [VerboseKey]) -> [Arg Name] -> [VerboseKey]
forall a b. (a -> b) -> a -> b
$ (Name -> Arg Term -> Arg Name)
-> [Name] -> [Arg Term] -> [Arg Name]
forall a b c. (a -> b -> c) -> [a] -> [b] -> [c]
zipWith Name -> Arg Term -> Arg Name
forall (f :: * -> *) a b. Functor f => a -> f b -> f a
(<$) ([Name] -> [Name]
forall a. [a] -> [a]
reverse ([Name] -> [Name]) -> [Name] -> [Name]
forall a b. (a -> b) -> a -> b
$ (Dom (Name, Type) -> Name) -> [Dom (Name, Type)] -> [Name]
forall a b. (a -> b) -> [a] -> [b]
map ((Name, Type) -> Name
forall a b. (a, b) -> a
fst ((Name, Type) -> Name)
-> (Dom (Name, Type) -> (Name, Type)) -> Dom (Name, Type) -> Name
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Dom (Name, Type) -> (Name, Type)
forall t e. Dom' t e -> e
unDom) [Dom (Name, Type)]
ctx) [Arg Term]
vs)
Definition -> m Definition
forall (m :: * -> *) a. Monad m => a -> m a
return (Definition -> m Definition) -> Definition -> m Definition
forall a b. (a -> b) -> a -> b
$ Definition
d Definition -> [Arg Term] -> Definition
forall t. Apply t => t -> [Arg Term] -> t
`apply` [Arg Term]
vs
instantiateRewriteRule :: (Functor m, HasConstInfo m, HasOptions m,
ReadTCState m, MonadTCEnv m, MonadDebug m)
=> RewriteRule -> m RewriteRule
instantiateRewriteRule :: RewriteRule -> m RewriteRule
instantiateRewriteRule RewriteRule
rew = do
VerboseKey
-> VerboseLevel -> TCM Doc -> m RewriteRule -> m RewriteRule
forall (m :: * -> *) a.
MonadDebug m =>
VerboseKey -> VerboseLevel -> TCM Doc -> m a -> m a
traceSDoc VerboseKey
"rewriting" VerboseLevel
95 (TCM Doc
"instantiating rewrite rule" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> QName -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty (RewriteRule -> QName
rewName RewriteRule
rew) TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> TCM Doc
"to the local context.") (m RewriteRule -> m RewriteRule) -> m RewriteRule -> m RewriteRule
forall a b. (a -> b) -> a -> b
$ do
[Arg Term]
vs <- QName -> m [Arg Term]
forall (m :: * -> *).
(Functor m, HasConstInfo m, HasOptions m, ReadTCState m,
MonadTCEnv m, MonadDebug m) =>
QName -> m [Arg Term]
freeVarsToApply (QName -> m [Arg Term]) -> QName -> m [Arg Term]
forall a b. (a -> b) -> a -> b
$ RewriteRule -> QName
rewName RewriteRule
rew
let rew' :: RewriteRule
rew' = RewriteRule
rew RewriteRule -> [Arg Term] -> RewriteRule
forall t. Apply t => t -> [Arg Term] -> t
`apply` [Arg Term]
vs
VerboseKey
-> VerboseLevel -> VerboseKey -> m RewriteRule -> m RewriteRule
forall (m :: * -> *) a.
MonadDebug m =>
VerboseKey -> VerboseLevel -> VerboseKey -> m a -> m a
traceSLn VerboseKey
"rewriting" VerboseLevel
95 (VerboseKey
"instantiated rewrite rule: ") (m RewriteRule -> m RewriteRule) -> m RewriteRule -> m RewriteRule
forall a b. (a -> b) -> a -> b
$ do
VerboseKey
-> VerboseLevel -> VerboseKey -> m RewriteRule -> m RewriteRule
forall (m :: * -> *) a.
MonadDebug m =>
VerboseKey -> VerboseLevel -> VerboseKey -> m a -> m a
traceSLn VerboseKey
"rewriting" VerboseLevel
95 (RewriteRule -> VerboseKey
forall a. Show a => a -> VerboseKey
show RewriteRule
rew') (m RewriteRule -> m RewriteRule) -> m RewriteRule -> m RewriteRule
forall a b. (a -> b) -> a -> b
$ do
RewriteRule -> m RewriteRule
forall (m :: * -> *) a. Monad m => a -> m a
return RewriteRule
rew'
instantiateRewriteRules :: (Functor m, HasConstInfo m, HasOptions m,
ReadTCState m, MonadTCEnv m, MonadDebug m)
=> RewriteRules -> m RewriteRules
instantiateRewriteRules :: RewriteRules -> m RewriteRules
instantiateRewriteRules = (RewriteRule -> m RewriteRule) -> RewriteRules -> m RewriteRules
forall (t :: * -> *) (m :: * -> *) a b.
(Traversable t, Monad m) =>
(a -> m b) -> t a -> m (t b)
mapM RewriteRule -> m RewriteRule
forall (m :: * -> *).
(Functor m, HasConstInfo m, HasOptions m, ReadTCState m,
MonadTCEnv m, MonadDebug m) =>
RewriteRule -> m RewriteRule
instantiateRewriteRule
makeAbstract :: Definition -> Maybe Definition
makeAbstract :: Definition -> Maybe Definition
makeAbstract Definition
d =
case Definition -> IsAbstract
defAbstract Definition
d of
IsAbstract
ConcreteDef -> Definition -> Maybe Definition
forall (m :: * -> *) a. Monad m => a -> m a
return Definition
d
IsAbstract
AbstractDef -> do
Defn
def <- Defn -> Maybe Defn
makeAbs (Defn -> Maybe Defn) -> Defn -> Maybe Defn
forall a b. (a -> b) -> a -> b
$ Definition -> Defn
theDef Definition
d
Definition -> Maybe Definition
forall (m :: * -> *) a. Monad m => a -> m a
return Definition
d { defArgOccurrences :: [Occurrence]
defArgOccurrences = []
, defPolarity :: [Polarity]
defPolarity = []
, theDef :: Defn
theDef = Defn
def
}
where
makeAbs :: Defn -> Maybe Defn
makeAbs Defn
Axiom = Defn -> Maybe Defn
forall a. a -> Maybe a
Just Defn
Axiom
makeAbs d :: Defn
d@DataOrRecSig{} = Defn -> Maybe Defn
forall a. a -> Maybe a
Just Defn
d
makeAbs d :: Defn
d@GeneralizableVar{} = Defn -> Maybe Defn
forall a. a -> Maybe a
Just Defn
d
makeAbs d :: Defn
d@Datatype {} = Defn -> Maybe Defn
forall a. a -> Maybe a
Just (Defn -> Maybe Defn) -> Defn -> Maybe Defn
forall a b. (a -> b) -> a -> b
$ Defn -> Defn
AbstractDefn Defn
d
makeAbs d :: Defn
d@Function {} = Defn -> Maybe Defn
forall a. a -> Maybe a
Just (Defn -> Maybe Defn) -> Defn -> Maybe Defn
forall a b. (a -> b) -> a -> b
$ Defn -> Defn
AbstractDefn Defn
d
makeAbs Constructor{} = Maybe Defn
forall a. Maybe a
Nothing
makeAbs d :: Defn
d@Record{} = Defn -> Maybe Defn
forall a. a -> Maybe a
Just (Defn -> Maybe Defn) -> Defn -> Maybe Defn
forall a b. (a -> b) -> a -> b
$ Defn -> Defn
AbstractDefn Defn
d
makeAbs Primitive{} = Maybe Defn
forall a. HasCallStack => a
__IMPOSSIBLE__
makeAbs AbstractDefn{}= Maybe Defn
forall a. HasCallStack => a
__IMPOSSIBLE__
{-# SPECIALIZE inAbstractMode :: TCM a -> TCM a #-}
inAbstractMode :: MonadTCEnv m => m a -> m a
inAbstractMode :: m a -> m a
inAbstractMode = (TCEnv -> TCEnv) -> m a -> m a
forall (m :: * -> *) a.
MonadTCEnv m =>
(TCEnv -> TCEnv) -> m a -> m a
localTC ((TCEnv -> TCEnv) -> m a -> m a) -> (TCEnv -> TCEnv) -> m a -> m a
forall a b. (a -> b) -> a -> b
$ \TCEnv
e -> TCEnv
e { envAbstractMode :: AbstractMode
envAbstractMode = AbstractMode
AbstractMode }
{-# SPECIALIZE inConcreteMode :: TCM a -> TCM a #-}
inConcreteMode :: MonadTCEnv m => m a -> m a
inConcreteMode :: m a -> m a
inConcreteMode = (TCEnv -> TCEnv) -> m a -> m a
forall (m :: * -> *) a.
MonadTCEnv m =>
(TCEnv -> TCEnv) -> m a -> m a
localTC ((TCEnv -> TCEnv) -> m a -> m a) -> (TCEnv -> TCEnv) -> m a -> m a
forall a b. (a -> b) -> a -> b
$ \TCEnv
e -> TCEnv
e { envAbstractMode :: AbstractMode
envAbstractMode = AbstractMode
ConcreteMode }
ignoreAbstractMode :: MonadTCEnv m => m a -> m a
ignoreAbstractMode :: m a -> m a
ignoreAbstractMode = (TCEnv -> TCEnv) -> m a -> m a
forall (m :: * -> *) a.
MonadTCEnv m =>
(TCEnv -> TCEnv) -> m a -> m a
localTC ((TCEnv -> TCEnv) -> m a -> m a) -> (TCEnv -> TCEnv) -> m a -> m a
forall a b. (a -> b) -> a -> b
$ \TCEnv
e -> TCEnv
e { envAbstractMode :: AbstractMode
envAbstractMode = AbstractMode
IgnoreAbstractMode }
{-# SPECIALIZE inConcreteOrAbstractMode :: QName -> (Definition -> TCM a) -> TCM a #-}
inConcreteOrAbstractMode :: (MonadTCEnv m, HasConstInfo m) => QName -> (Definition -> m a) -> m a
inConcreteOrAbstractMode :: QName -> (Definition -> m a) -> m a
inConcreteOrAbstractMode QName
q Definition -> m a
cont = do
Definition
def <- m Definition -> m Definition
forall (m :: * -> *) a. MonadTCEnv m => m a -> m a
ignoreAbstractMode (m Definition -> m Definition) -> m Definition -> m Definition
forall a b. (a -> b) -> a -> b
$ QName -> m Definition
forall (m :: * -> *). HasConstInfo m => QName -> m Definition
getConstInfo QName
q
case Definition -> IsAbstract
defAbstract Definition
def of
IsAbstract
AbstractDef -> m a -> m a
forall (m :: * -> *) a. MonadTCEnv m => m a -> m a
inAbstractMode (m a -> m a) -> m a -> m a
forall a b. (a -> b) -> a -> b
$ Definition -> m a
cont Definition
def
IsAbstract
ConcreteDef -> m a -> m a
forall (m :: * -> *) a. MonadTCEnv m => m a -> m a
inConcreteMode (m a -> m a) -> m a -> m a
forall a b. (a -> b) -> a -> b
$ Definition -> m a
cont Definition
def
treatAbstractly :: MonadTCEnv m => QName -> m Bool
treatAbstractly :: QName -> m Bool
treatAbstractly QName
q = (TCEnv -> Bool) -> m Bool
forall (m :: * -> *) a. MonadTCEnv m => (TCEnv -> a) -> m a
asksTC ((TCEnv -> Bool) -> m Bool) -> (TCEnv -> Bool) -> m Bool
forall a b. (a -> b) -> a -> b
$ QName -> TCEnv -> Bool
treatAbstractly' QName
q
treatAbstractly' :: QName -> TCEnv -> Bool
treatAbstractly' :: QName -> TCEnv -> Bool
treatAbstractly' QName
q TCEnv
env = case TCEnv -> AbstractMode
envAbstractMode TCEnv
env of
AbstractMode
ConcreteMode -> Bool
True
AbstractMode
IgnoreAbstractMode -> Bool
False
AbstractMode
AbstractMode -> Bool -> Bool
not (Bool -> Bool) -> Bool -> Bool
forall a b. (a -> b) -> a -> b
$ ModuleName
current ModuleName -> ModuleName -> Bool
`isLeChildModuleOf` ModuleName
m
where
current :: ModuleName
current = ModuleName -> ModuleName
dropAnon (ModuleName -> ModuleName) -> ModuleName -> ModuleName
forall a b. (a -> b) -> a -> b
$ TCEnv -> ModuleName
envCurrentModule TCEnv
env
m :: ModuleName
m = ModuleName -> ModuleName
dropAnon (ModuleName -> ModuleName) -> ModuleName -> ModuleName
forall a b. (a -> b) -> a -> b
$ QName -> ModuleName
qnameModule QName
q
dropAnon :: ModuleName -> ModuleName
dropAnon (MName [Name]
ms) = [Name] -> ModuleName
MName ([Name] -> ModuleName) -> [Name] -> ModuleName
forall a b. (a -> b) -> a -> b
$ (Name -> Bool) -> [Name] -> [Name]
forall a. (a -> Bool) -> [a] -> [a]
List.dropWhileEnd Name -> Bool
forall a. IsNoName a => a -> Bool
isNoName [Name]
ms
{-# SPECIALIZE typeOfConst :: QName -> TCM Type #-}
typeOfConst :: (HasConstInfo m, ReadTCState m) => QName -> m Type
typeOfConst :: QName -> m Type
typeOfConst QName
q = Definition -> Type
defType (Definition -> Type) -> m Definition -> m Type
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> (Definition -> m Definition
forall (m :: * -> *).
(Functor m, HasConstInfo m, HasOptions m, ReadTCState m,
MonadTCEnv m, MonadDebug m) =>
Definition -> m Definition
instantiateDef (Definition -> m Definition) -> m Definition -> m Definition
forall (m :: * -> *) a b. Monad m => (a -> m b) -> m a -> m b
=<< QName -> m Definition
forall (m :: * -> *). HasConstInfo m => QName -> m Definition
getConstInfo QName
q)
{-# SPECIALIZE relOfConst :: QName -> TCM Relevance #-}
relOfConst :: HasConstInfo m => QName -> m Relevance
relOfConst :: QName -> m Relevance
relOfConst QName
q = Definition -> Relevance
forall a. LensRelevance a => a -> Relevance
getRelevance (Definition -> Relevance) -> m Definition -> m Relevance
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> QName -> m Definition
forall (m :: * -> *). HasConstInfo m => QName -> m Definition
getConstInfo QName
q
{-# SPECIALIZE modalityOfConst :: QName -> TCM Modality #-}
modalityOfConst :: HasConstInfo m => QName -> m Modality
modalityOfConst :: QName -> m Modality
modalityOfConst QName
q = Definition -> Modality
forall a. LensModality a => a -> Modality
getModality (Definition -> Modality) -> m Definition -> m Modality
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> QName -> m Definition
forall (m :: * -> *). HasConstInfo m => QName -> m Definition
getConstInfo QName
q
droppedPars :: Definition -> Int
droppedPars :: Definition -> VerboseLevel
droppedPars Definition
d = case Definition -> Defn
theDef Definition
d of
Axiom{} -> VerboseLevel
0
DataOrRecSig{} -> VerboseLevel
0
GeneralizableVar{} -> VerboseLevel
0
def :: Defn
def@Function{} -> Defn -> VerboseLevel
projectionArgs Defn
def
Datatype {dataPars :: Defn -> VerboseLevel
dataPars = VerboseLevel
_} -> VerboseLevel
0
Record {recPars :: Defn -> VerboseLevel
recPars = VerboseLevel
_} -> VerboseLevel
0
Constructor{conPars :: Defn -> VerboseLevel
conPars = VerboseLevel
n} -> VerboseLevel
n
Primitive{} -> VerboseLevel
0
AbstractDefn{} -> VerboseLevel
forall a. HasCallStack => a
__IMPOSSIBLE__
{-# SPECIALIZE isProjection :: QName -> TCM (Maybe Projection) #-}
isProjection :: HasConstInfo m => QName -> m (Maybe Projection)
isProjection :: QName -> m (Maybe Projection)
isProjection QName
qn = Defn -> Maybe Projection
isProjection_ (Defn -> Maybe Projection)
-> (Definition -> Defn) -> Definition -> Maybe Projection
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Definition -> Defn
theDef (Definition -> Maybe Projection)
-> m Definition -> m (Maybe Projection)
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> QName -> m Definition
forall (m :: * -> *). HasConstInfo m => QName -> m Definition
getConstInfo QName
qn
isProjection_ :: Defn -> Maybe Projection
isProjection_ :: Defn -> Maybe Projection
isProjection_ Defn
def =
case Defn
def of
Function { funProjection :: Defn -> Maybe Projection
funProjection = Maybe Projection
result } -> Maybe Projection
result
Defn
_ -> Maybe Projection
forall a. Maybe a
Nothing
isStaticFun :: Defn -> Bool
isStaticFun :: Defn -> Bool
isStaticFun = (Defn -> Lens' Bool Defn -> Bool
forall o i. o -> Lens' i o -> i
^. Lens' Bool Defn
funStatic)
isInlineFun :: Defn -> Bool
isInlineFun :: Defn -> Bool
isInlineFun = (Defn -> Lens' Bool Defn -> Bool
forall o i. o -> Lens' i o -> i
^. Lens' Bool Defn
funInline)
isProperProjection :: Defn -> Bool
isProperProjection :: Defn -> Bool
isProperProjection Defn
d = Maybe Projection -> Bool -> (Projection -> Bool) -> Bool
forall a b. Maybe a -> b -> (a -> b) -> b
caseMaybe (Defn -> Maybe Projection
isProjection_ Defn
d) Bool
False ((Projection -> Bool) -> Bool) -> (Projection -> Bool) -> Bool
forall a b. (a -> b) -> a -> b
$ \ Projection
isP ->
if Projection -> VerboseLevel
projIndex Projection
isP VerboseLevel -> VerboseLevel -> Bool
forall a. Ord a => a -> a -> Bool
<= VerboseLevel
0 then Bool
False else Maybe QName -> Bool
forall a. Maybe a -> Bool
isJust (Maybe QName -> Bool) -> Maybe QName -> Bool
forall a b. (a -> b) -> a -> b
$ Projection -> Maybe QName
projProper Projection
isP
projectionArgs :: Defn -> Int
projectionArgs :: Defn -> VerboseLevel
projectionArgs = VerboseLevel
-> (Projection -> VerboseLevel) -> Maybe Projection -> VerboseLevel
forall b a. b -> (a -> b) -> Maybe a -> b
maybe VerboseLevel
0 (VerboseLevel -> VerboseLevel -> VerboseLevel
forall a. Ord a => a -> a -> a
max VerboseLevel
0 (VerboseLevel -> VerboseLevel)
-> (Projection -> VerboseLevel) -> Projection -> VerboseLevel
forall b c a. (b -> c) -> (a -> b) -> a -> c
. VerboseLevel -> VerboseLevel
forall a. Enum a => a -> a
pred (VerboseLevel -> VerboseLevel)
-> (Projection -> VerboseLevel) -> Projection -> VerboseLevel
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Projection -> VerboseLevel
projIndex) (Maybe Projection -> VerboseLevel)
-> (Defn -> Maybe Projection) -> Defn -> VerboseLevel
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Defn -> Maybe Projection
isProjection_
usesCopatterns :: (HasConstInfo m) => QName -> m Bool
usesCopatterns :: QName -> m Bool
usesCopatterns QName
q = Definition -> Bool
defCopatternLHS (Definition -> Bool) -> m Definition -> m Bool
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> QName -> m Definition
forall (m :: * -> *). HasConstInfo m => QName -> m Definition
getConstInfo QName
q
applyDef :: (HasConstInfo m)
=> ProjOrigin -> QName -> Arg Term -> m Term
applyDef :: ProjOrigin -> QName -> Arg Term -> m Term
applyDef ProjOrigin
o QName
f Arg Term
a = do
let fallback :: m Term
fallback = Term -> m Term
forall (m :: * -> *) a. Monad m => a -> m a
return (Term -> m Term) -> Term -> m Term
forall a b. (a -> b) -> a -> b
$ QName -> Elims -> Term
Def QName
f [Arg Term -> Elim' Term
forall a. Arg a -> Elim' a
Apply Arg Term
a]
m (Maybe Projection) -> m Term -> (Projection -> m Term) -> m Term
forall (m :: * -> *) a b.
Monad m =>
m (Maybe a) -> m b -> (a -> m b) -> m b
caseMaybeM (QName -> m (Maybe Projection)
forall (m :: * -> *).
HasConstInfo m =>
QName -> m (Maybe Projection)
isProjection QName
f) m Term
fallback ((Projection -> m Term) -> m Term)
-> (Projection -> m Term) -> m Term
forall a b. (a -> b) -> a -> b
$ \ Projection
isP -> do
if Projection -> VerboseLevel
projIndex Projection
isP VerboseLevel -> VerboseLevel -> Bool
forall a. Ord a => a -> a -> Bool
<= VerboseLevel
0 then m Term
fallback else do
if Maybe QName -> Bool
forall a. Maybe a -> Bool
isNothing (Projection -> Maybe QName
projProper Projection
isP) then m Term
fallback else do
Term -> m Term
forall (m :: * -> *) a. Monad m => a -> m a
return (Term -> m Term) -> Term -> m Term
forall a b. (a -> b) -> a -> b
$ Arg Term -> Term
forall e. Arg e -> e
unArg Arg Term
a Term -> Elims -> Term
forall t. Apply t => t -> Elims -> t
`applyE` [ProjOrigin -> QName -> Elim' Term
forall a. ProjOrigin -> QName -> Elim' a
Proj ProjOrigin
o (QName -> Elim' Term) -> QName -> Elim' Term
forall a b. (a -> b) -> a -> b
$ Projection -> QName
projOrig Projection
isP]