{-# OPTIONS_GHC -Wunused-imports #-}
module Agda.Compiler.CallCompiler where
import qualified Control.Exception as E
import Control.Monad.Trans
import System.Exit
import System.IO
import System.Process
import Agda.TypeChecking.Monad
import Agda.Utils.Impossible
callCompiler
  :: Bool
     
  -> FilePath
     
  -> [String]
     
  -> Maybe FilePath
     
     
  -> Maybe TextEncoding
     
     
  -> TCM ()
callCompiler :: Bool
-> FilePath
-> [FilePath]
-> Maybe FilePath
-> Maybe TextEncoding
-> TCM ()
callCompiler Bool
doCall FilePath
cmd [FilePath]
args Maybe FilePath
cwd Maybe TextEncoding
enc =
  if Bool
doCall then do
    merrors <- FilePath
-> [FilePath]
-> Maybe FilePath
-> Maybe TextEncoding
-> TCM (Maybe FilePath)
callCompiler' FilePath
cmd [FilePath]
args Maybe FilePath
cwd Maybe TextEncoding
enc
    case merrors of
      Maybe FilePath
Nothing     -> () -> TCM ()
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return ()
      Just FilePath
errors -> TypeError -> TCM ()
forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
TypeError -> m a
typeError (FilePath -> TypeError
CompilationError FilePath
errors)
  else
    FilePath -> VerboseLevel -> FilePath -> TCM ()
forall (m :: * -> *).
MonadDebug m =>
FilePath -> VerboseLevel -> FilePath -> m ()
alwaysReportSLn FilePath
"compile.cmd" VerboseLevel
1 (FilePath -> TCM ()) -> FilePath -> TCM ()
forall a b. (a -> b) -> a -> b
$ FilePath
"NOT calling: " FilePath -> FilePath -> FilePath
forall a. [a] -> [a] -> [a]
++ [FilePath] -> FilePath
unwords (FilePath
cmd FilePath -> [FilePath] -> [FilePath]
forall a. a -> [a] -> [a]
: [FilePath]
args)
callCompiler'
  :: FilePath
     
  -> [String]
     
  -> Maybe FilePath
     
     
  -> Maybe TextEncoding
     
     
  -> TCM (Maybe String)
callCompiler' :: FilePath
-> [FilePath]
-> Maybe FilePath
-> Maybe TextEncoding
-> TCM (Maybe FilePath)
callCompiler' FilePath
cmd [FilePath]
args Maybe FilePath
cwd Maybe TextEncoding
enc = do
  FilePath -> VerboseLevel -> FilePath -> TCM ()
forall (m :: * -> *).
MonadDebug m =>
FilePath -> VerboseLevel -> FilePath -> m ()
alwaysReportSLn FilePath
"compile.cmd" VerboseLevel
1 (FilePath -> TCM ()) -> FilePath -> TCM ()
forall a b. (a -> b) -> a -> b
$ FilePath
"Calling: " FilePath -> FilePath -> FilePath
forall a. [a] -> [a] -> [a]
++ [FilePath] -> FilePath
unwords (FilePath
cmd FilePath -> [FilePath] -> [FilePath]
forall a. a -> [a] -> [a]
: [FilePath]
args)
  (_, out, err, p) <-
    IO (Maybe Handle, Maybe Handle, Maybe Handle, ProcessHandle)
-> TCMT
     IO (Maybe Handle, Maybe Handle, Maybe Handle, ProcessHandle)
forall a. IO a -> TCMT IO a
forall (m :: * -> *) a. MonadIO m => IO a -> m a
liftIO (IO (Maybe Handle, Maybe Handle, Maybe Handle, ProcessHandle)
 -> TCMT
      IO (Maybe Handle, Maybe Handle, Maybe Handle, ProcessHandle))
-> IO (Maybe Handle, Maybe Handle, Maybe Handle, ProcessHandle)
-> TCMT
     IO (Maybe Handle, Maybe Handle, Maybe Handle, ProcessHandle)
forall a b. (a -> b) -> a -> b
$ CreateProcess
-> IO (Maybe Handle, Maybe Handle, Maybe Handle, ProcessHandle)
createProcess
               (FilePath -> [FilePath] -> CreateProcess
proc FilePath
cmd [FilePath]
args) { std_err = CreatePipe
                               , std_out = CreatePipe
                               , cwd     = cwd
                               }
  
  
  case out of
    Maybe Handle
Nothing  -> TCM ()
forall a. HasCallStack => a
__IMPOSSIBLE__
    Just Handle
out -> TCM () -> TCM ()
forkTCM (TCM () -> TCM ()) -> TCM () -> TCM ()
forall a b. (a -> b) -> a -> b
$ do
      
      IO () -> TCM ()
forall a. IO a -> TCMT IO a
forall (m :: * -> *) a. MonadIO m => IO a -> m a
liftIO (IO () -> TCM ()) -> IO () -> TCM ()
forall a b. (a -> b) -> a -> b
$ Handle -> Bool -> IO ()
hSetBinaryMode Handle
out Bool
False
      case Maybe TextEncoding
enc of
        Maybe TextEncoding
Nothing  -> () -> TCM ()
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return ()
        Just TextEncoding
enc -> IO () -> TCM ()
forall a. IO a -> TCMT IO a
forall (m :: * -> *) a. MonadIO m => IO a -> m a
liftIO (IO () -> TCM ()) -> IO () -> TCM ()
forall a b. (a -> b) -> a -> b
$ Handle -> TextEncoding -> IO ()
hSetEncoding Handle
out TextEncoding
enc
      progressInfo <- IO FilePath -> TCMT IO FilePath
forall a. IO a -> TCMT IO a
forall (m :: * -> *) a. MonadIO m => IO a -> m a
liftIO (IO FilePath -> TCMT IO FilePath)
-> IO FilePath -> TCMT IO FilePath
forall a b. (a -> b) -> a -> b
$ Handle -> IO FilePath
hGetContents Handle
out
      mapM_ (alwaysReportSLn "compile.output" 1) $ lines progressInfo
  errors <- liftIO $ case err of
    Maybe Handle
Nothing  -> IO FilePath
forall a. HasCallStack => a
__IMPOSSIBLE__
    Just Handle
err -> do
      
      Handle -> Bool -> IO ()
hSetBinaryMode Handle
err Bool
False
      case Maybe TextEncoding
enc of
        Maybe TextEncoding
Nothing  -> () -> IO ()
forall a. a -> IO a
forall (m :: * -> *) a. Monad m => a -> m a
return ()
        Just TextEncoding
enc -> IO () -> IO ()
forall a. IO a -> IO a
forall (m :: * -> *) a. MonadIO m => IO a -> m a
liftIO (IO () -> IO ()) -> IO () -> IO ()
forall a b. (a -> b) -> a -> b
$ Handle -> TextEncoding -> IO ()
hSetEncoding Handle
err TextEncoding
enc
      Handle -> IO FilePath
hGetContents Handle
err
  exitcode <- liftIO $ do
    
    
    _ <- E.evaluate (length errors)
    waitForProcess p
  case exitcode of
    ExitFailure VerboseLevel
_ -> Maybe FilePath -> TCM (Maybe FilePath)
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return (Maybe FilePath -> TCM (Maybe FilePath))
-> Maybe FilePath -> TCM (Maybe FilePath)
forall a b. (a -> b) -> a -> b
$ FilePath -> Maybe FilePath
forall a. a -> Maybe a
Just FilePath
errors
    ExitCode
_             -> Maybe FilePath -> TCM (Maybe FilePath)
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return Maybe FilePath
forall a. Maybe a
Nothing