module Agda.TypeChecking.Monad.Options where
import Prelude hiding (null)
import Control.Monad ( unless, when )
import Control.Monad.IO.Class ( MonadIO(..) )
import Control.Monad.Except
import Control.Monad.Reader
import Control.Monad.State
import Control.Monad.Writer
import qualified Data.Graph as Graph
import Data.List (sort)
import Data.Map (Map)
import qualified Data.Map as Map
import qualified Data.Set as Set
import System.Directory
import System.FilePath
import Agda.Syntax.Common
import Agda.Syntax.TopLevelModuleName
import Agda.TypeChecking.Monad.Debug (reportSDoc)
import Agda.TypeChecking.Warnings
import Agda.TypeChecking.Monad.Base
import Agda.TypeChecking.Monad.Imports
import Agda.TypeChecking.Monad.State
import Agda.TypeChecking.Monad.Benchmark
import Agda.TypeChecking.Monad.Trace
import Agda.Interaction.FindFile
import Agda.Interaction.Options hiding (setPragmaOptions)
import qualified Agda.Interaction.Options.Lenses as Lens
import Agda.Interaction.Library
import Agda.Interaction.Library.Base (libAbove, libFile)
import Agda.Utils.FileName
import qualified Agda.Utils.Graph.AdjacencyMap.Unidirectional as G
import Agda.Utils.Lens
import Agda.Utils.List
import Agda.Utils.List1 (List1)
import qualified Agda.Utils.List1 as List1
import Agda.Utils.Null
import Agda.Syntax.Common.Pretty
import Agda.Utils.Size
import Agda.Utils.WithDefault
import Agda.Utils.Impossible
setPragmaOptions :: PragmaOptions -> TCM ()
setPragmaOptions :: PragmaOptions -> TCM ()
setPragmaOptions PragmaOptions
opts = do
Bool -> TCM () -> TCM ()
forall (f :: * -> *). Applicative f => Bool -> f () -> f ()
when (PragmaOptions -> Bool
forall a. LensSafeMode a => a -> Bool
Lens.getSafeMode PragmaOptions
opts) (TCM () -> TCM ()) -> TCM () -> TCM ()
forall a b. (a -> b) -> a -> b
$
[FilePath] -> ([FilePath] -> TCM ()) -> TCM ()
forall (m :: * -> *) a.
(Monad m, Null a) =>
a -> (a -> m ()) -> m ()
unlessNull (PragmaOptions -> [FilePath]
unsafePragmaOptions PragmaOptions
opts) (([FilePath] -> TCM ()) -> TCM ())
-> ([FilePath] -> TCM ()) -> TCM ()
forall a b. (a -> b) -> a -> b
$ \ [FilePath]
unsafe ->
Warning -> TCM ()
forall (m :: * -> *).
(HasCallStack, MonadWarning m) =>
Warning -> m ()
warning (Warning -> TCM ()) -> Warning -> TCM ()
forall a b. (a -> b) -> a -> b
$ Set FilePath -> Warning
SafeFlagPragma (Set FilePath -> Warning) -> Set FilePath -> Warning
forall a b. (a -> b) -> a -> b
$ [FilePath] -> Set FilePath
forall a. Ord a => [a] -> Set a
Set.fromList [FilePath]
unsafe
(PragmaOptions -> f PragmaOptions) -> TCState -> f TCState
Lens' TCState PragmaOptions
stPragmaOptions Lens' TCState PragmaOptions -> PragmaOptions -> TCM ()
forall (m :: * -> *) a.
MonadTCState m =>
Lens' TCState a -> a -> m ()
`setTCLens` PragmaOptions
opts
TCM ()
updateBenchmarkingStatus
checkPragmaOptionConsistency :: PragmaOptions -> PragmaOptions -> TCM ()
checkPragmaOptionConsistency :: PragmaOptions -> PragmaOptions -> TCM ()
checkPragmaOptionConsistency PragmaOptions
oldOpts PragmaOptions
newOpts = do
(ImpliedPragmaOption -> TCM ()) -> [ImpliedPragmaOption] -> TCM ()
forall (t :: * -> *) (m :: * -> *) a b.
(Foldable t, Monad m) =>
(a -> m b) -> t a -> m ()
mapM_ ImpliedPragmaOption -> TCM ()
check [ImpliedPragmaOption]
impliedPragmaOptions
where
check :: ImpliedPragmaOption -> TCM ()
check (ImpliesPragmaOption FilePath
nameA Bool
valA PragmaOptions -> WithDefault a
flagA FilePath
nameB Bool
valB PragmaOptions -> WithDefault b
flagB)
| PragmaOptions -> WithDefault a
flagA PragmaOptions
newOpts WithDefault a -> WithDefault a -> Bool
forall a. Eq a => a -> a -> Bool
== PragmaOptions -> WithDefault a
flagA PragmaOptions
oldOpts
, PragmaOptions -> WithDefault b
flagB PragmaOptions
newOpts WithDefault b -> WithDefault b -> Bool
forall a. Eq a => a -> a -> Bool
== PragmaOptions -> WithDefault b
flagB PragmaOptions
oldOpts = () -> TCM ()
forall a. a -> TCMT IO a
forall (f :: * -> *) a. Applicative f => a -> f a
pure ()
| Value Bool
vA <- PragmaOptions -> WithDefault a
flagA PragmaOptions
newOpts, Bool
vA Bool -> Bool -> Bool
forall a. Eq a => a -> a -> Bool
== Bool
valA
, Value Bool
vB <- PragmaOptions -> WithDefault b
flagB PragmaOptions
newOpts, Bool
vB Bool -> Bool -> Bool
forall a. Eq a => a -> a -> Bool
/= Bool
valB = Warning -> TCM ()
forall (m :: * -> *).
(HasCallStack, MonadWarning m) =>
Warning -> m ()
warning (Warning -> TCM ()) -> Warning -> TCM ()
forall a b. (a -> b) -> a -> b
$ FilePath -> FilePath -> Warning
ConflictingPragmaOptions (FilePath
nameA FilePath -> Bool -> FilePath
.= Bool
valA) (FilePath
nameB FilePath -> Bool -> FilePath
.= Bool
valB)
| Bool
otherwise = () -> TCM ()
forall a. a -> TCMT IO a
forall (f :: * -> *) a. Applicative f => a -> f a
pure ()
where
FilePath
name .= :: FilePath -> Bool -> FilePath
.= Bool
True = FilePath
name
FilePath
name .= Bool
False = FilePath
"no-" FilePath -> FilePath -> FilePath
forall a. [a] -> [a] -> [a]
++ FilePath
name
setCommandLineOptions :: CommandLineOptions -> TCM ()
setCommandLineOptions :: CommandLineOptions -> TCM ()
setCommandLineOptions CommandLineOptions
opts = do
root <- IO AbsolutePath -> TCMT IO AbsolutePath
forall a. IO a -> TCMT IO a
forall (m :: * -> *) a. MonadIO m => IO a -> m a
liftIO (FilePath -> IO AbsolutePath
absolute (FilePath -> IO AbsolutePath) -> IO FilePath -> IO AbsolutePath
forall (m :: * -> *) a b. Monad m => (a -> m b) -> m a -> m b
=<< IO FilePath
getCurrentDirectory)
setCommandLineOptions' root opts
setCommandLineOptions'
:: AbsolutePath
-> CommandLineOptions
-> TCM ()
setCommandLineOptions' :: AbsolutePath -> CommandLineOptions -> TCM ()
setCommandLineOptions' AbsolutePath
root CommandLineOptions
opts = do
incs <- case CommandLineOptions -> [AbsolutePath]
optAbsoluteIncludePaths CommandLineOptions
opts of
[] -> do
opts' <- AbsolutePath -> CommandLineOptions -> TCM CommandLineOptions
setLibraryPaths AbsolutePath
root CommandLineOptions
opts
let incs = CommandLineOptions -> [FilePath]
optIncludePaths CommandLineOptions
opts'
setIncludeDirs incs root
List1.toList <$> getIncludeDirs
[AbsolutePath]
incs -> [AbsolutePath] -> TCMT IO [AbsolutePath]
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return [AbsolutePath]
incs
modifyTC $ Lens.setCommandLineOptions opts{ optAbsoluteIncludePaths = incs }
setPragmaOptions (optPragmaOptions opts)
updateBenchmarkingStatus
libToTCM :: LibM a -> TCM a
libToTCM :: forall a. LibM a -> TCM a
libToTCM LibM a
m = do
cachedConfs <- Lens' TCState (Map FilePath ProjectConfig)
-> TCMT IO (Map FilePath ProjectConfig)
forall (m :: * -> *) a. ReadTCState m => Lens' TCState a -> m a
useTC (Map FilePath ProjectConfig -> f (Map FilePath ProjectConfig))
-> TCState -> f TCState
Lens' TCState (Map FilePath ProjectConfig)
stProjectConfigs
cachedLibs <- useTC stAgdaLibFiles
((z, warns), (cachedConfs', cachedLibs')) <- liftIO $
(`runStateT` (cachedConfs, cachedLibs)) $ runWriterT $ runExceptT m
modifyTCLens stProjectConfigs $ const cachedConfs'
modifyTCLens stAgdaLibFiles $ const cachedLibs'
List1.unlessNull warns \ List1 LibWarning
warns -> List1 Warning -> TCM ()
forall (m :: * -> *).
(HasCallStack, MonadWarning m) =>
List1 Warning -> m ()
warnings (List1 Warning -> TCM ()) -> List1 Warning -> TCM ()
forall a b. (a -> b) -> a -> b
$ (LibWarning -> Warning) -> List1 LibWarning -> List1 Warning
forall a b. (a -> b) -> NonEmpty a -> NonEmpty b
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
fmap LibWarning -> Warning
LibraryWarning List1 LibWarning
warns
case z of
Left LibErrors
s -> TypeError -> TCMT IO a
forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
TypeError -> m a
typeError (TypeError -> TCMT IO a) -> TypeError -> TCMT IO a
forall a b. (a -> b) -> a -> b
$ LibErrors -> TypeError
LibraryError LibErrors
s
Right a
x -> a -> TCMT IO a
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return a
x
getAgdaLibFiles
:: AbsolutePath
-> TopLevelModuleName
-> TCM [AgdaLibFile]
getAgdaLibFiles :: AbsolutePath -> TopLevelModuleName -> TCM [AgdaLibFile]
getAgdaLibFiles AbsolutePath
f TopLevelModuleName
m = do
ls <- AbsolutePath -> TCM [AgdaLibFile]
getAgdaLibFilesWithoutTopLevelModuleName AbsolutePath
f
mapM_ (checkLibraryFileNotTooFarDown m) ls
return ls
getAgdaLibFilesWithoutTopLevelModuleName
:: AbsolutePath
-> TCM [AgdaLibFile]
getAgdaLibFilesWithoutTopLevelModuleName :: AbsolutePath -> TCM [AgdaLibFile]
getAgdaLibFilesWithoutTopLevelModuleName AbsolutePath
f = do
useLibs <- CommandLineOptions -> Bool
optUseLibs (CommandLineOptions -> Bool)
-> TCM CommandLineOptions -> TCMT IO Bool
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> TCM CommandLineOptions
forall (m :: * -> *). HasOptions m => m CommandLineOptions
commandLineOptions
if | useLibs -> libToTCM $ mkLibM [] $ getAgdaLibFiles' root
| otherwise -> return []
where
root :: FilePath
root = FilePath -> FilePath
takeDirectory (FilePath -> FilePath) -> FilePath -> FilePath
forall a b. (a -> b) -> a -> b
$ AbsolutePath -> FilePath
filePath AbsolutePath
f
checkLibraryFileNotTooFarDown ::
TopLevelModuleName ->
AgdaLibFile ->
TCM ()
checkLibraryFileNotTooFarDown :: TopLevelModuleName -> AgdaLibFile -> TCM ()
checkLibraryFileNotTooFarDown TopLevelModuleName
m AgdaLibFile
lib =
Bool -> TCM () -> TCM ()
forall (f :: * -> *). Applicative f => Bool -> f () -> f ()
when (AgdaLibFile
lib AgdaLibFile -> Lens' AgdaLibFile Int -> Int
forall o i. o -> Lens' o i -> i
^. (Int -> f Int) -> AgdaLibFile -> f AgdaLibFile
Lens' AgdaLibFile Int
libAbove Int -> Int -> Bool
forall a. Ord a => a -> a -> Bool
< TopLevelModuleName -> Int
forall a. Sized a => a -> Int
size TopLevelModuleName
m Int -> Int -> Int
forall a. Num a => a -> a -> a
- Int
1) (TCM () -> TCM ()) -> TCM () -> TCM ()
forall a b. (a -> b) -> a -> b
$ TypeError -> TCM ()
forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
TypeError -> m a
typeError (TypeError -> TCM ()) -> TypeError -> TCM ()
forall a b. (a -> b) -> a -> b
$ TopLevelModuleName -> AgdaLibFile -> TypeError
LibTooFarDown TopLevelModuleName
m AgdaLibFile
lib
getLibraryOptions
:: AbsolutePath
-> TopLevelModuleName
-> TCM [OptionsPragma]
getLibraryOptions :: AbsolutePath -> TopLevelModuleName -> TCM [OptionsPragma]
getLibraryOptions AbsolutePath
f TopLevelModuleName
m = (AgdaLibFile -> OptionsPragma) -> [AgdaLibFile] -> [OptionsPragma]
forall a b. (a -> b) -> [a] -> [b]
map AgdaLibFile -> OptionsPragma
_libPragmas ([AgdaLibFile] -> [OptionsPragma])
-> TCM [AgdaLibFile] -> TCM [OptionsPragma]
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> AbsolutePath -> TopLevelModuleName -> TCM [AgdaLibFile]
getAgdaLibFiles AbsolutePath
f TopLevelModuleName
m
setLibraryPaths
:: AbsolutePath
-> CommandLineOptions
-> TCM CommandLineOptions
setLibraryPaths :: AbsolutePath -> CommandLineOptions -> TCM CommandLineOptions
setLibraryPaths AbsolutePath
root CommandLineOptions
o =
CommandLineOptions -> TCM CommandLineOptions
setLibraryIncludes (CommandLineOptions -> TCM CommandLineOptions)
-> TCM CommandLineOptions -> TCM CommandLineOptions
forall (m :: * -> *) a b. Monad m => (a -> m b) -> m a -> m b
=<< AbsolutePath -> CommandLineOptions -> TCM CommandLineOptions
addDefaultLibraries AbsolutePath
root CommandLineOptions
o
setLibraryIncludes :: CommandLineOptions -> TCM CommandLineOptions
setLibraryIncludes :: CommandLineOptions -> TCM CommandLineOptions
setLibraryIncludes CommandLineOptions
o
| Bool -> Bool
not (CommandLineOptions -> Bool
optUseLibs CommandLineOptions
o) = CommandLineOptions -> TCM CommandLineOptions
forall a. a -> TCMT IO a
forall (f :: * -> *) a. Applicative f => a -> f a
pure CommandLineOptions
o
| Bool
otherwise = do
let libs :: [FilePath]
libs = CommandLineOptions -> [FilePath]
optLibraries CommandLineOptions
o
installed <- LibM [AgdaLibFile] -> TCM [AgdaLibFile]
forall a. LibM a -> TCM a
libToTCM (LibM [AgdaLibFile] -> TCM [AgdaLibFile])
-> LibM [AgdaLibFile] -> TCM [AgdaLibFile]
forall a b. (a -> b) -> a -> b
$ Maybe FilePath -> LibM [AgdaLibFile]
getInstalledLibraries (CommandLineOptions -> Maybe FilePath
optOverrideLibrariesFile CommandLineOptions
o)
paths <- libToTCM $ libraryIncludePaths (optOverrideLibrariesFile o) installed libs
return o{ optIncludePaths = paths ++ optIncludePaths o }
addDefaultLibraries
:: AbsolutePath
-> CommandLineOptions
-> TCM CommandLineOptions
addDefaultLibraries :: AbsolutePath -> CommandLineOptions -> TCM CommandLineOptions
addDefaultLibraries AbsolutePath
root CommandLineOptions
o
| Bool -> Bool
not ([FilePath] -> Bool
forall a. Null a => a -> Bool
null ([FilePath] -> Bool) -> [FilePath] -> Bool
forall a b. (a -> b) -> a -> b
$ CommandLineOptions -> [FilePath]
optLibraries CommandLineOptions
o) Bool -> Bool -> Bool
|| Bool -> Bool
not (CommandLineOptions -> Bool
optUseLibs CommandLineOptions
o) = CommandLineOptions -> TCM CommandLineOptions
forall a. a -> TCMT IO a
forall (f :: * -> *) a. Applicative f => a -> f a
pure CommandLineOptions
o
| Bool
otherwise = do
(libs, incs) <- LibM ([FilePath], [FilePath]) -> TCM ([FilePath], [FilePath])
forall a. LibM a -> TCM a
libToTCM (LibM ([FilePath], [FilePath]) -> TCM ([FilePath], [FilePath]))
-> LibM ([FilePath], [FilePath]) -> TCM ([FilePath], [FilePath])
forall a b. (a -> b) -> a -> b
$ FilePath -> Bool -> LibM ([FilePath], [FilePath])
getDefaultLibraries (AbsolutePath -> FilePath
filePath AbsolutePath
root) (CommandLineOptions -> Bool
optDefaultLibs CommandLineOptions
o)
return o{ optIncludePaths = incs ++ optIncludePaths o, optLibraries = libs }
addTrustedExecutables
:: CommandLineOptions
-> TCM CommandLineOptions
addTrustedExecutables :: CommandLineOptions -> TCM CommandLineOptions
addTrustedExecutables CommandLineOptions
o = do
trustedExes <- LibM (Map ExeName FilePath) -> TCM (Map ExeName FilePath)
forall a. LibM a -> TCM a
libToTCM (LibM (Map ExeName FilePath) -> TCM (Map ExeName FilePath))
-> LibM (Map ExeName FilePath) -> TCM (Map ExeName FilePath)
forall a b. (a -> b) -> a -> b
$ LibM (Map ExeName FilePath)
getTrustedExecutables
return o{ optTrustedExecutables = trustedExes }
setOptionsFromPragma :: OptionsPragma -> TCM ()
setOptionsFromPragma :: OptionsPragma -> TCM ()
setOptionsFromPragma = Bool -> OptionsPragma -> TCM ()
setOptionsFromPragma' Bool
False
checkAndSetOptionsFromPragma :: OptionsPragma -> TCM ()
checkAndSetOptionsFromPragma :: OptionsPragma -> TCM ()
checkAndSetOptionsFromPragma = Bool -> OptionsPragma -> TCM ()
setOptionsFromPragma' Bool
True
setOptionsFromPragma' :: Bool -> OptionsPragma -> TCM ()
setOptionsFromPragma' :: Bool -> OptionsPragma -> TCM ()
setOptionsFromPragma' Bool
checkConsistency OptionsPragma
ps = Range -> TCM () -> TCM ()
forall (m :: * -> *) x a.
(MonadTrace m, HasRange x) =>
x -> m a -> m a
setCurrentRange (OptionsPragma -> Range
pragmaRange OptionsPragma
ps) (TCM () -> TCM ()) -> TCM () -> TCM ()
forall a b. (a -> b) -> a -> b
$ do
opts <- TCM CommandLineOptions
forall (m :: * -> *). HasOptions m => m CommandLineOptions
commandLineOptions
let (z, warns) = runOptM (parsePragmaOptions ps opts)
mapM_ (warning . OptionWarning) warns
case z of
Left FilePath
err -> TypeError -> TCM ()
forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
TypeError -> m a
typeError (TypeError -> TCM ()) -> TypeError -> TCM ()
forall a b. (a -> b) -> a -> b
$ FilePath -> TypeError
OptionError FilePath
err
Right PragmaOptions
opts' -> do
Bool -> TCM () -> TCM ()
forall (f :: * -> *). Applicative f => Bool -> f () -> f ()
when Bool
checkConsistency (TCM () -> TCM ()) -> TCM () -> TCM ()
forall a b. (a -> b) -> a -> b
$ do
oldOpts <- TCMT IO PragmaOptions
forall (m :: * -> *). HasOptions m => m PragmaOptions
pragmaOptions
checkPragmaOptionConsistency oldOpts opts'
PragmaOptions -> TCM ()
setPragmaOptions PragmaOptions
opts'
enableDisplayForms :: MonadTCEnv m => m a -> m a
enableDisplayForms :: forall (m :: * -> *) a. MonadTCEnv m => m a -> m a
enableDisplayForms =
(TCEnv -> TCEnv) -> m a -> m a
forall a. (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 { envDisplayFormsEnabled = True }
disableDisplayForms :: MonadTCEnv m => m a -> m a
disableDisplayForms :: forall (m :: * -> *) a. MonadTCEnv m => m a -> m a
disableDisplayForms =
(TCEnv -> TCEnv) -> m a -> m a
forall a. (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 { envDisplayFormsEnabled = False }
displayFormsEnabled :: MonadTCEnv m => m Bool
displayFormsEnabled :: forall (m :: * -> *). MonadTCEnv m => m Bool
displayFormsEnabled = (TCEnv -> Bool) -> m Bool
forall (m :: * -> *) a. MonadTCEnv m => (TCEnv -> a) -> m a
asksTC TCEnv -> Bool
envDisplayFormsEnabled
getIncludeDirs :: HasOptions m => m (List1 AbsolutePath)
getIncludeDirs :: forall (m :: * -> *). HasOptions m => m (NonEmpty AbsolutePath)
getIncludeDirs = do
NonEmpty AbsolutePath -> [AbsolutePath] -> NonEmpty AbsolutePath
forall a. List1 a -> [a] -> List1 a
List1.fromListSafe NonEmpty AbsolutePath
forall a. HasCallStack => a
__IMPOSSIBLE__ ([AbsolutePath] -> NonEmpty AbsolutePath)
-> (CommandLineOptions -> [AbsolutePath])
-> CommandLineOptions
-> NonEmpty AbsolutePath
forall b c a. (b -> c) -> (a -> b) -> a -> c
. CommandLineOptions -> [AbsolutePath]
optAbsoluteIncludePaths (CommandLineOptions -> NonEmpty AbsolutePath)
-> m CommandLineOptions -> m (NonEmpty AbsolutePath)
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> m CommandLineOptions
forall (m :: * -> *). HasOptions m => m CommandLineOptions
commandLineOptions
setIncludeDirs :: [FilePath]
-> AbsolutePath
-> TCM ()
setIncludeDirs :: [FilePath] -> AbsolutePath -> TCM ()
setIncludeDirs [FilePath]
incs AbsolutePath
root = do
oldIncs <- (TCState -> [AbsolutePath]) -> TCMT IO [AbsolutePath]
forall (m :: * -> *) a. ReadTCState m => (TCState -> a) -> m a
getsTC TCState -> [AbsolutePath]
forall a. LensIncludePaths a => a -> [AbsolutePath]
Lens.getAbsoluteIncludePaths
incs <- return $ List1.fromListSafe (List1.singleton ".") incs
incs <- return $ fmap (mkAbsolute . (filePath root </>)) incs
primdir <- liftIO $ mkAbsolute <$> getPrimitiveLibDir
incs <- return $ List1.fromListSafe __IMPOSSIBLE__ $ nubOn id $ List1.toList $ incs <> List1.singleton primdir
reportSDoc "setIncludeDirs" 10 $ return $ vcat
[ "Old include directories:"
, nest 2 $ vcat $ map pretty oldIncs
, "New include directories:"
, nest 2 $ vcat $ fmap pretty incs
]
when (sort oldIncs /= sort (List1.toList incs)) $ do
ho <- getInteractionOutputCallback
tcWarnings <- useTC stTCWarnings
projectConfs <- useTC stProjectConfigs
agdaLibFiles <- useTC stAgdaLibFiles
decodedModules <- getDecodedModules
(keptDecodedModules, modFile) <- modulesToKeep incs decodedModules
resetAllState
setTCLens stTCWarnings tcWarnings
setTCLens stProjectConfigs projectConfs
setTCLens stAgdaLibFiles agdaLibFiles
setInteractionOutputCallback ho
setDecodedModules keptDecodedModules
setTCLens stModuleToSource modFile
Lens.putAbsoluteIncludePaths $ List1.toList incs
where
modulesToKeep
:: List1 AbsolutePath
-> DecodedModules
-> TCM (DecodedModules, ModuleToSource)
modulesToKeep :: NonEmpty AbsolutePath
-> DecodedModules -> TCM (DecodedModules, ModuleToSource)
modulesToKeep NonEmpty AbsolutePath
incs DecodedModules
old = DecodedModules
-> ModuleToSource
-> [ModuleInfo]
-> TCM (DecodedModules, ModuleToSource)
process DecodedModules
forall k a. Map k a
Map.empty ModuleToSource
forall k a. Map k a
Map.empty [ModuleInfo]
modules
where
dependencyGraph :: G.Graph TopLevelModuleName ()
dependencyGraph :: Graph TopLevelModuleName ()
dependencyGraph =
[TopLevelModuleName] -> Graph TopLevelModuleName ()
forall n e. Ord n => [n] -> Graph n e
G.fromNodes
[ Interface -> TopLevelModuleName
iTopLevelModuleName (Interface -> TopLevelModuleName)
-> Interface -> TopLevelModuleName
forall a b. (a -> b) -> a -> b
$ ModuleInfo -> Interface
miInterface ModuleInfo
m
| ModuleInfo
m <- DecodedModules -> [ModuleInfo]
forall k a. Map k a -> [a]
Map.elems DecodedModules
old
]
Graph TopLevelModuleName ()
-> Graph TopLevelModuleName () -> Graph TopLevelModuleName ()
forall n e. Ord n => Graph n e -> Graph n e -> Graph n e
`G.union`
[Edge TopLevelModuleName ()] -> Graph TopLevelModuleName ()
forall n e. Ord n => [Edge n e] -> Graph n e
G.fromEdges
[ G.Edge
{ source :: TopLevelModuleName
source = Interface -> TopLevelModuleName
iTopLevelModuleName (Interface -> TopLevelModuleName)
-> Interface -> TopLevelModuleName
forall a b. (a -> b) -> a -> b
$ ModuleInfo -> Interface
miInterface ModuleInfo
m
, target :: TopLevelModuleName
target = TopLevelModuleName
d
, label :: ()
label = ()
}
| ModuleInfo
m <- DecodedModules -> [ModuleInfo]
forall k a. Map k a -> [a]
Map.elems DecodedModules
old
, (TopLevelModuleName
d, Hash
_) <- Interface -> [(TopLevelModuleName, Hash)]
iImportedModules (Interface -> [(TopLevelModuleName, Hash)])
-> Interface -> [(TopLevelModuleName, Hash)]
forall a b. (a -> b) -> a -> b
$ ModuleInfo -> Interface
miInterface ModuleInfo
m
]
modules :: [ModuleInfo]
modules :: [ModuleInfo]
modules =
(SCC TopLevelModuleName -> ModuleInfo)
-> [SCC TopLevelModuleName] -> [ModuleInfo]
forall a b. (a -> b) -> [a] -> [b]
map (\case
Graph.CyclicSCC{} ->
ModuleInfo
forall a. HasCallStack => a
__IMPOSSIBLE__
Graph.AcyclicSCC TopLevelModuleName
m ->
case TopLevelModuleName -> DecodedModules -> Maybe ModuleInfo
forall k a. Ord k => k -> Map k a -> Maybe a
Map.lookup TopLevelModuleName
m DecodedModules
old of
Just ModuleInfo
m -> ModuleInfo
m
Maybe ModuleInfo
Nothing -> ModuleInfo
forall a. HasCallStack => a
__IMPOSSIBLE__) ([SCC TopLevelModuleName] -> [ModuleInfo])
-> [SCC TopLevelModuleName] -> [ModuleInfo]
forall a b. (a -> b) -> a -> b
$
Graph TopLevelModuleName () -> [SCC TopLevelModuleName]
forall n e. Ord n => Graph n e -> [SCC n]
G.sccs' Graph TopLevelModuleName ()
dependencyGraph
process ::
Map TopLevelModuleName ModuleInfo -> ModuleToSource ->
[ModuleInfo] -> TCM (DecodedModules, ModuleToSource)
process :: DecodedModules
-> ModuleToSource
-> [ModuleInfo]
-> TCM (DecodedModules, ModuleToSource)
process !DecodedModules
keep !ModuleToSource
modFile [] = (DecodedModules, ModuleToSource)
-> TCM (DecodedModules, ModuleToSource)
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return
( [(TopLevelModuleName, ModuleInfo)] -> DecodedModules
forall k a. Ord k => [(k, a)] -> Map k a
Map.fromList ([(TopLevelModuleName, ModuleInfo)] -> DecodedModules)
-> [(TopLevelModuleName, ModuleInfo)] -> DecodedModules
forall a b. (a -> b) -> a -> b
$
DecodedModules -> [(TopLevelModuleName, ModuleInfo)]
forall k a. Map k a -> [(k, a)]
Map.toList DecodedModules
keep
, ModuleToSource
modFile
)
process DecodedModules
keep ModuleToSource
modFile (ModuleInfo
m : [ModuleInfo]
ms) = do
let deps :: [TopLevelModuleName]
deps = ((TopLevelModuleName, Hash) -> TopLevelModuleName)
-> [(TopLevelModuleName, Hash)] -> [TopLevelModuleName]
forall a b. (a -> b) -> [a] -> [b]
map (TopLevelModuleName, Hash) -> TopLevelModuleName
forall a b. (a, b) -> a
fst ([(TopLevelModuleName, Hash)] -> [TopLevelModuleName])
-> [(TopLevelModuleName, Hash)] -> [TopLevelModuleName]
forall a b. (a -> b) -> a -> b
$ Interface -> [(TopLevelModuleName, Hash)]
iImportedModules (Interface -> [(TopLevelModuleName, Hash)])
-> Interface -> [(TopLevelModuleName, Hash)]
forall a b. (a -> b) -> a -> b
$ ModuleInfo -> Interface
miInterface ModuleInfo
m
depsKept :: Bool
depsKept = (TopLevelModuleName -> Bool) -> [TopLevelModuleName] -> Bool
forall (t :: * -> *) a. Foldable t => (a -> Bool) -> t a -> Bool
all (TopLevelModuleName -> DecodedModules -> Bool
forall k a. Ord k => k -> Map k a -> Bool
`Map.member` DecodedModules
keep) [TopLevelModuleName]
deps
(keep, modFile) <-
if Bool -> Bool
not Bool
depsKept then (DecodedModules, ModuleToSource)
-> TCM (DecodedModules, ModuleToSource)
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return (DecodedModules
keep, ModuleToSource
modFile) else do
let t :: TopLevelModuleName
t = Interface -> TopLevelModuleName
iTopLevelModuleName (Interface -> TopLevelModuleName)
-> Interface -> TopLevelModuleName
forall a b. (a -> b) -> a -> b
$ ModuleInfo -> Interface
miInterface ModuleInfo
m
oldF <- TopLevelModuleName -> TCM (Either FindError SourceFile)
findFile' TopLevelModuleName
t
(newF, modFile) <- liftIO $ findFile'' incs t modFile
return $ case (oldF, newF) of
(Right SourceFile
f1, Right SourceFile
f2) | SourceFile
f1 SourceFile -> SourceFile -> Bool
forall a. Eq a => a -> a -> Bool
== SourceFile
f2 ->
(TopLevelModuleName
-> ModuleInfo -> DecodedModules -> DecodedModules
forall k a. Ord k => k -> a -> Map k a -> Map k a
Map.insert TopLevelModuleName
t ModuleInfo
m DecodedModules
keep, ModuleToSource
modFile)
(Either FindError SourceFile, Either FindError SourceFile)
_ -> (DecodedModules
keep, ModuleToSource
modFile)
process keep modFile ms
isPropEnabled :: HasOptions m => m Bool
isPropEnabled :: forall (m :: * -> *). HasOptions m => m Bool
isPropEnabled = PragmaOptions -> Bool
optProp (PragmaOptions -> Bool) -> m PragmaOptions -> m Bool
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> m PragmaOptions
forall (m :: * -> *). HasOptions m => m PragmaOptions
pragmaOptions
isLevelUniverseEnabled :: HasOptions m => m Bool
isLevelUniverseEnabled :: forall (m :: * -> *). HasOptions m => m Bool
isLevelUniverseEnabled = PragmaOptions -> Bool
optLevelUniverse (PragmaOptions -> Bool) -> m PragmaOptions -> m Bool
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> m PragmaOptions
forall (m :: * -> *). HasOptions m => m PragmaOptions
pragmaOptions
isTwoLevelEnabled :: HasOptions m => m Bool
isTwoLevelEnabled :: forall (m :: * -> *). HasOptions m => m Bool
isTwoLevelEnabled = PragmaOptions -> Bool
optTwoLevel (PragmaOptions -> Bool) -> m PragmaOptions -> m Bool
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> m PragmaOptions
forall (m :: * -> *). HasOptions m => m PragmaOptions
pragmaOptions
{-# SPECIALIZE hasUniversePolymorphism :: TCM Bool #-}
hasUniversePolymorphism :: HasOptions m => m Bool
hasUniversePolymorphism :: forall (m :: * -> *). HasOptions m => m Bool
hasUniversePolymorphism = PragmaOptions -> Bool
optUniversePolymorphism (PragmaOptions -> Bool) -> m PragmaOptions -> m Bool
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> m PragmaOptions
forall (m :: * -> *). HasOptions m => m PragmaOptions
pragmaOptions
showImplicitArguments :: HasOptions m => m Bool
showImplicitArguments :: forall (m :: * -> *). HasOptions m => m Bool
showImplicitArguments = PragmaOptions -> Bool
optShowImplicit (PragmaOptions -> Bool) -> m PragmaOptions -> m Bool
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> m PragmaOptions
forall (m :: * -> *). HasOptions m => m PragmaOptions
pragmaOptions
showGeneralizedArguments :: HasOptions m => m Bool
showGeneralizedArguments :: forall (m :: * -> *). HasOptions m => m Bool
showGeneralizedArguments = (\PragmaOptions
opt -> PragmaOptions -> Bool
optShowGeneralized PragmaOptions
opt) (PragmaOptions -> Bool) -> m PragmaOptions -> m Bool
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> m PragmaOptions
forall (m :: * -> *). HasOptions m => m PragmaOptions
pragmaOptions
showIrrelevantArguments :: HasOptions m => m Bool
showIrrelevantArguments :: forall (m :: * -> *). HasOptions m => m Bool
showIrrelevantArguments = PragmaOptions -> Bool
optShowIrrelevant (PragmaOptions -> Bool) -> m PragmaOptions -> m Bool
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> m PragmaOptions
forall (m :: * -> *). HasOptions m => m PragmaOptions
pragmaOptions
showIdentitySubstitutions :: HasOptions m => m Bool
showIdentitySubstitutions :: forall (m :: * -> *). HasOptions m => m Bool
showIdentitySubstitutions = PragmaOptions -> Bool
optShowIdentitySubstitutions (PragmaOptions -> Bool) -> m PragmaOptions -> m Bool
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> m PragmaOptions
forall (m :: * -> *). HasOptions m => m PragmaOptions
pragmaOptions
withShowAllArguments :: ReadTCState m => m a -> m a
withShowAllArguments :: forall (m :: * -> *) a. ReadTCState m => m a -> m a
withShowAllArguments = Bool -> m a -> m a
forall (m :: * -> *) a. ReadTCState m => Bool -> m a -> m a
withShowAllArguments' Bool
True
withShowAllArguments' :: ReadTCState m => Bool -> m a -> m a
withShowAllArguments' :: forall (m :: * -> *) a. ReadTCState m => Bool -> m a -> m a
withShowAllArguments' Bool
yes = (PragmaOptions -> PragmaOptions) -> m a -> m a
forall (m :: * -> *) a.
ReadTCState m =>
(PragmaOptions -> PragmaOptions) -> m a -> m a
withPragmaOptions ((PragmaOptions -> PragmaOptions) -> m a -> m a)
-> (PragmaOptions -> PragmaOptions) -> m a -> m a
forall a b. (a -> b) -> a -> b
$ \ PragmaOptions
opts ->
PragmaOptions
opts { _optShowImplicit = Value yes, _optShowIrrelevant = Value yes }
withoutPrintingGeneralization :: ReadTCState m => m a -> m a
withoutPrintingGeneralization :: forall (m :: * -> *) a. ReadTCState m => m a -> m a
withoutPrintingGeneralization = (PragmaOptions -> PragmaOptions) -> m a -> m a
forall (m :: * -> *) a.
ReadTCState m =>
(PragmaOptions -> PragmaOptions) -> m a -> m a
withPragmaOptions ((PragmaOptions -> PragmaOptions) -> m a -> m a)
-> (PragmaOptions -> PragmaOptions) -> m a -> m a
forall a b. (a -> b) -> a -> b
$ \ PragmaOptions
opts ->
PragmaOptions
opts { _optShowGeneralized = Value False }
withPragmaOptions :: ReadTCState m => (PragmaOptions -> PragmaOptions) -> m a -> m a
withPragmaOptions :: forall (m :: * -> *) a.
ReadTCState m =>
(PragmaOptions -> PragmaOptions) -> m a -> m a
withPragmaOptions = Lens' TCState PragmaOptions
-> (PragmaOptions -> PragmaOptions) -> m a -> m a
forall a b. Lens' TCState a -> (a -> a) -> m b -> m b
forall (m :: * -> *) a b.
ReadTCState m =>
Lens' TCState a -> (a -> a) -> m b -> m b
locallyTCState (PragmaOptions -> f PragmaOptions) -> TCState -> f TCState
Lens' TCState PragmaOptions
stPragmaOptions
positivityCheckEnabled :: HasOptions m => m Bool
positivityCheckEnabled :: forall (m :: * -> *). HasOptions m => m Bool
positivityCheckEnabled = PragmaOptions -> Bool
optPositivityCheck (PragmaOptions -> Bool) -> m PragmaOptions -> m Bool
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> m PragmaOptions
forall (m :: * -> *). HasOptions m => m PragmaOptions
pragmaOptions
{-# SPECIALIZE typeInType :: TCM Bool #-}
typeInType :: HasOptions m => m Bool
typeInType :: forall (m :: * -> *). HasOptions m => m Bool
typeInType = Bool -> Bool
not (Bool -> Bool) -> (PragmaOptions -> Bool) -> PragmaOptions -> Bool
forall b c a. (b -> c) -> (a -> b) -> a -> c
. PragmaOptions -> Bool
optUniverseCheck (PragmaOptions -> Bool) -> m PragmaOptions -> m Bool
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> m PragmaOptions
forall (m :: * -> *). HasOptions m => m PragmaOptions
pragmaOptions
etaEnabled :: HasOptions m => m Bool
etaEnabled :: forall (m :: * -> *). HasOptions m => m Bool
etaEnabled = PragmaOptions -> Bool
optEta (PragmaOptions -> Bool) -> m PragmaOptions -> m Bool
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> m PragmaOptions
forall (m :: * -> *). HasOptions m => m PragmaOptions
pragmaOptions
maxInstanceSearchDepth :: HasOptions m => m Int
maxInstanceSearchDepth :: forall (m :: * -> *). HasOptions m => m Int
maxInstanceSearchDepth = PragmaOptions -> Int
optInstanceSearchDepth (PragmaOptions -> Int) -> m PragmaOptions -> m Int
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> m PragmaOptions
forall (m :: * -> *). HasOptions m => m PragmaOptions
pragmaOptions
maxInversionDepth :: HasOptions m => m Int
maxInversionDepth :: forall (m :: * -> *). HasOptions m => m Int
maxInversionDepth = PragmaOptions -> Int
optInversionMaxDepth (PragmaOptions -> Int) -> m PragmaOptions -> m Int
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> m PragmaOptions
forall (m :: * -> *). HasOptions m => m PragmaOptions
pragmaOptions
getLanguage :: HasOptions m => m Language
getLanguage :: forall (m :: * -> *). HasOptions m => m Language
getLanguage = do
opts <- m PragmaOptions
forall (m :: * -> *). HasOptions m => m PragmaOptions
pragmaOptions
return $
if not (optWithoutK opts) then WithK else
case optCubical opts of
Just Cubical
variant -> Cubical -> Language
Cubical Cubical
variant
Maybe Cubical
Nothing -> Language
WithoutK