module Agda.Interaction.Response.Base
( Response_boot (..)
, RemoveTokenBasedHighlighting (..)
, MakeCaseVariant (..)
, DisplayInfo_boot (..)
, GoalDisplayInfo_boot(..)
, Goals_boot
, Info_Error_boot(..)
, GoalTypeAux(..)
, ResponseContextEntry(..)
, Status (..)
, GiveResult (..)
) where
import Control.Monad.Trans ( MonadIO(liftIO) )
import Data.Set (Set)
import Data.Word (Word32)
import Agda.Interaction.Base
( CommandState
, CompilerBackend
, ComputeMode
, OutputConstraint_boot
, OutputConstraint'
, OutputForm_boot
, Rewrite
)
import Agda.Interaction.Highlighting.Precise
import qualified Agda.Syntax.Abstract as A
import Agda.Syntax.Common (InteractionId(..), Arg)
import Agda.Syntax.Concrete (Expr)
import Agda.Syntax.Concrete.Name (Name, QName, NameInScope)
import Agda.Syntax.Scope.Base (WhyInScopeData)
import qualified Agda.Syntax.Internal as I
import Agda.TypeChecking.Monad.Base.Types
(HighlightingMethod, ModuleToSource, NamedMeta, IPFace')
import Agda.Utils.Impossible
import Agda.Utils.Time
data Response_boot tcErr tcWarning warningsAndNonFatalErrors
= Resp_HighlightingInfo
HighlightingInfo
RemoveTokenBasedHighlighting
HighlightingMethod
ModuleToSource
| Resp_Status Status
| Resp_JumpToError FilePath Word32
| Resp_InteractionPoints [InteractionId]
| Resp_GiveAction InteractionId GiveResult
| Resp_MakeCase InteractionId MakeCaseVariant [String]
| Resp_SolveAll [(InteractionId, Expr)]
| Resp_Mimer InteractionId (Maybe String)
| Resp_DisplayInfo (DisplayInfo_boot tcErr tcWarning warningsAndNonFatalErrors)
| Resp_RunningInfo Int String
| Resp_ClearRunningInfo
| Resp_ClearHighlighting TokenBased
| Resp_DoneAborting
| Resp_DoneExiting
data RemoveTokenBasedHighlighting
= RemoveHighlighting
| KeepHighlighting
data MakeCaseVariant = Function | ExtendedLambda
data DisplayInfo_boot tcErr tcWarning warningsAndNonFatalErrors
= Info_CompilationOk CompilerBackend warningsAndNonFatalErrors
| Info_Constraints [OutputForm_boot tcErr Expr Expr]
| Info_AllGoalsWarnings (Goals_boot tcErr) warningsAndNonFatalErrors
| Info_Time CPUTime
| Info_Error (Info_Error_boot tcErr tcWarning)
| Info_Intro_NotFound
| Info_Intro_ConstructorUnknown [String]
| Info_Auto String
| Info_ModuleContents [Name] I.Telescope [(Name, I.Type)]
| Info_SearchAbout [(Name, I.Type)] String
| Info_WhyInScope WhyInScopeData
| Info_NormalForm CommandState ComputeMode (Maybe CPUTime) A.Expr
| Info_InferredType CommandState (Maybe CPUTime) A.Expr
| Info_Context InteractionId [ResponseContextEntry]
| Info_Version
| Info_GoalSpecific InteractionId (GoalDisplayInfo_boot tcErr)
data GoalDisplayInfo_boot tcErr
= Goal_HelperFunction (OutputConstraint' A.Expr A.Expr)
| Goal_NormalForm ComputeMode A.Expr
| Goal_GoalType Rewrite GoalTypeAux [ResponseContextEntry] [IPFace' Expr] [OutputForm_boot tcErr Expr Expr]
| Goal_CurrentGoal Rewrite
| Goal_InferredType A.Expr
type Goals_boot tcErr =
( [OutputConstraint_boot tcErr A.Expr InteractionId]
, [OutputConstraint_boot tcErr A.Expr NamedMeta]
)
data Info_Error_boot tcErr tcWarning
= Info_GenericError tcErr
| Info_CompilationError (Set tcWarning)
| Info_HighlightingParseError InteractionId
| Info_HighlightingScopeCheckError InteractionId
data GoalTypeAux
= GoalOnly
| GoalAndHave A.Expr [IPFace' Expr]
| GoalAndElaboration I.Term
data ResponseContextEntry = ResponseContextEntry
{ ResponseContextEntry -> Name
respOrigName :: Name
, ResponseContextEntry -> Name
respReifName :: Name
, ResponseContextEntry -> Arg Expr
respType :: Arg A.Expr
, ResponseContextEntry -> Maybe Expr
respLetValue :: Maybe A.Expr
, ResponseContextEntry -> NameInScope
respInScope :: NameInScope
}
data Status = Status
{ Status -> Bool
sShowImplicitArguments :: Bool
, Status -> Bool
sShowIrrelevantArguments :: Bool
, Status -> Bool
sChecked :: Bool
}
data GiveResult
= Give_String String
| Give_Paren
| Give_NoParen