Safe Haskell | None |
---|---|
Language | Haskell2010 |
Agda.TypeChecking.Monad.Base
Contents
Synopsis
- pattern Axiom :: Bool -> Defn
- pattern Constructor :: Int -> Int -> ConHead -> QName -> IsAbstract -> CompKit -> Maybe [QName] -> [IsForced] -> Maybe [Bool] -> Bool -> Bool -> Defn
- pattern DDot :: Term -> DisplayTerm
- pattern DTerm :: Term -> DisplayTerm
- pattern DataOrRecSig :: Int -> Defn
- pattern Datatype :: Nat -> Nat -> Maybe Clause -> [QName] -> Sort -> Maybe [QName] -> IsAbstract -> [QName] -> Maybe QName -> Maybe QName -> Defn
- pattern Function :: [Clause] -> Maybe CompiledClauses -> Maybe SplitTree -> Maybe Compiled -> [Clause] -> FunctionInverse -> Maybe [QName] -> Either ProjectionLikenessMissing Projection -> SmallSet FunctionFlag -> Maybe Bool -> Maybe ExtLamInfo -> Maybe QName -> Maybe QName -> IsOpaque -> Defn
- pattern PProp :: NLPat -> NLPSort
- pattern PSSet :: NLPat -> NLPSort
- pattern PType :: NLPat -> NLPSort
- pattern Primitive :: IsAbstract -> PrimitiveId -> [Clause] -> FunctionInverse -> Maybe CompiledClauses -> IsOpaque -> Defn
- pattern PrimitiveSort :: BuiltinSort -> Sort -> Defn
- pattern Record :: Nat -> Maybe Clause -> ConHead -> Bool -> [Dom QName] -> Telescope -> Maybe [QName] -> EtaEquality -> PatternOrCopattern -> Maybe Induction -> Maybe Bool -> IsAbstract -> CompKit -> Defn
- pattern SortProp :: BuiltinSort
- pattern SortPropOmega :: BuiltinSort
- pattern SortSet :: BuiltinSort
- pattern SortSetOmega :: BuiltinSort
- pattern SortStrictSet :: BuiltinSort
- pattern SortStrictSetOmega :: BuiltinSort
- _mvInfo :: Lens' MetaVariable MetaInfo
- _recEtaEquality :: RecordData -> HasEta
- aDefToMode :: IsAbstract -> AbstractMode
- aModeToDef :: AbstractMode -> Maybe IsAbstract
- absurdLambdaName :: String
- allReductions :: AllowedReductions
- apReduce :: ReduceM (a -> b) -> ReduceM a -> ReduceM b
- apTCMT :: forall (m :: Type -> Type) a b. Applicative m => TCMT m (a -> b) -> TCMT m a -> TCMT m b
- askR :: ReduceM ReduceEnv
- asksTC :: MonadTCEnv m => (TCEnv -> a) -> m a
- axiomConstTransp :: Defn -> Bool
- beforeReduce :: ReduceM a -> ReduceM b -> ReduceM a
- bindReduce :: ReduceM a -> (a -> ReduceM b) -> ReduceM b
- bindTCMT :: forall (m :: Type -> Type) a b. Monad m => TCMT m a -> (a -> TCMT m b) -> TCMT m b
- buildClosure :: (MonadTCEnv m, ReadTCState m) => a -> m (Closure a)
- catchError_ :: TCM a -> (TCErr -> TCM a) -> TCM a
- conAbstr :: Defn -> IsAbstract
- conArity :: Defn -> Int
- conComp :: Defn -> CompKit
- conData :: Defn -> QName
- conErased :: Defn -> Maybe [Bool]
- conErasure :: Defn -> Bool
- conForced :: Defn -> [IsForced]
- conInline :: Defn -> Bool
- conPars :: Defn -> Int
- conProj :: Defn -> Maybe [QName]
- conSrcCon :: Defn -> ConHead
- constTranspAxiom :: Defn
- cubicalCompatibleOption :: HasOptions m => m Bool
- cubicalOption :: HasOptions m => m (Maybe Cubical)
- currentModality :: MonadTCEnv m => m Modality
- dataAbstr :: Defn -> IsAbstract
- dataClause :: Defn -> Maybe Clause
- dataCons :: Defn -> [QName]
- dataIxs :: Defn -> Nat
- dataMutual :: Defn -> Maybe [QName]
- dataPars :: Defn -> Nat
- dataPathCons :: Defn -> [QName]
- dataSort :: Defn -> Sort
- dataTransp :: Defn -> Maybe QName
- dataTranspIx :: Defn -> Maybe QName
- datarecPars :: Defn -> Int
- defAbstract :: Definition -> IsAbstract
- defClauses :: Definition -> [Clause]
- defCompiled :: Definition -> Maybe CompiledClauses
- defCompilerPragmas :: BackendName -> Definition -> [CompilerPragma]
- defConstructors :: Defn -> [QName]
- defForced :: Definition -> [IsForced]
- defInverse :: Definition -> FunctionInverse
- defIsDataOrRecord :: Defn -> Bool
- defIsRecord :: Defn -> Bool
- defNonterminating :: Definition -> Bool
- defOpaque :: Definition -> IsOpaque
- defParameters :: Definition -> Maybe Nat
- defTerminationUnconfirmed :: Definition -> Bool
- defaultAxiom :: Defn
- defaultDefn :: ArgInfo -> QName -> Type -> Language -> Defn -> Definition
- defaultDisplayForm :: QName -> [LocalDisplayForm]
- defaultInteractionOutputCallback :: InteractionOutputCallback
- defaultUnquoteFlags :: UnquoteFlags
- dirToCmp :: (Comparison -> a -> a -> c) -> CompareDirection -> a -> a -> c
- eAbstractMode :: Lens' TCEnv AbstractMode
- eActiveBackendName :: Lens' TCEnv (Maybe BackendName)
- eActiveProblems :: Lens' TCEnv (Set ProblemId)
- eAllowedReductions :: Lens' TCEnv AllowedReductions
- eAnonymousModules :: Lens' TCEnv [(ModuleName, Nat)]
- eAppDef :: Lens' TCEnv (Maybe QName)
- eAssignMetas :: Lens' TCEnv Bool
- eCall :: Lens' TCEnv (Maybe (Closure Call))
- eCallByNeed :: Lens' TCEnv Bool
- eCheckingWhere :: Lens' TCEnv WhereClause_
- eCheckpoints :: Lens' TCEnv (Map CheckpointId Substitution)
- eCompareBlocked :: Lens' TCEnv Bool
- eConflComputingOverlap :: Lens' TCEnv Bool
- eContext :: Lens' TCEnv Context
- eCoverageCheck :: Lens' TCEnv CoverageCheck
- eCurrentCheckpoint :: Lens' TCEnv CheckpointId
- eCurrentModule :: Lens' TCEnv ModuleName
- eCurrentPath :: Lens' TCEnv (Maybe FileId)
- eCurrentlyElaborating :: Lens' TCEnv Bool
- eDisplayFormsEnabled :: Lens' TCEnv Bool
- eExpandLast :: Lens' TCEnv ExpandHidden
- eExpandLastBool :: Lens' TCEnv Bool
- eFoldLetBindings :: Lens' TCEnv Bool
- eGeneralizeMetas :: Lens' TCEnv DoGeneralize
- eGeneralizedVars :: Lens' TCEnv (Map QName GeneralizedValue)
- eHardCompileTimeMode :: Lens' TCEnv Bool
- eHighlightingLevel :: Lens' TCEnv HighlightingLevel
- eHighlightingMethod :: Lens' TCEnv HighlightingMethod
- eHighlightingRange :: Lens' TCEnv Range
- eImportPath :: Lens' TCEnv [TopLevelModuleName]
- eInjectivityDepth :: Lens' TCEnv Int
- eInsideDotPattern :: Lens' TCEnv Bool
- eInstanceDepth :: Lens' TCEnv Int
- eIsDebugPrinting :: Lens' TCEnv Bool
- eLetBindings :: Lens' TCEnv LetBindings
- eMakeCase :: Lens' TCEnv Bool
- eMutualBlock :: Lens' TCEnv (Maybe MutualId)
- ePrintDomainFreePi :: Lens' TCEnv Bool
- ePrintMetasBare :: Lens' TCEnv Bool
- ePrintingPatternLambdas :: Lens' TCEnv [QName]
- eQuantity :: Lens' TCEnv Quantity
- eRange :: Lens' TCEnv Range
- eReconstructed :: Lens' TCEnv Bool
- eReduceDefs :: Lens' TCEnv ReduceDefs
- eReduceDefsPair :: Lens' TCEnv (Bool, [QName])
- eRelevance :: Lens' TCEnv Relevance
- eSimplification :: Lens' TCEnv Simplification
- eSolvingConstraints :: Lens' TCEnv Bool
- eSplitOnStrict :: Lens' TCEnv Bool
- eTerminationCheck :: Lens' TCEnv (TerminationCheck ())
- eUnquoteFlags :: Lens' TCEnv UnquoteFlags
- eUnquoteNormalise :: Lens' TCEnv Bool
- eWorkingOnTypes :: Lens' TCEnv Bool
- emptyCompKit :: CompKit
- emptyFunction :: HasOptions m => m Defn
- emptyFunctionData :: HasOptions m => m FunctionData
- emptyFunctionData_ :: Bool -> FunctionData
- emptyFunction_ :: Bool -> Defn
- emptySignature :: Signature
- enableCaching :: HasOptions m => m Bool
- execError :: (HasCallStack, MonadTCError m) => ExecError -> m a
- extendedLambdaName :: String
- finally_ :: TCM a -> TCM b -> TCM a
- flipCmp :: CompareDirection -> CompareDirection
- fmapReduce :: (a -> b) -> ReduceM a -> ReduceM b
- fmapTCMT :: forall (m :: Type -> Type) a b. Functor m => (a -> b) -> TCMT m a -> TCMT m b
- forkTCM :: TCM () -> TCM ()
- freshName :: MonadFresh NameId m => Range -> String -> m Name
- freshNoName :: MonadFresh NameId m => Range -> m Name
- freshNoName_ :: MonadFresh NameId m => m Name
- freshRecordName :: MonadFresh NameId m => m Name
- fromCmp :: Comparison -> CompareDirection
- fromReduceDefs :: ReduceDefs -> (Bool, [QName])
- funAbstr :: Lens' Defn IsAbstract
- funAbstr_ :: Lens' FunctionData IsAbstract
- funAbstract :: Lens' Defn Bool
- funAbstract_ :: Lens' FunctionData Bool
- funClauses :: Defn -> [Clause]
- funCompiled :: Defn -> Maybe CompiledClauses
- funCovering :: Defn -> [Clause]
- funErasure :: Lens' Defn Bool
- funExtLam :: Defn -> Maybe ExtLamInfo
- funFirstOrder :: Lens' Defn Bool
- funFlag :: FunctionFlag -> Lens' Defn Bool
- funFlag_ :: FunctionFlag -> Lens' FunctionData Bool
- funFlags :: Defn -> SmallSet FunctionFlag
- funInline :: Lens' Defn Bool
- funInv :: Defn -> FunctionInverse
- funIsKanOp :: Defn -> Maybe QName
- funMacro :: Lens' Defn Bool
- funMacro_ :: Lens' FunctionData Bool
- funMutual :: Defn -> Maybe [QName]
- funOpaque :: Defn -> IsOpaque
- funProj :: Lens' Defn Bool
- funProj_ :: Lens' FunctionData Bool
- funProjection :: Defn -> Either ProjectionLikenessMissing Projection
- funSplitTree :: Defn -> Maybe SplitTree
- funStatic :: Lens' Defn Bool
- funTerminates :: Defn -> Maybe Bool
- funTreeless :: Defn -> Maybe Compiled
- funWith :: Defn -> Maybe QName
- generalizedFieldName :: String
- genericDocError :: (HasCallStack, MonadTCError m) => Doc -> m a
- genericError :: (HasCallStack, MonadTCError m) => String -> m a
- getGeneralizedFieldName :: QName -> Maybe String
- getMetaEnv :: MetaVariable -> TCEnv
- getMetaInfo :: MetaVariable -> Closure Range
- getMetaScope :: MetaVariable -> ScopeInfo
- getMetaSig :: MetaVariable -> Signature
- getPartialDefs :: ReadTCState m => m (Set QName)
- getUserWarnings :: ReadTCState m => m UserWarnings
- getsTC :: ReadTCState m => (TCState -> a) -> m a
- ghcBackendName :: BackendName
- guardednessOption :: HasOptions m => m Bool
- highMetaPriority :: MetaPriority
- iFullHash :: Interface -> Hash
- ifTopLevelAndHighlightingLevelIs :: MonadTCEnv tcm => HighlightingLevel -> tcm () -> tcm ()
- ifTopLevelAndHighlightingLevelIsOr :: MonadTCEnv tcm => HighlightingLevel -> Bool -> tcm () -> tcm ()
- illegalRewriteWarningName :: IllegalRewriteRuleReason -> WarningName
- initEnv :: TCEnv
- initFileDict :: AbsolutePath -> FileDictWithBuiltins
- initPersistentState :: AbsolutePath -> PersistentTCState
- initPersistentStateFromSessionState :: SessionTCState -> PersistentTCState
- initPostScopeState :: PostScopeState
- initPreScopeState :: PreScopeState
- initSessionState :: AbsolutePath -> SessionTCState
- initState :: AbsolutePath -> TCState
- initStateFromPersistentState :: PersistentTCState -> TCState
- initStateFromSessionState :: SessionTCState -> TCState
- initStateIO :: IO TCState
- initialMetaId :: MetaId
- intSignature :: Lens' Interface Signature
- interactionError :: (HasCallStack, MonadTCError m) => InteractionError -> m a
- internalError :: (HasCallStack, MonadTCError m) => String -> m a
- isAbsurdLambdaName :: QName -> Bool
- isCopatternLHS :: [Clause] -> Bool
- isDontExpandLast :: ExpandHidden -> Bool
- isEmptyFunction :: Defn -> Bool
- isExpandLast :: ExpandHidden -> Bool
- isExtendedLambda :: Defn -> Bool
- isExtendedLambdaName :: QName -> Bool
- isMacro :: Defn -> Bool
- isReconstructed :: MonadTCEnv m => m Bool
- isSourceCodeWarning :: WarningName -> Bool
- isWithFunction :: Defn -> Bool
- itableCounts :: Lens' InstanceTable (Map QName Int)
- itableTree :: Lens' InstanceTable (DiscrimTree QName)
- jsBackendName :: BackendName
- lensAreWeCaching :: Lens' PostScopeState Bool
- lensAwakeConstraints :: Lens' PostScopeState Constraints
- lensBackends :: Lens' SessionTCState [Backend]
- lensBenchmark :: Lens' SessionTCState Benchmark
- lensBuiltinModuleIds :: Lens' SessionTCState BuiltinModuleIds
- lensConcreteNames :: Lens' PostScopeState ConcreteNames
- lensConsideringInstance :: Lens' PostScopeState Bool
- lensConstructor :: Lens' Defn ConstructorData
- lensCopiedNames :: Lens' PreScopeState (HashMap QName QName)
- lensDirty :: Lens' PostScopeState Bool
- lensDisambiguatedNames :: Lens' PostScopeState DisambiguatedNames
- lensFileDict :: Lens' SessionTCState FileDictWithBuiltins
- lensFileDictBuilder :: Lens' SessionTCState FileDictBuilder
- lensForeignCode :: Lens' PostScopeState BackendForeignCode
- lensFreshCheckpointId :: Lens' PostScopeState CheckpointId
- lensFreshInt :: Lens' PostScopeState Int
- lensFreshInteractionId :: Lens' PreScopeState InteractionId
- lensFreshMetaId :: Lens' PostScopeState MetaId
- lensFreshMutualId :: Lens' PostScopeState MutualId
- lensFreshNameId :: Lens' PostScopeState NameId
- lensFreshOpaqueId :: Lens' PostScopeState OpaqueId
- lensFreshProblemId :: Lens' PostScopeState ProblemId
- lensFunction :: Lens' Defn FunctionData
- lensGeneralizedVars :: Lens' PreScopeState (Maybe (Set QName))
- lensImportedBuiltins :: Lens' PreScopeState BuiltinThings
- lensImportedDisplayForms :: Lens' PreScopeState DisplayForms
- lensImportedMetaStore :: Lens' PreScopeState RemoteMetaStore
- lensImportedModules :: Lens' PreScopeState ImportedModules
- lensImportedPartialDefs :: Lens' PreScopeState (Set QName)
- lensImportedUserWarnings :: Lens' PreScopeState UserWarnings
- lensImports :: Lens' PreScopeState Signature
- lensImportsDisplayForms :: Lens' PostScopeState DisplayForms
- lensInstanceHack :: Lens' PostScopeState Bool
- lensInstantiateBlocking :: Lens' PostScopeState Bool
- lensInteractionPoints :: Lens' PostScopeState InteractionPoints
- lensLibCache :: Lens' PreScopeState LibCache
- lensLoadedFileCache :: Lens' PersistentTCState (Maybe LoadedFileCache)
- lensLocalBuiltins :: Lens' PostScopeState BuiltinThings
- lensLocalPartialDefs :: Lens' PostScopeState (Set QName)
- lensLocalUserWarnings :: Lens' PreScopeState UserWarnings
- lensModuleCheckpoints :: Lens' PostScopeState (Map ModuleName CheckpointId)
- lensModuleToSourceId :: Lens' PreScopeState ModuleToSourceId
- lensMutualBlocks :: Lens' PostScopeState MutualBlocks
- lensNameCopies :: Lens' PreScopeState (HashMap QName (HashSet QName))
- lensOccursCheckDefs :: Lens' PostScopeState (Set QName)
- lensOpaqueBlocks :: Lens' PostScopeState (Map OpaqueId OpaqueBlock)
- lensOpaqueIds :: Lens' PostScopeState (Map QName OpaqueId)
- lensOpenMetaStore :: Lens' PostScopeState LocalMetaStore
- lensPatternSynImports :: Lens' PreScopeState PatternSynDefns
- lensPatternSyns :: Lens' PreScopeState PatternSynDefns
- lensPersistentSession :: Lens' PersistentTCState SessionTCState
- lensPersistentState :: Lens' TCState PersistentTCState
- lensPostScopeState :: Lens' TCState PostScopeState
- lensPostponeInstanceSearch :: Lens' PostScopeState Bool
- lensPreScopeState :: Lens' TCState PreScopeState
- lensPreTokens :: Lens' PreScopeState HighlightingInfo
- lensPrimitiveLibDir :: Lens' SessionTCState PrimitiveLibDir
- lensRecEta :: Lens' RecordData EtaEquality
- lensRecTel :: Lens' RecordData Telescope
- lensRecord :: Lens' Defn RecordData
- lensScope :: Lens' PreScopeState ScopeInfo
- lensSessionState :: Lens' TCState SessionTCState
- lensShadowingNames :: Lens' PostScopeState ShadowingNames
- lensSignature :: Lens' PostScopeState Signature
- lensSleepingConstraints :: Lens' PostScopeState Constraints
- lensSolvedMetaStore :: Lens' PostScopeState LocalMetaStore
- lensStatistics :: Lens' PostScopeState Statistics
- lensSyntaxInfo :: Lens' PostScopeState HighlightingInfo
- lensTCWarnings :: Lens' PostScopeState (Set TCWarning)
- lensTemporaryInstances :: Lens' PostScopeState (Set QName)
- lensTheDef :: Lens' Definition Defn
- lensTopLevelModuleNames :: Lens' PersistentTCState (BiMap RawTopLevelModuleName ModuleNameHash)
- lensUsedNames :: Lens' PostScopeState UsedNames
- lensVisitedModules :: Lens' PreScopeState VisitedModules
- lensWarningOnImport :: Lens' PreScopeState (Maybe Text)
- localR :: (ReduceEnv -> ReduceEnv) -> ReduceM a -> ReduceM a
- locallyReconstructed :: MonadTCEnv m => m a -> m a
- locallyReduceAllDefs :: MonadTCEnv m => m a -> m a
- locallyReduceDefs :: MonadTCEnv m => ReduceDefs -> m a -> m a
- locallyTC :: MonadTCEnv m => Lens' TCEnv a -> (a -> a) -> m b -> m b
- locatedTypeError :: MonadTCError m => (a -> TypeError) -> HasCallStack => a -> m b
- lowMetaPriority :: MetaPriority
- mapRedEnv :: (TCEnv -> TCEnv) -> ReduceEnv -> ReduceEnv
- mapRedEnvSt :: (TCEnv -> TCEnv) -> (TCState -> TCState) -> ReduceEnv -> ReduceEnv
- mapRedSt :: (TCState -> TCState) -> ReduceEnv -> ReduceEnv
- mapTCMT :: (forall a1. m a1 -> n a1) -> TCMT m a -> TCMT n a
- metaFrozen :: Lens' MetaVariable Frozen
- modifySystem :: (System -> System) -> ExtLamInfo -> ExtLamInfo
- modifyTC' :: MonadTCState m => (TCState -> TCState) -> m ()
- modifyTCLens :: MonadTCState m => Lens' TCState a -> (a -> a) -> m ()
- modifyTCLens' :: MonadTCState m => Lens' TCState a -> (a -> a) -> m ()
- modifyTCLensM :: (MonadTCState m, ReadTCState m) => Lens' TCState a -> (a -> m a) -> m ()
- nextFresh :: HasFresh i => TCState -> (i, TCState)
- noCompiledRep :: CompiledRepresentation
- normalMetaPriority :: MetaPriority
- notReduced :: a -> MaybeReduced a
- onLetBindingType :: (Dom Type -> Dom Type) -> LetBinding -> LetBinding
- onReduceEnv :: (ReduceEnv -> ReduceEnv) -> ReduceM a -> ReduceM a
- patternInTeleName :: String
- primAbstr :: Defn -> IsAbstract
- primClauses :: Defn -> [Clause]
- primCompiled :: Defn -> Maybe CompiledClauses
- primFun :: QName -> Arity -> ([Arg Term] -> ReduceM (Reduced MaybeReducedArgs Term)) -> PrimFun
- primInv :: Defn -> FunctionInverse
- primName :: Defn -> PrimitiveId
- primOpaque :: Defn -> IsOpaque
- primSortName :: Defn -> BuiltinSort
- primSortSort :: Defn -> Sort
- projArgInfo :: Projection -> ArgInfo
- projDropPars :: Projection -> ProjOrigin -> Term
- pureTCM :: forall (m :: Type -> Type) a. MonadIO m => (TCState -> TCEnv -> a) -> TCMT m a
- reallyAllReductions :: AllowedReductions
- recAbstr :: Defn -> IsAbstract
- recClause :: Defn -> Maybe Clause
- recComp :: Defn -> CompKit
- recCon :: Defn -> QName
- recConHead :: Defn -> ConHead
- recEtaEquality :: Defn -> HasEta
- recEtaEquality' :: Defn -> EtaEquality
- recFields :: Defn -> [Dom QName]
- recInduction :: Defn -> Maybe Induction
- recMutual :: Defn -> Maybe [QName]
- recNamedCon :: Defn -> Bool
- recPars :: Defn -> Nat
- recPatternMatching :: Defn -> PatternOrCopattern
- recRecursive :: Defn -> Bool
- recRecursive_ :: RecordData -> Bool
- recTel :: Defn -> Telescope
- recTerminates :: Defn -> Maybe Bool
- recordFieldWarningToError :: RecordFieldWarning -> TypeError
- redBind :: ReduceM (Reduced a a') -> (a -> ReduceM b) -> (a' -> ReduceM (Reduced b b')) -> ReduceM (Reduced b b')
- redReturn :: a -> ReduceM (Reduced a' a)
- reduceAllDefs :: ReduceDefs
- reduceEnv :: Lens' ReduceEnv TCEnv
- reduceSt :: Lens' ReduceEnv TCState
- reduced :: Blocked (Arg Term) -> MaybeReduced (Arg Term)
- returnTCMT :: forall (m :: Type -> Type) a. Applicative m => a -> TCMT m a
- runBlocked :: Monad m => BlockT m a -> m (Either Blocker a)
- runReduceM :: ReduceM a -> TCM a
- runSafeTCM :: TCM a -> TCState -> IO (a, TCState)
- runTCM :: MonadIO m => TCEnv -> TCState -> TCMT m a -> m (a, TCState)
- runTCMTop :: TCM a -> IO (Either TCErr a)
- runTCMTop' :: MonadIO m => TCMT m a -> m a
- secTelescope :: Lens' Section Telescope
- setEtaEquality :: EtaEquality -> HasEta -> EtaEquality
- setTCLens :: MonadTCState m => Lens' TCState a -> a -> m ()
- setTCLens' :: MonadTCState m => Lens' TCState a -> a -> m ()
- shouldReduceDef :: MonadTCEnv m => QName -> m Bool
- sigDefinitions :: Lens' Signature Definitions
- sigInstances :: Lens' Signature InstanceTable
- sigRewriteRules :: Lens' Signature RewriteRuleMap
- sigSections :: Lens' Signature Sections
- sizedTypesOption :: HasOptions m => m Bool
- srcFilePath :: MonadFileId m => SourceFile -> m AbsolutePath
- srcFromPath :: MonadFileId m => AbsolutePath -> m SourceFile
- stAreWeCaching :: Lens' TCState Bool
- stAwakeConstraints :: Lens' TCState Constraints
- stBackends :: Lens' TCState [Backend]
- stBenchmark :: Lens' TCState Benchmark
- stBuiltinModuleIds :: Lens' TCState BuiltinModuleIds
- stBuiltinThings :: TCState -> BuiltinThings
- stConcreteNames :: Lens' TCState ConcreteNames
- stConsideringInstance :: Lens' TCState Bool
- stCopiedNames :: Lens' TCState (HashMap QName QName)
- stCurrentModule :: Lens' TCState (Maybe (ModuleName, TopLevelModuleName))
- stDirty :: Lens' TCState Bool
- stDisambiguatedNames :: Lens' TCState DisambiguatedNames
- stFileDict :: Lens' TCState FileDictWithBuiltins
- stForeignCode :: Lens' TCState BackendForeignCode
- stFreshCheckpointId :: Lens' TCState CheckpointId
- stFreshInt :: Lens' TCState Int
- stFreshInteractionId :: Lens' TCState InteractionId
- stFreshMetaId :: Lens' TCState MetaId
- stFreshMutualId :: Lens' TCState MutualId
- stFreshNameId :: Lens' TCState NameId
- stFreshOpaqueId :: Lens' TCState OpaqueId
- stFreshProblemId :: Lens' TCState ProblemId
- stGeneralizedVars :: Lens' TCState (Maybe (Set QName))
- stImportedBuiltins :: Lens' TCState BuiltinThings
- stImportedDisplayForms :: Lens' TCState DisplayForms
- stImportedMetaStore :: Lens' TCState RemoteMetaStore
- stImportedModules :: Lens' TCState ImportedModules
- stImportedPartialDefs :: Lens' TCState (Set QName)
- stImportedUserWarnings :: Lens' TCState UserWarnings
- stImports :: Lens' TCState Signature
- stImportsDisplayForms :: Lens' TCState DisplayForms
- stInstanceDefs :: Lens' TCState TempInstanceTable
- stInstanceHack :: Lens' TCState Bool
- stInstanceTree :: Lens' TCState (DiscrimTree QName)
- stInstantiateBlocking :: Lens' TCState Bool
- stInteractionPoints :: Lens' TCState InteractionPoints
- stLibCache :: Lens' TCState LibCache
- stLoadedFileCache :: Lens' TCState (Maybe LoadedFileCache)
- stLocalBuiltins :: Lens' TCState BuiltinThings
- stLocalPartialDefs :: Lens' TCState (Set QName)
- stLocalUserWarnings :: Lens' TCState UserWarnings
- stModuleCheckpoints :: Lens' TCState (Map ModuleName CheckpointId)
- stModuleToSource :: Lens' TCState ModuleToSource
- stModuleToSourceId :: Lens' TCState ModuleToSourceId
- stMutualBlocks :: Lens' TCState MutualBlocks
- stNameCopies :: Lens' TCState (HashMap QName (HashSet QName))
- stOccursCheckDefs :: Lens' TCState (Set QName)
- stOpaqueBlocks :: Lens' TCState (Map OpaqueId OpaqueBlock)
- stOpaqueIds :: Lens' TCState (Map QName OpaqueId)
- stOpenMetaStore :: Lens' TCState LocalMetaStore
- stPatternSynImports :: Lens' TCState PatternSynDefns
- stPatternSyns :: Lens' TCState PatternSynDefns
- stPostponeInstanceSearch :: Lens' TCState Bool
- stPragmaOptions :: Lens' TCState PragmaOptions
- stPrimitiveLibDir :: Lens' TCState PrimitiveLibDir
- stScope :: Lens' TCState ScopeInfo
- stShadowingNames :: Lens' TCState ShadowingNames
- stSignature :: Lens' TCState Signature
- stSleepingConstraints :: Lens' TCState Constraints
- stSolvedMetaStore :: Lens' TCState LocalMetaStore
- stStatistics :: Lens' TCState Statistics
- stSyntaxInfo :: Lens' TCState HighlightingInfo
- stTCWarnings :: Lens' TCState (Set TCWarning)
- stTemporaryInstances :: Lens' TCState (Set QName)
- stTokens :: Lens' TCState HighlightingInfo
- stTopLevelModuleNames :: Lens' TCState (BiMap RawTopLevelModuleName ModuleNameHash)
- stUsedNames :: Lens' TCState UsedNames
- stVisitedModules :: Lens' TCState VisitedModules
- stWarningOnImport :: Lens' TCState (Maybe Text)
- stateTCLens :: (MonadTCState m, ReadTCState m) => Lens' TCState a -> (a -> (r, a)) -> m r
- stateTCLensM :: (MonadTCState m, ReadTCState m) => Lens' TCState a -> (a -> m (r, a)) -> m r
- suffixNameSuggestion :: MetaNameSuggestion -> ArgName -> MetaNameSuggestion
- syntaxError :: (HasCallStack, MonadTCError m) => String -> m a
- tcWarningOrigin :: TCWarning -> SrcFile
- thenReduce :: ReduceM a -> ReduceM b -> ReduceM b
- thenTCMT :: forall (m :: Type -> Type) a b. Applicative m => TCMT m a -> TCMT m b -> TCMT m b
- toExpandLast :: Bool -> ExpandHidden
- toReduceDefs :: (Bool, [QName]) -> ReduceDefs
- topLevelModuleFilePath :: ModuleToSource -> TopLevelModuleName -> AbsolutePath
- typeError :: (HasCallStack, MonadTCError m) => TypeError -> m a
- typeError' :: MonadTCError m => CallStack -> TypeError -> m a
- typeError'_ :: (MonadTCEnv m, ReadTCState m) => CallStack -> TypeError -> m TCErr
- typeError_ :: (HasCallStack, MonadTCEnv m, ReadTCState m) => TypeError -> m TCErr
- unKleisli :: (a -> ReduceM b) -> ReduceM (a -> b)
- unionBuiltin :: Builtin a -> Builtin a -> Builtin a
- unquoteError :: (HasCallStack, MonadTCError m) => UnquoteError -> m a
- unquoteNormalise :: Lens' UnquoteFlags Bool
- useR :: ReadTCState m => Lens' TCState a -> m a
- useTC :: ReadTCState m => Lens' TCState a -> m a
- viewTC :: MonadTCEnv m => Lens' TCEnv a -> m a
- warningName :: Warning -> WarningName
- withoutKOption :: HasOptions m => m Bool
- data AbstractMode
- data AllowedReduction
- type AllowedReductions = SmallSet AllowedReduction
- data ArgsCheckState a = ACState {
- acCheckedArgs :: [CheckedArg]
- acFun :: Expr
- acType :: Type
- acData :: a
- data AxiomData = AxiomData {}
- type Backend = Backend_boot Definition TCM
- type Backend' opts env menv mod def = Backend'_boot Definition TCM opts env menv mod def
- type BackendForeignCode = Map BackendName ForeignCodeStack
- newtype BlockT (m :: Type -> Type) a = BlockT {}
- data Builtin pf
- data BuiltinDescriptor
- = BuiltinData (TCM Type) [BuiltinId]
- | BuiltinDataCons (TCM Type)
- | BuiltinPrim PrimitiveId (Term -> TCM ())
- | BuiltinSort BuiltinSort
- | BuiltinPostulate Relevance (TCM Type)
- | BuiltinUnknown (Maybe (TCM Type)) (Term -> Type -> TCM ())
- data BuiltinInfo = BuiltinInfo {}
- data BuiltinSort
- type BuiltinThings = BuiltinThings' PrimFun
- type BuiltinThings' pf = Map SomeBuiltin (Builtin pf)
- type CachedTypeCheckLog = [(TypeCheckAction, PostScopeState)]
- data Call
- = CheckClause Type SpineClause
- | CheckLHS SpineLHS
- | CheckPattern Pattern Telescope Type
- | CheckPatternLinearityType Name
- | CheckPatternLinearityValue Name
- | CheckLetBinding LetBinding
- | InferExpr Expr
- | CheckExprCall Comparison Expr Type
- | CheckDotPattern Expr Term
- | CheckProjection Range QName Type
- | IsTypeCall Comparison Expr Sort
- | IsType_ Expr
- | InferVar Name
- | InferDef QName
- | CheckArguments Expr [NamedArg Expr] Type (Maybe Type)
- | CheckMetaSolution Range MetaId Type Term
- | CheckTargetType Range Type Type
- | CheckDataDef Range QName [LamBinding] [Constructor]
- | CheckRecDef Range QName [LamBinding] [Constructor]
- | CheckConstructor QName Telescope Sort Constructor
- | CheckConArgFitsIn QName Bool Type Sort
- | CheckFunDefCall Range QName [Clause] Bool
- | CheckPragma Range Pragma
- | CheckPrimitive Range QName Expr
- | CheckIsEmpty Range Type
- | CheckConfluence QName QName
- | CheckModuleParameters ModuleName Telescope
- | CheckWithFunctionType Type
- | CheckSectionApplication Range Erased ModuleName ModuleApplication
- | CheckNamedWhere ModuleName
- | CheckIApplyConfluence Range QName Term Term Term Type
- | ScopeCheckExpr Expr
- | ScopeCheckDeclaration NiceDeclaration
- | ScopeCheckLHS QName Pattern
- | NoHighlighting
- | ModuleContents
- | SetRange Range
- data CallInfo = CallInfo {}
- data Candidate = Candidate {}
- data CandidateKind
- data CannotQuote
- data CheckedArg = CheckedArg {
- caElim :: Elim
- caRange :: Maybe Range
- caConstraint :: Maybe (Abs Constraint)
- data CheckedTarget
- newtype CheckpointId = CheckpointId Int
- data Closure a = Closure {}
- data CompKit = CompKit {}
- data CompareAs
- data CompareDirection
- type CompiledRepresentation = Map BackendName [CompilerPragma]
- data CompilerPragma = CompilerPragma Range String
- type ConcreteNames = Map Name (List1 Name)
- data Constraint
- = ValueCmp Comparison CompareAs Term Term
- | ValueCmpOnFace Comparison Term Type Term Term
- | ElimCmp [Polarity] [IsForced] Type Term [Elim] [Elim]
- | SortCmp Comparison Sort Sort
- | LevelCmp Comparison Level Level
- | HasBiggerSort Sort
- | HasPTSRule (Dom Type) (Abs Sort)
- | CheckDataSort QName Sort
- | CheckMetaInst MetaId
- | CheckType Type
- | UnBlock MetaId
- | IsEmpty Range Type
- | CheckSizeLtSat Term
- | FindInstance MetaId (Maybe [Candidate])
- | ResolveInstanceHead QName
- | CheckFunDef DefInfo QName [Clause] TCErr
- | UnquoteTactic Term Term Type
- | CheckLockedVars Term Type (Arg Term) Type
- | UsableAtModality WhyCheckModality (Maybe Sort) Modality Term
- type Constraints = [ProblemConstraint]
- data ConstructorData = ConstructorData {
- _conPars :: Int
- _conArity :: Int
- _conSrcCon :: ConHead
- _conData :: QName
- _conAbstr :: IsAbstract
- _conComp :: CompKit
- _conProj :: Maybe [QName]
- _conForced :: [IsForced]
- _conErased :: Maybe [Bool]
- _conErasure :: !Bool
- _conInline :: !Bool
- data ConstructorDisambiguationData = ConstructorDisambiguationData {
- bcdConName :: QName
- bcdCandidates :: List1 (QName, Type, ConHead)
- bcdArguments :: Args
- bcdType :: Type
- type CurrentTypeCheckLog = [(TypeCheckAction, PostScopeState)]
- data DataOrRecSigData = DataOrRecSigData {
- _datarecPars :: Int
- type DataOrRecordE = DataOrRecord' InductionAndEta
- data DatatypeData = DatatypeData {
- _dataPars :: Nat
- _dataIxs :: Nat
- _dataClause :: Maybe Clause
- _dataCons :: [QName]
- _dataSort :: Sort
- _dataMutual :: Maybe [QName]
- _dataAbstr :: IsAbstract
- _dataPathCons :: [QName]
- _dataTranspIx :: Maybe QName
- _dataTransp :: Maybe QName
- type DecodedModules = Map TopLevelModuleName ModuleInfo
- data Definition = Defn {
- defArgInfo :: ArgInfo
- defName :: QName
- defType :: Type
- defPolarity :: [Polarity]
- defArgOccurrences :: [Occurrence]
- defGeneralizedParams :: [Maybe Name]
- defDisplay :: [LocalDisplayForm]
- defMutual :: MutualId
- defCompiledRep :: CompiledRepresentation
- defInstance :: Maybe InstanceInfo
- defCopy :: Bool
- defMatchable :: Set QName
- defNoCompilation :: Bool
- defInjective :: Bool
- defCopatternLHS :: Bool
- defBlocked :: Blocked_
- defLanguage :: !Language
- theDef :: Defn
- type Definitions = HashMap QName Definition
- data Defn
- data DisambiguatedName = DisambiguatedName NameKind QName
- type DisambiguatedNames = IntMap DisambiguatedName
- data DisplayForm = Display {
- dfPatternVars :: Nat
- dfPats :: Elims
- dfRHS :: DisplayTerm
- type DisplayForms = HashMap QName [LocalDisplayForm]
- data DisplayTerm
- data DoGeneralize
- data EtaEquality
- = Specified {
- theEtaEquality :: !HasEta
- | Inferred {
- theEtaEquality :: !HasEta
- = Specified {
- data ExecError
- data ExpandHidden
- data ExtLamInfo = ExtLamInfo {
- extLamModule :: ModuleName
- extLamAbsurd :: Bool
- extLamSys :: !(Maybe System)
- type Face = [(Term, Bool)]
- newtype Fields = Fields [(Name, Type)]
- data ForeignCode = ForeignCode Range String
- newtype ForeignCodeStack = ForeignCodeStack {}
- class FreshName a where
- freshName_ :: MonadFresh NameId m => a -> m Name
- data Frozen
- data FunctionData = FunctionData {
- _funClauses :: [Clause]
- _funCompiled :: Maybe CompiledClauses
- _funSplitTree :: Maybe SplitTree
- _funTreeless :: Maybe Compiled
- _funCovering :: [Clause]
- _funInv :: FunctionInverse
- _funMutual :: Maybe [QName]
- _funProjection :: Either ProjectionLikenessMissing Projection
- _funFlags :: SmallSet FunctionFlag
- _funTerminates :: Maybe Bool
- _funExtLam :: Maybe ExtLamInfo
- _funWith :: Maybe QName
- _funIsKanOp :: Maybe QName
- _funOpaque :: IsOpaque
- data FunctionFlag
- type FunctionInverse = FunctionInverse' Clause
- data FunctionInverse' c
- = NotInjective
- | Inverse (InversionMap c)
- data GHCBackendError
- data GeneralizedValue = GeneralizedValue {}
- class Enum i => HasFresh i where
- freshLens :: Lens' TCState i
- nextFresh' :: i -> i
- type IPBoundary = IPBoundary' Term
- newtype IPBoundary' t = IPBoundary {
- getBoundary :: Map (IntMap Bool) t
- data IPClause
- = IPClause {
- ipcQName :: QName
- ipcClauseNo :: Int
- ipcType :: Type
- ipcWithSub :: Maybe Substitution
- ipcClause :: SpineClause
- ipcClosure :: Closure ()
- | IPNoClause
- = IPClause {
- data IllegalRewriteRuleReason
- = LHSNotDefinitionOrConstructor
- | VariablesNotBoundByLHS IntSet
- | VariablesBoundMoreThanOnce IntSet
- | LHSReduces Term Term
- | HeadSymbolIsProjectionLikeFunction QName
- | HeadSymbolIsTypeConstructor QName
- | HeadSymbolContainsMetas QName
- | ConstructorParametersNotGeneral ConHead Args
- | ContainsUnsolvedMetaVariables (Set1 MetaId)
- | BlockedOnProblems (Set1 ProblemId)
- | RequiresDefinitions (Set1 QName)
- | DoesNotTargetRewriteRelation
- | BeforeFunctionDefinition
- | BeforeMutualFunctionDefinition QName
- | DuplicateRewriteRule
- type ImportedModules = Set TopLevelModuleName
- data IncorrectTypeForRewriteRelationReason
- data InductionAndEta = InductionAndEta {}
- data InstanceInfo = InstanceInfo {}
- data InstanceTable = InstanceTable {}
- data Instantiation = Instantiation {}
- data InteractionError
- type InteractionOutputCallback = Response_boot TCErr TCWarning WarningsAndNonFatalErrors -> TCM ()
- data InteractionPoint = InteractionPoint {}
- type InteractionPoints = BiMap InteractionId InteractionPoint
- data Interface = Interface {
- iSourceHash :: !Hash
- iSource :: Text
- iFileType :: FileType
- iImportedModules :: [(TopLevelModuleName, Hash)]
- iModuleName :: ModuleName
- iTopLevelModuleName :: TopLevelModuleName
- iScope :: Map ModuleName Scope
- iInsideScope :: ScopeInfo
- iSignature :: Signature
- iMetaBindings :: RemoteMetaStore
- iDisplayForms :: DisplayForms
- iUserWarnings :: UserWarnings
- iImportWarning :: Maybe Text
- iBuiltin :: BuiltinThings' (PrimitiveId, QName)
- iForeignCode :: Map BackendName ForeignCodeStack
- iHighlighting :: HighlightingInfo
- iDefaultPragmaOptions :: [OptionsPragma]
- iFilePragmaOptions :: [OptionsPragma]
- iOptionsUsed :: PragmaOptions
- iPatternSyns :: PatternSynDefns
- iWarnings :: Set TCWarning
- iPartialDefs :: Set QName
- iOpaqueBlocks :: Map OpaqueId OpaqueBlock
- iOpaqueNames :: Map QName OpaqueId
- data InvalidFileNameReason
- type InversionMap c = Map TermHead [c]
- data IsAmbiguous
- data IsForced
- data IsReduced
- = NotReduced
- | Reduced (Blocked ())
- data JSBackendError = BadCompilePragma
- data Judgement a
- = HasType {
- jMetaId :: a
- jComparison :: Comparison
- jMetaType :: Type
- | IsSort { }
- = HasType {
- data LHSOrPatSyn
- class LensClosure b a | b -> a where
- lensClosure :: Lens' b (Closure a)
- class LensTCEnv a where
- data LetBinding = LetBinding {}
- type LetBindings = Map Name (Open LetBinding)
- data Listener
- data LoadedFileCache = LoadedFileCache {}
- type LocalDisplayForm = Open DisplayForm
- type LocalMetaStore = Map MetaId MetaVariable
- data MaybeReduced a = MaybeRed {
- isReduced :: IsReduced
- ignoreReduced :: a
- type MaybeReducedArgs = [MaybeReduced (Arg Term)]
- type MaybeReducedElims = [MaybeReduced Elim]
- data MetaInfo = MetaInfo {}
- data MetaInstantiation
- newtype MetaPriority = MetaPriority Int
- data MetaVariable = MetaVar {}
- data MissingTypeSignatureInfo
- data ModuleCheckMode
- data ModuleInfo = ModuleInfo {}
- class Monad m => MonadBlock (m :: Type -> Type) where
- patternViolation :: Blocker -> m a
- catchPatternErr :: (Blocker -> m a) -> m a -> m a
- class Monad m => MonadFresh i (m :: Type -> Type) where
- fresh :: m i
- class (Applicative m, MonadTCEnv m, ReadTCState m, HasOptions m) => MonadReduce (m :: Type -> Type) where
- liftReduce :: ReduceM a -> m a
- class Monad m => MonadStConcreteNames (m :: Type -> Type) where
- runStConcreteNames :: StateT ConcreteNames m a -> m a
- useConcreteNames :: m ConcreteNames
- modifyConcreteNames :: (ConcreteNames -> ConcreteNames) -> m ()
- class Monad m => MonadTCEnv (m :: Type -> Type) where
- type MonadTCError (m :: Type -> Type) = (MonadTCEnv m, ReadTCState m, MonadError TCErr m)
- class (Applicative tcm, MonadIO tcm, MonadTCEnv tcm, MonadTCState tcm, HasOptions tcm) => MonadTCM (tcm :: Type -> Type) where
- class Monad m => MonadTCState (m :: Type -> Type) where
- data MutualBlock = MutualBlock {}
- type MutualBlocks = IntMap MutualBlock
- newtype MutualId = MutualId Word32
- data NLPSort
- data NLPType = NLPType {}
- data NLPat
- data NegativeUnification
- data NumGeneralizableArgs
- data OpaqueBlock = OpaqueBlock {}
- data Open a = OpenThing {}
- data Overapplied
- type PElims = [Elim' NLPat]
- data PersistentTCState = PersistentTCSt {
- stPersistentSession :: !SessionTCState
- stDecodedModules :: !DecodedModules
- stPersistentTopLevelModuleNames :: !(BiMap RawTopLevelModuleName ModuleNameHash)
- stPersistentOptions :: CommandLineOptions
- stInteractionOutputCallback :: InteractionOutputCallback
- stAccumStatistics :: !Statistics
- stPersistLoadedFileCache :: !(Maybe LoadedFileCache)
- data PostScopeState = PostScopeState {
- stPostSyntaxInfo :: !HighlightingInfo
- stPostDisambiguatedNames :: !DisambiguatedNames
- stPostOpenMetaStore :: !LocalMetaStore
- stPostSolvedMetaStore :: !LocalMetaStore
- stPostInteractionPoints :: !InteractionPoints
- stPostAwakeConstraints :: !Constraints
- stPostSleepingConstraints :: !Constraints
- stPostDirty :: !Bool
- stPostOccursCheckDefs :: !(Set QName)
- stPostSignature :: !Signature
- stPostModuleCheckpoints :: !(Map ModuleName CheckpointId)
- stPostImportsDisplayForms :: !DisplayForms
- stPostForeignCode :: !BackendForeignCode
- stPostCurrentModule :: !(Maybe (ModuleName, TopLevelModuleName))
- stPostPendingInstances :: !(Set QName)
- stPostTemporaryInstances :: !(Set QName)
- stPostConcreteNames :: !ConcreteNames
- stPostUsedNames :: !UsedNames
- stPostShadowingNames :: !ShadowingNames
- stPostStatistics :: !Statistics
- stPostTCWarnings :: !(Set TCWarning)
- stPostMutualBlocks :: !MutualBlocks
- stPostLocalBuiltins :: !BuiltinThings
- stPostFreshMetaId :: !MetaId
- stPostFreshMutualId :: !MutualId
- stPostFreshProblemId :: !ProblemId
- stPostFreshCheckpointId :: !CheckpointId
- stPostFreshInt :: !Int
- stPostFreshNameId :: !NameId
- stPostFreshOpaqueId :: !OpaqueId
- stPostAreWeCaching :: !Bool
- stPostPostponeInstanceSearch :: !Bool
- stPostConsideringInstance :: !Bool
- stPostInstantiateBlocking :: !Bool
- stPostLocalPartialDefs :: !(Set QName)
- stPostOpaqueBlocks :: Map OpaqueId OpaqueBlock
- stPostOpaqueIds :: Map QName OpaqueId
- stPostInstanceHack :: !Bool
- data PreScopeState = PreScopeState {
- stPreTokens :: !HighlightingInfo
- stPreImports :: !Signature
- stPreImportedModules :: !ImportedModules
- stPreModuleToSourceId :: !ModuleToSourceId
- stPreVisitedModules :: !VisitedModules
- stPreScope :: !ScopeInfo
- stPrePatternSyns :: !PatternSynDefns
- stPrePatternSynImports :: !PatternSynDefns
- stPreGeneralizedVars :: !(Maybe (Set QName))
- stPrePragmaOptions :: !PragmaOptions
- stPreImportedBuiltins :: !BuiltinThings
- stPreImportedDisplayForms :: !DisplayForms
- stPreFreshInteractionId :: !InteractionId
- stPreImportedUserWarnings :: !UserWarnings
- stPreLocalUserWarnings :: !UserWarnings
- stPreWarningOnImport :: !(Maybe Text)
- stPreImportedPartialDefs :: !(Set QName)
- stPreLibCache :: !LibCache
- stPreImportedMetaStore :: !RemoteMetaStore
- stPreCopiedNames :: !(HashMap QName QName)
- stPreNameCopies :: !(HashMap QName (HashSet QName))
- data PrimFun = PrimFun {
- primFunName :: QName
- primFunArity :: Arity
- primFunArgOccurrences :: [Occurrence]
- primFunImplementation :: [Arg Term] -> Int -> ReduceM (Reduced MaybeReducedArgs Term)
- data PrimitiveData = PrimitiveData {}
- data PrimitiveImpl = PrimImpl Type PrimFun
- data PrimitiveSortData = PrimitiveSortData {}
- data PrincipalArgTypeMetas = PrincipalArgTypeMetas {
- patmMetas :: Args
- patmRemainder :: Type
- data ProblemConstraint = PConstr {}
- newtype ProjLams = ProjLams {
- getProjLams :: [Arg ArgName]
- data Projection = Projection {}
- data ProjectionLikenessMissing
- class Monad m => ReadTCState (m :: Type -> Type) where
- getTCState :: m TCState
- locallyTCState :: Lens' TCState a -> (a -> a) -> m b -> m b
- withTCState :: (TCState -> TCState) -> m a -> m a
- data RecordData = RecordData {
- _recPars :: Nat
- _recClause :: Maybe Clause
- _recConHead :: ConHead
- _recNamedCon :: Bool
- _recFields :: [Dom QName]
- _recTel :: Telescope
- _recMutual :: Maybe [QName]
- _recEtaEquality' :: EtaEquality
- _recPatternMatching :: PatternOrCopattern
- _recInduction :: Maybe Induction
- _recTerminates :: Maybe Bool
- _recAbstr :: IsAbstract
- _recComp :: CompKit
- data ReduceDefs
- = OnlyReduceDefs (Set QName)
- | DontReduceDefs (Set QName)
- data ReduceEnv = ReduceEnv {}
- newtype ReduceM a = ReduceM {}
- data Reduced no yes
- = NoReduction no
- | YesReduction Simplification yes
- type RemoteMetaStore = HashMap MetaId RemoteMetaVariable
- data RemoteMetaVariable = RemoteMetaVariable {}
- data RewriteRule = RewriteRule {}
- type RewriteRuleMap = HashMap QName RewriteRules
- type RewriteRules = [RewriteRule]
- data RunMetaOccursCheck
- newtype Section = Section {}
- type Sections = Map ModuleName Section
- data SessionTCState = SessionTCState {}
- type ShadowingNames = Map Name (Set1 RawName)
- data Signature = Sig {}
- data Simplification
- data SplitError
- = NotADatatype (Closure Type)
- | BlockedType Blocker (Closure Type)
- | ErasedDatatype ErasedDatatypeReason (Closure Type)
- | CoinductiveDatatype (Closure Type)
- | UnificationStuck { }
- | CosplitCatchall
- | CosplitNoTarget
- | CosplitNoRecordType (Closure Type)
- | CannotCreateMissingClause QName (Telescope, [NamedArg DeBruijnPattern]) Doc (Closure (Abs Type))
- | GenericSplitError String
- type Statistics = Map String Integer
- data System = System {
- systemTel :: Telescope
- systemClauses :: [(Face, Term)]
- data TCEnv = TCEnv {
- envContext :: Context
- envLetBindings :: LetBindings
- envCurrentModule :: ModuleName
- envCurrentPath :: Maybe FileId
- envAnonymousModules :: [(ModuleName, Nat)]
- envImportPath :: [TopLevelModuleName]
- envMutualBlock :: Maybe MutualId
- envTerminationCheck :: TerminationCheck ()
- envCoverageCheck :: CoverageCheck
- envMakeCase :: Bool
- envSolvingConstraints :: Bool
- envCheckingWhere :: WhereClause_
- envWorkingOnTypes :: Bool
- envAssignMetas :: Bool
- envActiveProblems :: Set ProblemId
- envUnquoteProblem :: Maybe ProblemId
- envAbstractMode :: AbstractMode
- envRelevance :: Relevance
- envQuantity :: Quantity
- envHardCompileTimeMode :: Bool
- envSplitOnStrict :: Bool
- envDisplayFormsEnabled :: Bool
- envFoldLetBindings :: Bool
- envRange :: Range
- envHighlightingRange :: Range
- envClause :: IPClause
- envCall :: Maybe (Closure Call)
- envHighlightingLevel :: HighlightingLevel
- envHighlightingMethod :: HighlightingMethod
- envExpandLast :: ExpandHidden
- envAppDef :: Maybe QName
- envSimplification :: Simplification
- envAllowedReductions :: AllowedReductions
- envReduceDefs :: ReduceDefs
- envReconstructed :: Bool
- envInjectivityDepth :: Int
- envCompareBlocked :: Bool
- envPrintDomainFreePi :: Bool
- envPrintMetasBare :: Bool
- envInsideDotPattern :: Bool
- envUnquoteFlags :: UnquoteFlags
- envInstanceDepth :: !Int
- envIsDebugPrinting :: Bool
- envPrintingPatternLambdas :: [QName]
- envCallByNeed :: Bool
- envCurrentCheckpoint :: CheckpointId
- envCheckpoints :: Map CheckpointId Substitution
- envGeneralizeMetas :: DoGeneralize
- envGeneralizedVars :: Map QName GeneralizedValue
- envActiveBackendName :: Maybe BackendName
- envConflComputingOverlap :: Bool
- envCurrentlyElaborating :: Bool
- envSyntacticEqualityFuel :: !(Maybe Int)
- envCurrentOpaqueId :: !(Maybe OpaqueId)
- envTermCheckReducing :: Bool
- data TCErr
- type TCM = TCMT IO
- newtype TCMT (m :: Type -> Type) a = TCM {}
- data TCState = TCSt {}
- data TCWarning = TCWarning {}
- type TempInstanceTable = (InstanceTable, Set QName)
- data TermHead
- data TerminationError = TerminationError {}
- data TypeCheckAction
- data TypeCheckingProblem
- = CheckExpr Comparison Expr Type
- | CheckArgs Comparison ExpandHidden Expr [NamedArg Expr] Type Type (ArgsCheckState CheckedTarget -> TCM Term)
- | CheckProjAppToKnownPrincipalArg Comparison Expr ProjOrigin (List1 QName) Expr Args Type Int Term Type PrincipalArgTypeMetas
- | CheckLambda Comparison (Arg (List1 (WithHiding Name), Maybe Type)) Expr Type
- | DisambiguateConstructor ConstructorDisambiguationData (ConHead -> TCM Term)
- | DoQuoteTerm Comparison Term Type
- data TypeError
- = InternalError String
- | NotImplemented String
- | NotSupported String
- | CompilationError String
- | SyntaxError String
- | OptionError OptionError
- | NicifierError DeclarationException'
- | DoNotationError String
- | IdiomBracketError String
- | NoKnownRecordWithSuchFields [Name]
- | ShouldEndInApplicationOfTheDatatype Type
- | ConstructorPatternInWrongDatatype QName QName
- | CantResolveOverloadedConstructorsTargetingSameDatatype QName (List1 QName)
- | ConstructorDoesNotTargetGivenType QName Type
- | InvalidDottedExpression
- | LiteralTooBig
- | NegativeLiteralInPattern
- | WrongHidingInLHS
- | WrongHidingInLambda Type
- | WrongHidingInApplication Type
- | WrongHidingInProjection QName
- | IllegalHidingInPostfixProjection (NamedArg Expr)
- | WrongNamedArgument (NamedArg Expr) (List1 NamedName)
- | WrongAnnotationInLambda
- | WrongIrrelevanceInLambda
- | WrongQuantityInLambda
- | WrongCohesionInLambda
- | WrongPolarityInLambda
- | QuantityMismatch Quantity Quantity
- | HidingMismatch Hiding Hiding
- | RelevanceMismatch Relevance Relevance
- | ForcedConstructorNotInstantiated Pattern
- | IllformedProjectionPatternAbstract Pattern
- | IllformedProjectionPatternConcrete Pattern
- | CannotEliminateWithPattern (Maybe Blocker) (NamedArg Pattern) Type
- | CannotEliminateWithProjection (Arg Type) Bool QName
- | WrongNumberOfConstructorArguments QName Nat Nat
- | ShouldBeEmpty Type [DeBruijnPattern]
- | ShouldBeASort Type
- | ShouldBePi Type
- | ShouldBePath Type
- | ShouldBeRecordType Type
- | ShouldBeRecordPattern DeBruijnPattern
- | CannotApply Expr Type
- | InvalidTypeSort Sort
- | SplitOnCoinductive
- | SplitOnIrrelevant (Dom Type)
- | SplitOnUnusableCohesion (Dom Type)
- | SplitOnUnusablePolarity (Dom Type)
- | SplitOnNonVariable Term Type
- | SplitOnNonEtaRecord QName
- | SplitOnAbstract QName
- | SplitOnUnchecked QName
- | SplitOnPartial (Dom Type)
- | SplitInProp DataOrRecordE
- | DefinitionIsIrrelevant QName
- | DefinitionIsErased QName
- | ProjectionIsIrrelevant QName
- | VariableIsIrrelevant Name
- | VariableIsErased Name
- | VariableIsOfUnusableCohesion Name Cohesion
- | LambdaIsErased
- | RecordIsErased
- | InvalidModalTelescopeUse Term Modality Modality Definition
- | VariableIsOfUnusablePolarity Name PolarityModality
- | UnequalLevel Comparison Level Level
- | UnequalTerms Comparison Term Term CompareAs
- | UnequalRelevance Comparison Term Term
- | UnequalQuantity Comparison Term Term
- | UnequalCohesion Comparison Term Term
- | UnequalPolarity Comparison Term Term
- | UnequalFiniteness Comparison Term Term
- | UnequalHiding Term Term
- | UnequalSorts Sort Sort
- | NotLeqSort Sort Sort
- | MetaCannotDependOn MetaId Term Nat
- | MetaIrrelevantSolution MetaId Term
- | MetaErasedSolution MetaId Term
- | GenericError String
- | GenericDocError Doc
- | SortOfSplitVarError (Maybe Blocker) Doc
- | WrongSharpArity QName
- | BuiltinMustBeConstructor BuiltinId Expr
- | BuiltinMustBeData BuiltinId Int
- | BuiltinMustBeDef BuiltinId
- | BuiltinMustBeFunction BuiltinId
- | BuiltinMustBePostulate BuiltinId
- | NoSuchBuiltinName String
- | InvalidBuiltin String
- | DuplicateBuiltinBinding BuiltinId Term Term
- | NoBindingForBuiltin BuiltinId
- | NoBindingForPrimitive PrimitiveId
- | NoSuchPrimitiveFunction String
- | DuplicatePrimitiveBinding PrimitiveId QName QName
- | WrongArgInfoForPrimitive PrimitiveId ArgInfo ArgInfo
- | ShadowedModule Name (List1 ModuleName)
- | BuiltinInParameterisedModule BuiltinId
- | IllegalDeclarationInDataDefinition (List1 Declaration)
- | IllegalLetInTelescope TypedBinding
- | IllegalPatternInTelescope Binder
- | AbsentRHSRequiresAbsurdPattern
- | TooManyFields QName [Name] (List1 Name)
- | DuplicateFields (List1 Name)
- | DuplicateConstructors (List1 Name)
- | DuplicateOverlapPragma QName OverlapMode OverlapMode
- | WithOnFreeVariable Expr Term
- | UnexpectedWithPatterns (List1 Pattern)
- | WithClausePatternMismatch Pattern (NamedArg DeBruijnPattern)
- | IllTypedPatternAfterWithAbstraction Pattern
- | TooFewPatternsInWithClause
- | TooManyPatternsInWithClause
- | PathAbstractionFailed (Abs Type)
- | FieldOutsideRecord
- | ModuleArityMismatch ModuleName Telescope (Either (List1 (NamedArg Expr)) Args)
- | GeneralizeCyclicDependency
- | ReferencesFutureVariables Term (List1 Int) (Arg Term) Int
- | DoesNotMentionTicks Term Type (Arg Term)
- | MismatchedProjectionsError QName QName
- | AttributeKindNotEnabled String String String
- | InvalidProjectionParameter (NamedArg Expr)
- | TacticAttributeNotAllowed
- | CannotRewriteByNonEquation Type
- | MacroResultTypeMismatch Type
- | NamedWhereModuleInRefinedContext [Term] [String]
- | ComatchingDisabledForRecord QName
- | IncorrectTypeForRewriteRelation Term IncorrectTypeForRewriteRelationReason
- | CannotGenerateHCompClause Type
- | CannotGenerateTransportClause QName (Closure (Abs Type))
- | CubicalNotErasure QName
- | CubicalPrimitiveNotFullyApplied QName
- | ExpectedIntervalLiteral Expr
- | FaceConstraintDisjunction
- | FaceConstraintUnsatisfiable
- | PatternInPathLambda
- | PatternInSystem
- | UnexpectedParameter LamBinding
- | NoParameterOfName ArgName
- | UnexpectedModalityAnnotationInParameter LamBinding
- | ExpectedBindingForParameter (Dom Type) (Abs Type)
- | UnexpectedTypeSignatureForParameter (List1 (NamedArg Binder))
- | SortDoesNotAdmitDataDefinitions QName Sort
- | SortCannotDependOnItsIndex QName Type
- | UnusableAtModality WhyCheckModality Modality Term
- | SplitError SplitError
- | ImpossibleConstructor QName NegativeUnification
- | DatatypeIndexPolarity
- | RecursiveRecordNeedsInductivity QName
- | CannotSolveSizeConstraints (List1 (ProblemConstraint, HypSizeConstraint)) Doc
- | ContradictorySizeConstraint (ProblemConstraint, HypSizeConstraint)
- | EmptyTypeOfSizes Term
- | FunctionTypeInSizeUniv Term
- | PostulatedSizeInModule
- | LibraryError LibErrors
- | LibTooFarDown TopLevelModuleName AgdaLibFile
- | SolvedButOpenHoles
- | CyclicModuleDependency (List2 TopLevelModuleName)
- | FileNotFound TopLevelModuleName [AbsolutePath]
- | OverlappingProjects AbsolutePath TopLevelModuleName TopLevelModuleName
- | AmbiguousTopLevelModuleName TopLevelModuleName (List2 AbsolutePath)
- | ModuleNameUnexpected TopLevelModuleName TopLevelModuleName
- | ModuleNameDoesntMatchFileName TopLevelModuleName [AbsolutePath]
- | ModuleDefinedInOtherFile TopLevelModuleName AbsolutePath AbsolutePath
- | InvalidFileName AbsolutePath InvalidFileNameReason
- | ModuleNameHashCollision RawTopLevelModuleName (Maybe RawTopLevelModuleName)
- | BothWithAndRHS
- | AbstractConstructorNotInScope QName
- | CopatternHeadNotProjection QName
- | NotAllowedInDotPatterns NotAllowedInDotPatterns
- | NotInScope QName
- | NoSuchModule QName
- | AmbiguousName QName AmbiguousNameReason
- | AmbiguousModule QName (List1 ModuleName)
- | AmbiguousField Name (List2 ModuleName)
- | AmbiguousConstructor QName (List2 QName)
- | ClashingDefinition QName QName (Maybe NiceDeclaration)
- | ClashingModule ModuleName ModuleName
- | DefinitionInDifferentModule QName
- | DuplicateImports QName (List1 ImportedName)
- | InvalidPattern Pattern
- | InvalidPun ConstructorOrPatternSynonym QName
- | RepeatedNamesInImportDirective (List1 (List2 ImportedName))
- | RepeatedVariablesInPattern (List1 Name)
- | GeneralizeNotSupportedHere QName
- | GeneralizedVarInLetOpenedModule QName
- | MultipleFixityDecls (List1 (Name, Pair Fixity'))
- | MultiplePolarityPragmas (List1 Name)
- | ExplicitPolarityVsPragma QName
- | ConstructorNameOfNonRecord ResolvedName
- | CannotQuote CannotQuote
- | CannotQuoteTerm CannotQuoteTerm
- | DeclarationsAfterTopLevelModule
- | IllegalDeclarationBeforeTopLevelModule
- | MissingTypeSignature MissingTypeSignatureInfo
- | NotAnExpression Expr
- | NotAValidLetBinding (Maybe NotAValidLetBinding)
- | NotAValidLetExpression NotAValidLetExpression
- | NotValidBeforeField NiceDeclaration
- | OpenEverythingInRecordWhere
- | PrivateRecordField
- | QualifiedLocalModule
- | AsPatternInPatternSynonym
- | DotPatternInPatternSynonym
- | BadArgumentsToPatternSynonym AmbiguousQName
- | TooFewArgumentsToPatternSynonym AmbiguousQName
- | CannotResolveAmbiguousPatternSynonym (List1 (QName, PatternSynDefn))
- | IllegalInstanceVariableInPatternSynonym Name
- | PatternSynonymArgumentShadows ConstructorOrPatternSynonym Name (List1 AbstractName)
- | UnusedVariableInPatternSynonym Name
- | UnboundVariablesInPatternSynonym (List1 Name)
- | NoParseForApplication (List2 Expr)
- | AmbiguousParseForApplication (List2 Expr) (List1 Expr)
- | NoParseForLHS LHSOrPatSyn [Pattern] Pattern
- | AmbiguousParseForLHS LHSOrPatSyn Pattern (List2 Pattern)
- | AmbiguousProjection QName (List1 QName)
- | AmbiguousOverloadedProjection (List1 QName) Doc
- | OperatorInformation [NotationSection] TypeError
- | InstanceNoCandidate Type [(Term, TCErr)]
- | ExecError ExecError
- | UnquoteFailed UnquoteError
- | DeBruijnIndexOutOfScope Nat Telescope [Name]
- | NeedOptionAllowExec
- | NeedOptionCopatterns
- | NeedOptionCubical Cubical String
- | NeedOptionPatternMatching
- | NeedOptionProp
- | NeedOptionRewriting
- | NeedOptionSizedTypes String
- | NeedOptionTwoLevel
- | NeedOptionUniversePolymorphism
- | NonFatalErrors (Set1 TCWarning)
- | InstanceSearchDepthExhausted Term Type Int
- | TriedToCopyConstrainedPrim QName
- | InvalidInstanceHeadType Type WhyInvalidInstanceType
- | InteractionError InteractionError
- | BackendDoesNotSupportOnlyScopeChecking BackendName
- | CubicalCompilationNotSupported Cubical
- | CustomBackendError BackendName Doc
- | GHCBackendError GHCBackendError
- | JSBackendError JSBackendError
- | UnknownBackend BackendName (Set BackendName)
- data UnificationFailure
- data UnquoteError
- = BlockedOnMeta TCState Blocker
- | CannotDeclareHiddenFunction QName
- | CommitAfterDef
- | ConInsteadOfDef QName String String
- | DefineDataNotData QName
- | DefInsteadOfCon QName String String
- | MissingDeclaration QName
- | MissingDefinition QName
- | NakedUnquote
- | NonCanonical String Term
- | PatLamWithoutClauses Term
- | StaleMeta TopLevelModuleName MetaId
- | TooManyParameters Nat Expr
- | UnboundName QName
- data UnquoteFlags = UnquoteFlags {}
- type UsedNames = Map RawName (Set1 RawName)
- type UserWarnings = Map QName Text
- type VisitedModules = Map TopLevelModuleName ModuleInfo
- data Warning
- = NicifierIssue DeclarationWarning
- | TerminationIssue (List1 TerminationError)
- | UnreachableClauses QName (List1 Range)
- | CoverageIssue QName (List1 (Telescope, [NamedArg DeBruijnPattern]))
- | CoverageNoExactSplit QName (List1 Clause)
- | InlineNoExactSplit QName Clause
- | NotStrictlyPositive QName (Seq OccursWhere)
- | ConstructorDoesNotFitInData DataOrRecord_ QName Sort Sort TCErr
- | CoinductiveEtaRecord QName
- | UnsolvedMetaVariables (Set1 Range)
- | UnsolvedInteractionMetas (Set1 Range)
- | UnsolvedConstraints (List1 ProblemConstraint)
- | InteractionMetaBoundaries (Set1 Range)
- | CantGeneralizeOverSorts (Set1 MetaId)
- | AbsurdPatternRequiresAbsentRHS
- | OldBuiltin BuiltinId BuiltinId
- | BuiltinDeclaresIdentifier BuiltinId
- | DuplicateRecordDirective RecordDirective
- | EmptyRewritePragma
- | EmptyWhere
- | FixingRelevance String Relevance Relevance
- | FixingCohesion String Cohesion Cohesion
- | FixingPolarity String PolarityModality PolarityModality
- | IllformedAsClause String
- | InvalidCharacterLiteral Char
- | ClashesViaRenaming NameOrModule (Set1 Name)
- | UselessPatternDeclarationForRecord String
- | UselessPragma Range Doc
- | UselessPublic UselessPublicReason
- | UselessHiding (List1 ImportedName)
- | UselessInline QName
- | UselessTactic
- | WrongInstanceDeclaration
- | InstanceWithExplicitArg QName
- | InstanceNoOutputTypeName Doc
- | InstanceArgWithExplicitArg Doc
- | InversionDepthReached QName
- | SafeFlagPostulate QName
- | SafeFlagPragma (Set String)
- | SafeFlagWithoutKFlagPrimEraseEquality
- | WithoutKFlagPrimEraseEquality
- | ConflictingPragmaOptions String String
- | OptionWarning OptionWarning
- | ParseWarning ParseWarning
- | LibraryWarning LibWarning
- | DeprecationWarning String String String
- | UserWarning Text
- | DuplicateUsing (List1 ImportedName)
- | FixityInRenamingModule (List1 Range)
- | ModuleDoesntExport QName [Name] [Name] (List1 ImportedName)
- | InfectiveImport Doc
- | CoInfectiveImport Doc
- | ConfluenceCheckingIncompleteBecauseOfMeta QName
- | ConfluenceForCubicalNotSupported
- | NotARewriteRule QName IsAmbiguous
- | IllegalRewriteRule QName IllegalRewriteRuleReason
- | RewriteNonConfluent Term Term Term Doc
- | RewriteMaybeNonConfluent Term Term [Doc]
- | RewriteAmbiguousRules Term Term Term
- | RewriteMissingRule Term Term Term
- | PragmaCompileErased BackendName QName
- | PragmaCompileList
- | PragmaCompileMaybe
- | PragmaCompileUnparsable String
- | PragmaCompileWrong QName String
- | PragmaCompileWrongName QName IsAmbiguous
- | PragmaExpectsDefinedSymbol String QName
- | PragmaExpectsUnambiguousConstructorOrFunction String QName IsAmbiguous
- | PragmaExpectsUnambiguousProjectionOrFunction String QName IsAmbiguous
- | NoMain TopLevelModuleName
- | NotInScopeW QName
- | UnsupportedIndexedMatch Doc
- | AsPatternShadowsConstructorOrPatternSynonym ConstructorOrPatternSynonym
- | PatternShadowsConstructor Name QName
- | PlentyInHardCompileTimeMode QωOrigin
- | RecordFieldWarning RecordFieldWarning
- | MissingTypeSignatureForOpaque QName IsOpaque
- | NotAffectedByOpaque
- | UnfoldingWrongName QName
- | UnfoldTransparentName QName
- | UselessOpaque
- | HiddenNotInArgumentPosition Expr
- | InstanceNotInArgumentPosition Expr
- | MacroInLetBindings
- | AbstractInLetBindings
- | InvalidDisplayForm QName String
- | UnusedVariablesInDisplayForm (List1 Name)
- | TooManyArgumentsToSort QName (List1 (NamedArg Expr))
- | WithClauseProjectionFixityMismatch { }
- | TooManyPolarities QName PragmaPolarities
- | TopLevelPolarity QName PolarityModality
- | FaceConstraintCannotBeHidden ArgInfo
- | FaceConstraintCannotBeNamed NamedName
- | CustomBackendWarning String Doc
- data WarningsAndNonFatalErrors = WarningsAndNonFatalErrors {}
- data WhyCheckModality
- data WhyInvalidInstanceType
- data WhyNotAHaskellType
- module Agda.TypeChecking.Monad.Base.Types
- data CannotQuoteTerm
- data ErasedDatatypeReason
- data NotAValidLetBinding
- data NotAValidLetExpression = MissingBody
- data NotAllowedInDotPatterns
- data FileId
- class Monad m => MonadFileId (m :: Type -> Type) where
- fileFromId :: FileId -> m File
- idFromFile :: File -> m FileId
- class (Functor m, Applicative m, Monad m) => HasOptions (m :: Type -> Type) where
- data RecordFieldWarning
- data UselessPublicReason
Documentation
pattern Constructor :: Int -> Int -> ConHead -> QName -> IsAbstract -> CompKit -> Maybe [QName] -> [IsForced] -> Maybe [Bool] -> Bool -> Bool -> Defn Source #
pattern DDot :: Term -> DisplayTerm Source #
pattern DTerm :: Term -> DisplayTerm Source #
pattern DataOrRecSig :: Int -> Defn Source #
pattern Datatype :: Nat -> Nat -> Maybe Clause -> [QName] -> Sort -> Maybe [QName] -> IsAbstract -> [QName] -> Maybe QName -> Maybe QName -> Defn Source #
pattern Function :: [Clause] -> Maybe CompiledClauses -> Maybe SplitTree -> Maybe Compiled -> [Clause] -> FunctionInverse -> Maybe [QName] -> Either ProjectionLikenessMissing Projection -> SmallSet FunctionFlag -> Maybe Bool -> Maybe ExtLamInfo -> Maybe QName -> Maybe QName -> IsOpaque -> Defn Source #
pattern Primitive :: IsAbstract -> PrimitiveId -> [Clause] -> FunctionInverse -> Maybe CompiledClauses -> IsOpaque -> Defn Source #
pattern PrimitiveSort :: BuiltinSort -> Sort -> Defn Source #
pattern Record :: Nat -> Maybe Clause -> ConHead -> Bool -> [Dom QName] -> Telescope -> Maybe [QName] -> EtaEquality -> PatternOrCopattern -> Maybe Induction -> Maybe Bool -> IsAbstract -> CompKit -> Defn Source #
pattern SortProp :: BuiltinSort Source #
pattern SortPropOmega :: BuiltinSort Source #
pattern SortSet :: BuiltinSort Source #
pattern SortSetOmega :: BuiltinSort Source #
pattern SortStrictSet :: BuiltinSort Source #
pattern SortStrictSetOmega :: BuiltinSort Source #
_recEtaEquality :: RecordData -> HasEta Source #
aDefToMode :: IsAbstract -> AbstractMode Source #
aModeToDef :: AbstractMode -> Maybe IsAbstract Source #
absurdLambdaName :: String Source #
Name of absurdLambda definitions.
allReductions :: AllowedReductions Source #
Not quite all reductions (skip non-terminating reductions)
apTCMT :: forall (m :: Type -> Type) a b. Applicative m => TCMT m (a -> b) -> TCMT m a -> TCMT m b Source #
asksTC :: MonadTCEnv m => (TCEnv -> a) -> m a Source #
axiomConstTransp :: Defn -> Bool Source #
bindTCMT :: forall (m :: Type -> Type) a b. Monad m => TCMT m a -> (a -> TCMT m b) -> TCMT m b Source #
buildClosure :: (MonadTCEnv m, ReadTCState m) => a -> m (Closure a) Source #
catchError_ :: TCM a -> (TCErr -> TCM a) -> TCM a Source #
Preserve the state of the failing computation.
conAbstr :: Defn -> IsAbstract Source #
conErasure :: Defn -> Bool Source #
cubicalCompatibleOption :: HasOptions m => m Bool Source #
cubicalOption :: HasOptions m => m (Maybe Cubical) Source #
currentModality :: MonadTCEnv m => m Modality Source #
The current modality.
Note that the returned cohesion component is always unitCohesion
.
dataAbstr :: Defn -> IsAbstract Source #
dataPathCons :: Defn -> [QName] Source #
datarecPars :: Defn -> Int Source #
defAbstract :: Definition -> IsAbstract Source #
defClauses :: Definition -> [Clause] Source #
defCompilerPragmas :: BackendName -> Definition -> [CompilerPragma] Source #
defConstructors :: Defn -> [QName] Source #
defForced :: Definition -> [IsForced] Source #
defIsDataOrRecord :: Defn -> Bool Source #
defIsRecord :: Defn -> Bool Source #
defNonterminating :: Definition -> Bool Source #
Has the definition failed the termination checker?
defOpaque :: Definition -> IsOpaque Source #
defParameters :: Definition -> Maybe Nat Source #
defTerminationUnconfirmed :: Definition -> Bool Source #
Has the definition not termination checked or did the check fail?
defaultAxiom :: Defn Source #
defaultDefn :: ArgInfo -> QName -> Type -> Language -> Defn -> Definition Source #
Create a definition with sensible defaults.
defaultDisplayForm :: QName -> [LocalDisplayForm] Source #
By default, we have no display form.
defaultInteractionOutputCallback :: InteractionOutputCallback Source #
The default InteractionOutputCallback
function prints certain
things to stdout (other things generate internal errors).
dirToCmp :: (Comparison -> a -> a -> c) -> CompareDirection -> a -> a -> c Source #
Turn a Comparison
function into a CompareDirection
function.
Property: dirToCmp f (fromCmp cmp) = f cmp
eAnonymousModules :: Lens' TCEnv [(ModuleName, Nat)] Source #
eQuantity :: Lens' TCEnv Quantity Source #
Note that this lens does not satisfy all lens laws: If hard
compile-time mode is enabled, then quantities other than zero are
replaced by __IMPOSSIBLE__
.
eTerminationCheck :: Lens' TCEnv (TerminationCheck ()) Source #
emptyFunction :: HasOptions m => m Defn Source #
emptyFunctionData :: HasOptions m => m FunctionData Source #
A template for creating Function
definitions, with sensible
defaults.
emptyFunction_ :: Bool -> Defn Source #
enableCaching :: HasOptions m => m Bool Source #
execError :: (HasCallStack, MonadTCError m) => ExecError -> m a Source #
extendedLambdaName :: String Source #
Base name for extended lambda patterns
finally_ :: TCM a -> TCM b -> TCM a Source #
Execute a finalizer even when an exception is thrown. Does not catch any errors. In case both the regular computation and the finalizer throw an exception, the one of the finalizer is propagated.
flipCmp :: CompareDirection -> CompareDirection Source #
Flip the direction of comparison.
fmapReduce :: (a -> b) -> ReduceM a -> ReduceM b Source #
forkTCM :: TCM () -> TCM () Source #
Runs the given computation in a separate thread, with a copy of the current state and environment.
Note that Agda sometimes uses actual, mutable state. If the
computation given to forkTCM
tries to modify this state, then
bad things can happen, because accesses are not mutually exclusive.
The forkTCM
function has been added mainly to allow the thread to
read (a snapshot of) the current state in a convenient way.
Note also that exceptions which are raised in the thread are not propagated to the parent, so the thread should not do anything important.
freshNoName :: MonadFresh NameId m => Range -> m Name Source #
freshNoName_ :: MonadFresh NameId m => m Name Source #
freshRecordName :: MonadFresh NameId m => m Name Source #
fromCmp :: Comparison -> CompareDirection Source #
Embed Comparison
into CompareDirection
.
fromReduceDefs :: ReduceDefs -> (Bool, [QName]) Source #
funAbstr :: Lens' Defn IsAbstract Source #
Toggle the FunAbstract
flag.
funAbstr_ :: Lens' FunctionData IsAbstract Source #
Toggle the FunAbstract
flag.
funAbstract :: Lens' Defn Bool Source #
Toggle the FunAbstract
flag.
funAbstract_ :: Lens' FunctionData Bool Source #
Toggle the FunAbstract
flag.
funClauses :: Defn -> [Clause] Source #
funCompiled :: Defn -> Maybe CompiledClauses Source #
funCovering :: Defn -> [Clause] Source #
funErasure :: Lens' Defn Bool Source #
Toggle the FunErasure
flag.
funFirstOrder :: Lens' Defn Bool Source #
Toggle the FunFirstOrder
flag.
funFlag_ :: FunctionFlag -> Lens' FunctionData Bool Source #
funInv :: Defn -> FunctionInverse Source #
generalizedFieldName :: String Source #
Base name for generalized variable projections
genericDocError :: (HasCallStack, MonadTCError m) => Doc -> m a Source #
genericError :: (HasCallStack, MonadTCError m) => String -> m a Source #
getGeneralizedFieldName :: QName -> Maybe String Source #
Check whether we have a generalized variable field
getMetaEnv :: MetaVariable -> TCEnv Source #
getMetaInfo :: MetaVariable -> Closure Range Source #
getMetaScope :: MetaVariable -> ScopeInfo Source #
getMetaSig :: MetaVariable -> Signature Source #
getPartialDefs :: ReadTCState m => m (Set QName) Source #
getUserWarnings :: ReadTCState m => m UserWarnings Source #
getsTC :: ReadTCState m => (TCState -> a) -> m a Source #
guardednessOption :: HasOptions m => m Bool Source #
iFullHash :: Interface -> Hash Source #
Combines the source hash and the (full) hashes of the imported modules.
ifTopLevelAndHighlightingLevelIs :: MonadTCEnv tcm => HighlightingLevel -> tcm () -> tcm () Source #
ifTopLevelAndHighlightingLevelIs l m
runs m
when we're
type-checking the top-level module (or before we've started doing
this) and the highlighting level is at least l
.
ifTopLevelAndHighlightingLevelIsOr :: MonadTCEnv tcm => HighlightingLevel -> Bool -> tcm () -> tcm () Source #
ifTopLevelAndHighlightingLevelIs l b m
runs m
when we're
type-checking the top-level module (or before we've started doing
this) and either the highlighting level is at least l
or b
is
True
.
initFileDict :: AbsolutePath -> FileDictWithBuiltins Source #
Empty session state.
initPersistentState :: AbsolutePath -> PersistentTCState Source #
Empty persistent state.
initPreScopeState :: PreScopeState Source #
Empty state of type checker.
initState :: AbsolutePath -> TCState Source #
initStateIO :: IO TCState Source #
initialMetaId :: MetaId Source #
An initial MetaId
.
intSignature :: Lens' Interface Signature Source #
A lens for the iSignature
field of the Interface
type.
interactionError :: (HasCallStack, MonadTCError m) => InteractionError -> m a Source #
internalError :: (HasCallStack, MonadTCError m) => String -> m a Source #
isAbsurdLambdaName :: QName -> Bool Source #
Check whether we have an definition from an absurd lambda.
isCopatternLHS :: [Clause] -> Bool Source #
isDontExpandLast :: ExpandHidden -> Bool Source #
isEmptyFunction :: Defn -> Bool Source #
Checking whether we are dealing with a function yet to be defined.
isExpandLast :: ExpandHidden -> Bool Source #
isExtendedLambda :: Defn -> Bool Source #
isExtendedLambdaName :: QName -> Bool Source #
Check whether we have an definition from an extended lambda.
isReconstructed :: MonadTCEnv m => m Bool Source #
isSourceCodeWarning :: WarningName -> Bool Source #
Should warnings of that type be serialized?
Only when changes in the source code can silence or influence the warning.
isWithFunction :: Defn -> Bool Source #
itableCounts :: Lens' InstanceTable (Map QName Int) Source #
lensTopLevelModuleNames :: Lens' PersistentTCState (BiMap RawTopLevelModuleName ModuleNameHash) Source #
locallyReconstructed :: MonadTCEnv m => m a -> m a Source #
locallyReduceAllDefs :: MonadTCEnv m => m a -> m a Source #
locallyReduceDefs :: MonadTCEnv m => ReduceDefs -> m a -> m a Source #
locallyTC :: MonadTCEnv m => Lens' TCEnv a -> (a -> a) -> m b -> m b Source #
Modify the lens-indicated part of the TCEnv
in a subcomputation.
locatedTypeError :: MonadTCError m => (a -> TypeError) -> HasCallStack => a -> m b Source #
Utility function for 1-arg constructed type errors.
Note that the HasCallStack
constraint is on the *resulting* function.
modifySystem :: (System -> System) -> ExtLamInfo -> ExtLamInfo Source #
modifyTC' :: MonadTCState m => (TCState -> TCState) -> m () Source #
A variant of modifyTC
in which the computation is strict in the
new state.
modifyTCLens :: MonadTCState m => Lens' TCState a -> (a -> a) -> m () Source #
Modify the part of the TCState
focused on by the lens.
modifyTCLens' :: MonadTCState m => Lens' TCState a -> (a -> a) -> m () Source #
Modify the part of the TCState
focused on by the lens
(strictly).
modifyTCLensM :: (MonadTCState m, ReadTCState m) => Lens' TCState a -> (a -> m a) -> m () Source #
Modify a part of the state monadically.
This is an instance of %==
.
notReduced :: a -> MaybeReduced a Source #
onLetBindingType :: (Dom Type -> Dom Type) -> LetBinding -> LetBinding Source #
patternInTeleName :: String Source #
Base name for patterns in telescopes
primAbstr :: Defn -> IsAbstract Source #
primClauses :: Defn -> [Clause] Source #
primCompiled :: Defn -> Maybe CompiledClauses Source #
primFun :: QName -> Arity -> ([Arg Term] -> ReduceM (Reduced MaybeReducedArgs Term)) -> PrimFun Source #
primInv :: Defn -> FunctionInverse Source #
primName :: Defn -> PrimitiveId Source #
primOpaque :: Defn -> IsOpaque Source #
primSortName :: Defn -> BuiltinSort Source #
primSortSort :: Defn -> Sort Source #
projArgInfo :: Projection -> ArgInfo Source #
The info of the principal (record) argument.
projDropPars :: Projection -> ProjOrigin -> Term Source #
Building the projection function (which drops the parameters).
recAbstr :: Defn -> IsAbstract Source #
recConHead :: Defn -> ConHead Source #
recEtaEquality :: Defn -> HasEta Source #
recEtaEquality' :: Defn -> EtaEquality Source #
recNamedCon :: Defn -> Bool Source #
recRecursive :: Defn -> Bool Source #
Is the record type recursive?
recRecursive_ :: RecordData -> Bool Source #
redBind :: ReduceM (Reduced a a') -> (a -> ReduceM b) -> (a' -> ReduceM (Reduced b b')) -> ReduceM (Reduced b b') Source #
Conceptually: redBind m f k = either (return . Left . f) k =<< m
returnTCMT :: forall (m :: Type -> Type) a. Applicative m => a -> TCMT m a Source #
runReduceM :: ReduceM a -> TCM a Source #
runSafeTCM :: TCM a -> TCState -> IO (a, TCState) Source #
runSafeTCM
runs a safe TCM
action (a TCM
action which
cannot fail, except that it might raise IOException
s) in the
initial environment.
runTCM :: MonadIO m => TCEnv -> TCState -> TCMT m a -> m (a, TCState) Source #
Running the type checking monad (most general form).
runTCMTop :: TCM a -> IO (Either TCErr a) Source #
Running the type checking monad on toplevel (with initial state).
runTCMTop' :: MonadIO m => TCMT m a -> m a Source #
setEtaEquality :: EtaEquality -> HasEta -> EtaEquality Source #
Make sure we do not overwrite a user specification.
setTCLens :: MonadTCState m => Lens' TCState a -> a -> m () infix 4 Source #
Overwrite the part of the TCState
focused on by the lens.
setTCLens' :: MonadTCState m => Lens' TCState a -> a -> m () Source #
Overwrite the part of the TCState
focused on by the lens
(strictly).
shouldReduceDef :: MonadTCEnv m => QName -> m Bool Source #
sizedTypesOption :: HasOptions m => m Bool Source #
srcFilePath :: MonadFileId m => SourceFile -> m AbsolutePath Source #
Get the file name of a SourceFile
.
srcFromPath :: MonadFileId m => AbsolutePath -> m SourceFile Source #
Get make a SourceFile
from a file name.
stCurrentModule :: Lens' TCState (Maybe (ModuleName, TopLevelModuleName)) Source #
Note that the lens is "strict".
stateTCLens :: (MonadTCState m, ReadTCState m) => Lens' TCState a -> (a -> (r, a)) -> m r Source #
Modify the part of the TCState
focused on by the lens, and return some result.
stateTCLensM :: (MonadTCState m, ReadTCState m) => Lens' TCState a -> (a -> m (r, a)) -> m r Source #
Modify a part of the state monadically, and return some result.
This is an instance of %%=
.
suffixNameSuggestion :: MetaNameSuggestion -> ArgName -> MetaNameSuggestion Source #
Append an ArgName
to a MetaNameSuggestion
, for computing the
name suggestions of eta-expansion metas. If the MetaNameSuggestion
is empty or an underscore, the field name is taken as the suggestion.
syntaxError :: (HasCallStack, MonadTCError m) => String -> m a Source #
tcWarningOrigin :: TCWarning -> SrcFile Source #
thenTCMT :: forall (m :: Type -> Type) a b. Applicative m => TCMT m a -> TCMT m b -> TCMT m b Source #
toExpandLast :: Bool -> ExpandHidden Source #
toReduceDefs :: (Bool, [QName]) -> ReduceDefs Source #
topLevelModuleFilePath :: ModuleToSource -> TopLevelModuleName -> AbsolutePath Source #
Lookup the path of a top level module name, which must be a known one.
typeError :: (HasCallStack, MonadTCError m) => TypeError -> m a Source #
typeError' :: MonadTCError m => CallStack -> TypeError -> m a Source #
typeError'_ :: (MonadTCEnv m, ReadTCState m) => CallStack -> TypeError -> m TCErr Source #
typeError_ :: (HasCallStack, MonadTCEnv m, ReadTCState m) => TypeError -> m TCErr Source #
unionBuiltin :: Builtin a -> Builtin a -> Builtin a Source #
Union two Builtin
s. Only defined for BuiltinRewriteRelations
.
unquoteError :: (HasCallStack, MonadTCError m) => UnquoteError -> m a Source #
warningName :: Warning -> WarningName Source #
withoutKOption :: HasOptions m => m Bool Source #
data AbstractMode Source #
Constructors
AbstractMode | Abstract things in the current module can be accessed. |
ConcreteMode | No abstract things can be accessed. |
IgnoreAbstractMode | All abstract things can be accessed. |
Instances
data AllowedReduction Source #
Controlling reduce
.
Constructors
ProjectionReductions | (Projection and) projection-like functions may be reduced. |
InlineReductions | Functions marked INLINE may be reduced. |
CopatternReductions | Copattern definitions may be reduced. |
FunctionReductions | Non-recursive functions and primitives may be reduced. |
RecursiveReductions | Even recursive functions may be reduced. |
LevelReductions | Reduce |
TypeLevelReductions | Allow |
UnconfirmedReductions | Functions whose termination has not (yet) been confirmed. |
NonTerminatingReductions | Functions that have failed termination checking. |
Instances
SmallSetElement AllowedReduction Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
NFData AllowedReduction Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: AllowedReduction -> () # | |||||
Bounded AllowedReduction Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Enum AllowedReduction Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods succ :: AllowedReduction -> AllowedReduction # pred :: AllowedReduction -> AllowedReduction # toEnum :: Int -> AllowedReduction # fromEnum :: AllowedReduction -> Int # enumFrom :: AllowedReduction -> [AllowedReduction] # enumFromThen :: AllowedReduction -> AllowedReduction -> [AllowedReduction] # enumFromTo :: AllowedReduction -> AllowedReduction -> [AllowedReduction] # enumFromThenTo :: AllowedReduction -> AllowedReduction -> AllowedReduction -> [AllowedReduction] # | |||||
Generic AllowedReduction Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
Methods from :: AllowedReduction -> Rep AllowedReduction x # to :: Rep AllowedReduction x -> AllowedReduction # | |||||
Ix AllowedReduction Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods range :: (AllowedReduction, AllowedReduction) -> [AllowedReduction] # index :: (AllowedReduction, AllowedReduction) -> AllowedReduction -> Int # unsafeIndex :: (AllowedReduction, AllowedReduction) -> AllowedReduction -> Int # inRange :: (AllowedReduction, AllowedReduction) -> AllowedReduction -> Bool # rangeSize :: (AllowedReduction, AllowedReduction) -> Int # unsafeRangeSize :: (AllowedReduction, AllowedReduction) -> Int # | |||||
Show AllowedReduction Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods showsPrec :: Int -> AllowedReduction -> ShowS # show :: AllowedReduction -> String # showList :: [AllowedReduction] -> ShowS # | |||||
Eq AllowedReduction Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods (==) :: AllowedReduction -> AllowedReduction -> Bool # (/=) :: AllowedReduction -> AllowedReduction -> Bool # | |||||
Ord AllowedReduction Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods compare :: AllowedReduction -> AllowedReduction -> Ordering # (<) :: AllowedReduction -> AllowedReduction -> Bool # (<=) :: AllowedReduction -> AllowedReduction -> Bool # (>) :: AllowedReduction -> AllowedReduction -> Bool # (>=) :: AllowedReduction -> AllowedReduction -> Bool # max :: AllowedReduction -> AllowedReduction -> AllowedReduction # min :: AllowedReduction -> AllowedReduction -> AllowedReduction # | |||||
type Rep AllowedReduction Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep AllowedReduction = D1 ('MetaData "AllowedReduction" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (((C1 ('MetaCons "ProjectionReductions" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "InlineReductions" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "CopatternReductions" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "FunctionReductions" 'PrefixI 'False) (U1 :: Type -> Type))) :+: ((C1 ('MetaCons "RecursiveReductions" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "LevelReductions" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "TypeLevelReductions" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "UnconfirmedReductions" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "NonTerminatingReductions" 'PrefixI 'False) (U1 :: Type -> Type))))) |
data ArgsCheckState a Source #
Constructors
ACState | |
Fields
|
Instances
Show a => Show (ArgsCheckState a) Source # | |
Defined in Agda.TypeChecking.Monad.Base Methods showsPrec :: Int -> ArgsCheckState a -> ShowS # show :: ArgsCheckState a -> String # showList :: [ArgsCheckState a] -> ShowS # |
Constructors
AxiomData | |
Fields
|
Instances
NFData AxiomData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Generic AxiomData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
| |||||
Show AxiomData Source # | |||||
type Rep AxiomData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base |
type Backend = Backend_boot Definition TCM Source #
type Backend' opts env menv mod def = Backend'_boot Definition TCM opts env menv mod def Source #
newtype BlockT (m :: Type -> Type) a Source #
Instances
MonadTrans BlockT Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
HasOptions m => HasOptions (BlockT m) Source # | |
Defined in Agda.TypeChecking.Monad.Base Methods | |
Monad m => MonadBlock (BlockT m) Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
MonadReduce m => MonadReduce (BlockT m) Source # | |
Defined in Agda.TypeChecking.Monad.Base Methods liftReduce :: ReduceM a -> BlockT m a Source # | |
MonadTCEnv m => MonadTCEnv (BlockT m) Source # | |
MonadTCM m => MonadTCM (BlockT m) Source # | |
MonadTCState m => MonadTCState (BlockT m) Source # | |
ReadTCState m => ReadTCState (BlockT m) Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
HasBuiltins m => HasBuiltins (BlockT m) Source # | |
Defined in Agda.TypeChecking.Monad.Builtin Methods getBuiltinThing :: SomeBuiltin -> BlockT m (Maybe (Builtin PrimFun)) Source # | |
MonadAddContext m => MonadAddContext (BlockT m) Source # | |
Defined in Agda.TypeChecking.Monad.Context Methods addCtx :: Name -> Dom Type -> BlockT m a -> BlockT m a Source # addLetBinding' :: Origin -> Name -> Term -> Dom Type -> BlockT m a -> BlockT m a Source # updateContext :: Substitution -> (Context -> Context) -> BlockT m a -> BlockT m a Source # withFreshName :: Range -> ArgName -> (Name -> BlockT m a) -> BlockT m a Source # | |
MonadDebug m => MonadDebug (BlockT m) Source # | |
Defined in Agda.TypeChecking.Monad.Debug Methods formatDebugMessage :: VerboseKey -> VerboseLevel -> TCM Doc -> BlockT m String Source # traceDebugMessage :: VerboseKey -> VerboseLevel -> String -> BlockT m a -> BlockT m a Source # verboseBracket :: VerboseKey -> VerboseLevel -> String -> BlockT m a -> BlockT m a Source # getVerbosity :: BlockT m Verbosity Source # getProfileOptions :: BlockT m ProfileOptions Source # isDebugPrinting :: BlockT m Bool Source # nowDebugPrinting :: BlockT m a -> BlockT m a Source # | |
PureTCM m => PureTCM (BlockT m) Source # | |
Defined in Agda.TypeChecking.Monad.Pure | |
HasConstInfo m => HasConstInfo (BlockT m) Source # | |
Defined in Agda.TypeChecking.Monad.Signature Methods getConstInfo :: QName -> BlockT m Definition Source # getConstInfo' :: QName -> BlockT m (Either SigError Definition) Source # getRewriteRulesFor :: QName -> BlockT m RewriteRules Source # | |
Monad m => Applicative (BlockT m) Source # | |
Functor m => Functor (BlockT m) Source # | |
Monad m => Monad (BlockT m) Source # | |
MonadFail m => MonadFail (BlockT m) Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
MonadIO m => MonadIO (BlockT m) Source # | |
Defined in Agda.TypeChecking.Monad.Base |
Constructors
Builtin Term | |
Prim pf | |
BuiltinRewriteRelations (Set QName) |
|
Instances
Functor Builtin Source # | |||||
Foldable Builtin Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods fold :: Monoid m => Builtin m -> m # foldMap :: Monoid m => (a -> m) -> Builtin a -> m # foldMap' :: Monoid m => (a -> m) -> Builtin a -> m # foldr :: (a -> b -> b) -> b -> Builtin a -> b # foldr' :: (a -> b -> b) -> b -> Builtin a -> b # foldl :: (b -> a -> b) -> b -> Builtin a -> b # foldl' :: (b -> a -> b) -> b -> Builtin a -> b # foldr1 :: (a -> a -> a) -> Builtin a -> a # foldl1 :: (a -> a -> a) -> Builtin a -> a # elem :: Eq a => a -> Builtin a -> Bool # maximum :: Ord a => Builtin a -> a # minimum :: Ord a => Builtin a -> a # | |||||
Traversable Builtin Source # | |||||
NamesIn a => NamesIn (Builtin a) Source # | |||||
Defined in Agda.Syntax.Internal.Names | |||||
InstantiateFull a => InstantiateFull (Builtin a) Source # | |||||
Defined in Agda.TypeChecking.Reduce | |||||
EmbPrj a => EmbPrj (Builtin a) Source # | |||||
NFData pf => NFData (Builtin pf) Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Generic (Builtin pf) Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
| |||||
Show pf => Show (Builtin pf) Source # | |||||
type Rep (Builtin pf) Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep (Builtin pf) = D1 ('MetaData "Builtin" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "Builtin" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)) :+: (C1 ('MetaCons "Prim" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 pf)) :+: C1 ('MetaCons "BuiltinRewriteRelations" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Set QName))))) |
data BuiltinDescriptor Source #
Constructors
BuiltinData (TCM Type) [BuiltinId] | |
BuiltinDataCons (TCM Type) | |
BuiltinPrim PrimitiveId (Term -> TCM ()) | |
BuiltinSort BuiltinSort | |
BuiltinPostulate Relevance (TCM Type) | |
BuiltinUnknown (Maybe (TCM Type)) (Term -> Type -> TCM ()) | Builtin of any kind.
Type can be checked ( |
data BuiltinInfo Source #
Constructors
BuiltinInfo | |
Fields |
data BuiltinSort Source #
Constructors
SortUniv Univ | |
SortOmega Univ | |
SortIntervalUniv | |
SortLevelUniv |
Instances
KillRange BuiltinSort Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods | |||||
EmbPrj BuiltinSort Source # | |||||
NFData BuiltinSort Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: BuiltinSort -> () # | |||||
Generic BuiltinSort Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
| |||||
Show BuiltinSort Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods showsPrec :: Int -> BuiltinSort -> ShowS # show :: BuiltinSort -> String # showList :: [BuiltinSort] -> ShowS # | |||||
Eq BuiltinSort Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
type Rep BuiltinSort Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep BuiltinSort = D1 ('MetaData "BuiltinSort" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) ((C1 ('MetaCons "SortUniv" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Univ)) :+: C1 ('MetaCons "SortOmega" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Univ))) :+: (C1 ('MetaCons "SortIntervalUniv" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "SortLevelUniv" 'PrefixI 'False) (U1 :: Type -> Type))) |
type BuiltinThings = BuiltinThings' PrimFun Source #
type BuiltinThings' pf = Map SomeBuiltin (Builtin pf) Source #
type CachedTypeCheckLog = [(TypeCheckAction, PostScopeState)] Source #
A log of what the type checker does and states after the action is completed. The cached version is stored first executed action first.
Constructors
CheckClause Type SpineClause | |
CheckLHS SpineLHS | |
CheckPattern Pattern Telescope Type | |
CheckPatternLinearityType Name | |
CheckPatternLinearityValue Name | |
CheckLetBinding LetBinding | |
InferExpr Expr | |
CheckExprCall Comparison Expr Type | |
CheckDotPattern Expr Term | |
CheckProjection Range QName Type | |
IsTypeCall Comparison Expr Sort | |
IsType_ Expr | |
InferVar Name | |
InferDef QName | |
CheckArguments Expr [NamedArg Expr] Type (Maybe Type) | |
CheckMetaSolution Range MetaId Type Term | |
CheckTargetType Range Type Type | |
CheckDataDef Range QName [LamBinding] [Constructor] | |
CheckRecDef Range QName [LamBinding] [Constructor] | |
CheckConstructor QName Telescope Sort Constructor | |
CheckConArgFitsIn QName Bool Type Sort | |
CheckFunDefCall Range QName [Clause] Bool | Highlight (interactively) if and only if the boolean is |
CheckPragma Range Pragma | |
CheckPrimitive Range QName Expr | |
CheckIsEmpty Range Type | |
CheckConfluence QName QName | |
CheckModuleParameters ModuleName Telescope | |
CheckWithFunctionType Type | |
CheckSectionApplication Range Erased ModuleName ModuleApplication | |
CheckNamedWhere ModuleName | |
CheckIApplyConfluence | Checking a clause for confluence with endpoint reductions. Always
|
ScopeCheckExpr Expr | |
ScopeCheckDeclaration NiceDeclaration | |
ScopeCheckLHS QName Pattern | |
NoHighlighting | |
ModuleContents | Interaction command: show module contents. |
SetRange Range | used by |
Instances
Pretty Call Source # | |||||
HasRange Call Source # | |||||
PrettyTCM Call Source # | |||||
Defined in Agda.TypeChecking.Pretty.Call | |||||
NFData Call Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Generic Call Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
| |||||
type Rep Call Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep Call = D1 ('MetaData "Call" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (((((C1 ('MetaCons "CheckClause" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 SpineClause)) :+: C1 ('MetaCons "CheckLHS" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 SpineLHS))) :+: (C1 ('MetaCons "CheckPattern" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Pattern) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Telescope) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type))) :+: C1 ('MetaCons "CheckPatternLinearityType" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Name)))) :+: ((C1 ('MetaCons "CheckPatternLinearityValue" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Name)) :+: C1 ('MetaCons "CheckLetBinding" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 LetBinding))) :+: (C1 ('MetaCons "InferExpr" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr)) :+: (C1 ('MetaCons "CheckExprCall" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type))) :+: C1 ('MetaCons "CheckDotPattern" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)))))) :+: (((C1 ('MetaCons "CheckProjection" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type))) :+: C1 ('MetaCons "IsTypeCall" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Sort)))) :+: (C1 ('MetaCons "IsType_" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr)) :+: C1 ('MetaCons "InferVar" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Name)))) :+: ((C1 ('MetaCons "InferDef" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "CheckArguments" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [NamedArg Expr])) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe Type))))) :+: (C1 ('MetaCons "CheckMetaSolution" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 MetaId)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term))) :+: (C1 ('MetaCons "CheckTargetType" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type))) :+: C1 ('MetaCons "CheckDataDef" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [LamBinding]) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Constructor])))))))) :+: ((((C1 ('MetaCons "CheckRecDef" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [LamBinding]) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Constructor]))) :+: C1 ('MetaCons "CheckConstructor" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Telescope)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Sort) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Constructor)))) :+: (C1 ('MetaCons "CheckConArgFitsIn" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Sort))) :+: C1 ('MetaCons "CheckFunDefCall" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Clause]) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool))))) :+: ((C1 ('MetaCons "CheckPragma" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Pragma)) :+: C1 ('MetaCons "CheckPrimitive" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr)))) :+: (C1 ('MetaCons "CheckIsEmpty" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)) :+: (C1 ('MetaCons "CheckConfluence" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "CheckModuleParameters" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ModuleName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Telescope)))))) :+: (((C1 ('MetaCons "CheckWithFunctionType" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)) :+: C1 ('MetaCons "CheckSectionApplication" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Erased)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ModuleName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ModuleApplication)))) :+: (C1 ('MetaCons "CheckNamedWhere" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ModuleName)) :+: (C1 ('MetaCons "CheckIApplyConfluence" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term))) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)))) :+: C1 ('MetaCons "ScopeCheckExpr" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr))))) :+: ((C1 ('MetaCons "ScopeCheckDeclaration" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 NiceDeclaration)) :+: C1 ('MetaCons "ScopeCheckLHS" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Pattern))) :+: (C1 ('MetaCons "NoHighlighting" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "ModuleContents" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "SetRange" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range)))))))) |
Information about a call.
Constructors
CallInfo | |
Fields
|
Instances
Pretty CallInfo Source # | We only | ||||
HasRange CallInfo Source # | |||||
PrettyTCM CallInfo Source # | |||||
Defined in Agda.TypeChecking.Pretty.Call | |||||
NFData CallInfo Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Generic CallInfo Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
| |||||
Show CallInfo Source # | |||||
type Rep CallInfo Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep CallInfo = D1 ('MetaData "CallInfo" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "CallInfo" 'PrefixI 'True) (S1 ('MetaSel ('Just "callInfoTarget") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Just "callInfoCall") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Closure Term)))) |
A candidate solution for an instance meta is a term with its type. It may be the case that the candidate is not fully applied yet or of the wrong type, hence the need for the type.
Constructors
Candidate | |
Fields |
Instances
HasOverlapMode Candidate Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods | |||||
Free Candidate Source # | |||||
PrettyTCM Candidate Source # | |||||
Defined in Agda.TypeChecking.Pretty | |||||
Instantiate Candidate Source # | |||||
Defined in Agda.TypeChecking.Reduce | |||||
InstantiateFull Candidate Source # | |||||
Defined in Agda.TypeChecking.Reduce | |||||
Normalise Candidate Source # | |||||
Defined in Agda.TypeChecking.Reduce | |||||
Reduce Candidate Source # | |||||
Simplify Candidate Source # | |||||
Defined in Agda.TypeChecking.Reduce | |||||
Subst Candidate Source # | |||||
Defined in Agda.TypeChecking.Substitute Associated Types
Methods applySubst :: Substitution' (SubstArg Candidate) -> Candidate -> Candidate Source # | |||||
NFData Candidate Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Generic Candidate Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
| |||||
Show Candidate Source # | |||||
Eq Candidate Source # | |||||
Ord Candidate Source # | |||||
type SubstArg Candidate Source # | |||||
Defined in Agda.TypeChecking.Substitute | |||||
type Rep Candidate Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep Candidate = D1 ('MetaData "Candidate" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "Candidate" 'PrefixI 'True) ((S1 ('MetaSel ('Just "candidateKind") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 CandidateKind) :*: S1 ('MetaSel ('Just "candidateTerm") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)) :*: (S1 ('MetaSel ('Just "candidateType") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type) :*: S1 ('MetaSel ('Just "candidateOverlap") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 OverlapMode)))) |
data CandidateKind Source #
Constructors
LocalCandidate | |
GlobalCandidate QName |
Instances
NFData CandidateKind Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: CandidateKind -> () # | |||||
Generic CandidateKind Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
| |||||
Show CandidateKind Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods showsPrec :: Int -> CandidateKind -> ShowS # show :: CandidateKind -> String # showList :: [CandidateKind] -> ShowS # | |||||
Eq CandidateKind Source # | |||||
Defined in Agda.TypeChecking.Substitute Methods (==) :: CandidateKind -> CandidateKind -> Bool # (/=) :: CandidateKind -> CandidateKind -> Bool # | |||||
Ord CandidateKind Source # | |||||
Defined in Agda.TypeChecking.Substitute Methods compare :: CandidateKind -> CandidateKind -> Ordering # (<) :: CandidateKind -> CandidateKind -> Bool # (<=) :: CandidateKind -> CandidateKind -> Bool # (>) :: CandidateKind -> CandidateKind -> Bool # (>=) :: CandidateKind -> CandidateKind -> Bool # max :: CandidateKind -> CandidateKind -> CandidateKind # min :: CandidateKind -> CandidateKind -> CandidateKind # | |||||
type Rep CandidateKind Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep CandidateKind = D1 ('MetaData "CandidateKind" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "LocalCandidate" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "GlobalCandidate" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName))) |
data CannotQuote Source #
Extra information for error CannotQuote
.
Constructors
CannotQuoteAmbiguous (List2 QName) |
|
CannotQuoteExpression Expr |
|
CannotQuoteHidden |
|
CannotQuoteNothing |
|
CannotQuotePattern (NamedArg Pattern) |
|
Instances
NFData CannotQuote Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: CannotQuote -> () # | |||||
Generic CannotQuote Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
| |||||
Show CannotQuote Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods showsPrec :: Int -> CannotQuote -> ShowS # show :: CannotQuote -> String # showList :: [CannotQuote] -> ShowS # | |||||
type Rep CannotQuote Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep CannotQuote = D1 ('MetaData "CannotQuote" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) ((C1 ('MetaCons "CannotQuoteAmbiguous" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List2 QName))) :+: C1 ('MetaCons "CannotQuoteExpression" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr))) :+: (C1 ('MetaCons "CannotQuoteHidden" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "CannotQuoteNothing" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "CannotQuotePattern" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (NamedArg Pattern)))))) |
data CheckedArg Source #
Constructors
CheckedArg | |
Fields
|
Instances
Show CheckedArg Source # | |
Defined in Agda.TypeChecking.Monad.Base Methods showsPrec :: Int -> CheckedArg -> ShowS # show :: CheckedArg -> String # showList :: [CheckedArg] -> ShowS # |
data CheckedTarget Source #
Solving a CheckArgs
constraint may or may not check the target type. If
it did, it returns a handle to any unsolved constraints.
Constructors
CheckedTarget (Maybe ProblemId) | |
NotCheckedTarget |
newtype CheckpointId Source #
Constructors
CheckpointId Int |
Instances
Pretty CheckpointId Source # | |
Defined in Agda.TypeChecking.Monad.Base Methods pretty :: CheckpointId -> Doc Source # prettyPrec :: Int -> CheckpointId -> Doc Source # prettyList :: [CheckpointId] -> Doc Source # | |
HasFresh CheckpointId Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
PrettyTCM CheckpointId Source # | |
Defined in Agda.TypeChecking.Pretty Methods prettyTCM :: MonadPretty m => CheckpointId -> m Doc Source # | |
EmbPrj CheckpointId Source # | |
NFData CheckpointId Source # | |
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: CheckpointId -> () # | |
Enum CheckpointId Source # | |
Defined in Agda.TypeChecking.Monad.Base Methods succ :: CheckpointId -> CheckpointId # pred :: CheckpointId -> CheckpointId # toEnum :: Int -> CheckpointId # fromEnum :: CheckpointId -> Int # enumFrom :: CheckpointId -> [CheckpointId] # enumFromThen :: CheckpointId -> CheckpointId -> [CheckpointId] # enumFromTo :: CheckpointId -> CheckpointId -> [CheckpointId] # enumFromThenTo :: CheckpointId -> CheckpointId -> CheckpointId -> [CheckpointId] # | |
Num CheckpointId Source # | |
Defined in Agda.TypeChecking.Monad.Base Methods (+) :: CheckpointId -> CheckpointId -> CheckpointId # (-) :: CheckpointId -> CheckpointId -> CheckpointId # (*) :: CheckpointId -> CheckpointId -> CheckpointId # negate :: CheckpointId -> CheckpointId # abs :: CheckpointId -> CheckpointId # signum :: CheckpointId -> CheckpointId # fromInteger :: Integer -> CheckpointId # | |
Integral CheckpointId Source # | |
Defined in Agda.TypeChecking.Monad.Base Methods quot :: CheckpointId -> CheckpointId -> CheckpointId # rem :: CheckpointId -> CheckpointId -> CheckpointId # div :: CheckpointId -> CheckpointId -> CheckpointId # mod :: CheckpointId -> CheckpointId -> CheckpointId # quotRem :: CheckpointId -> CheckpointId -> (CheckpointId, CheckpointId) # divMod :: CheckpointId -> CheckpointId -> (CheckpointId, CheckpointId) # toInteger :: CheckpointId -> Integer # | |
Real CheckpointId Source # | |
Defined in Agda.TypeChecking.Monad.Base Methods toRational :: CheckpointId -> Rational # | |
Show CheckpointId Source # | |
Defined in Agda.TypeChecking.Monad.Base Methods showsPrec :: Int -> CheckpointId -> ShowS # show :: CheckpointId -> String # showList :: [CheckpointId] -> ShowS # | |
Eq CheckpointId Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
Ord CheckpointId Source # | |
Defined in Agda.TypeChecking.Monad.Base Methods compare :: CheckpointId -> CheckpointId -> Ordering # (<) :: CheckpointId -> CheckpointId -> Bool # (<=) :: CheckpointId -> CheckpointId -> Bool # (>) :: CheckpointId -> CheckpointId -> Bool # (>=) :: CheckpointId -> CheckpointId -> Bool # max :: CheckpointId -> CheckpointId -> CheckpointId # min :: CheckpointId -> CheckpointId -> CheckpointId # |
Constructors
Closure | |
Fields
|
Instances
Functor Closure Source # | |||||
Foldable Closure Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods fold :: Monoid m => Closure m -> m # foldMap :: Monoid m => (a -> m) -> Closure a -> m # foldMap' :: Monoid m => (a -> m) -> Closure a -> m # foldr :: (a -> b -> b) -> b -> Closure a -> b # foldr' :: (a -> b -> b) -> b -> Closure a -> b # foldl :: (b -> a -> b) -> b -> Closure a -> b # foldl' :: (b -> a -> b) -> b -> Closure a -> b # foldr1 :: (a -> a -> a) -> Closure a -> a # foldl1 :: (a -> a -> a) -> Closure a -> a # elem :: Eq a => a -> Closure a -> Bool # maximum :: Ord a => Closure a -> a # minimum :: Ord a => Closure a -> a # | |||||
LensIsAbstract (Closure a) Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods lensIsAbstract :: Lens' (Closure a) IsAbstract Source # | |||||
HasRange a => HasRange (Closure a) Source # | |||||
KillRange a => KillRange (Closure a) Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods killRange :: KillRangeT (Closure a) Source # | |||||
MentionsMeta a => MentionsMeta (Closure a) Source # | |||||
Defined in Agda.TypeChecking.MetaVars.Mention | |||||
LensTCEnv (Closure a) Source # | |||||
PrettyTCM a => PrettyTCM (Closure a) Source # | |||||
Defined in Agda.TypeChecking.Pretty | |||||
Instantiate a => Instantiate (Closure a) Source # | |||||
Defined in Agda.TypeChecking.Reduce | |||||
InstantiateFull a => InstantiateFull (Closure a) Source # | |||||
Defined in Agda.TypeChecking.Reduce | |||||
Normalise a => Normalise (Closure a) Source # | |||||
Defined in Agda.TypeChecking.Reduce | |||||
Reduce a => Reduce (Closure a) Source # | |||||
Simplify a => Simplify (Closure a) Source # | |||||
Defined in Agda.TypeChecking.Reduce | |||||
NFData a => NFData (Closure a) Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Generic (Closure a) Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
| |||||
Show a => Show (Closure a) Source # | |||||
LensClosure (Closure a) a Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
type Rep (Closure a) Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep (Closure a) = D1 ('MetaData "Closure" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "Closure" 'PrefixI 'True) ((S1 ('MetaSel ('Just "clSignature") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Signature) :*: S1 ('MetaSel ('Just "clEnv") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 TCEnv)) :*: (S1 ('MetaSel ('Just "clScope") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ScopeInfo) :*: (S1 ('MetaSel ('Just "clModuleCheckpoints") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Map ModuleName CheckpointId)) :*: S1 ('MetaSel ('Just "clValue") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 a))))) |
Constructors
CompKit | |
Fields
|
Instances
NamesIn CompKit Source # | |||||
Defined in Agda.Syntax.Internal.Names | |||||
KillRange CompKit Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods | |||||
EmbPrj CompKit Source # | |||||
NFData CompKit Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Generic CompKit Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
| |||||
Show CompKit Source # | |||||
Eq CompKit Source # | |||||
Ord CompKit Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
type Rep CompKit Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep CompKit = D1 ('MetaData "CompKit" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "CompKit" 'PrefixI 'True) (S1 ('MetaSel ('Just "nameOfHComp") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe QName)) :*: S1 ('MetaSel ('Just "nameOfTransp") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe QName)))) |
We can either compare two terms at a given type, or compare two types without knowing (or caring about) their sorts.
Constructors
AsTermsOf Type |
|
AsSizes | Replaces |
AsTypes |
Instances
Pretty CompareAs Source # | |||||
TermLike CompareAs Source # | |||||
AllMetas CompareAs Source # | |||||
Free CompareAs Source # | |||||
MentionsMeta CompareAs Source # | |||||
Defined in Agda.TypeChecking.MetaVars.Mention | |||||
IsSizeType CompareAs Source # | |||||
Defined in Agda.TypeChecking.Monad.SizedTypes Methods isSizeType :: (HasOptions m, HasBuiltins m) => CompareAs -> m (Maybe BoundedSize) Source # | |||||
PrettyTCM CompareAs Source # | |||||
Defined in Agda.TypeChecking.Pretty | |||||
Instantiate CompareAs Source # | |||||
Defined in Agda.TypeChecking.Reduce | |||||
InstantiateFull CompareAs Source # | |||||
Defined in Agda.TypeChecking.Reduce | |||||
IsMeta CompareAs Source # | |||||
Normalise CompareAs Source # | |||||
Defined in Agda.TypeChecking.Reduce | |||||
Reduce CompareAs Source # | |||||
Simplify CompareAs Source # | |||||
Defined in Agda.TypeChecking.Reduce | |||||
Subst CompareAs Source # | |||||
Defined in Agda.TypeChecking.Substitute Associated Types
Methods applySubst :: Substitution' (SubstArg CompareAs) -> CompareAs -> CompareAs Source # | |||||
NFData CompareAs Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Generic CompareAs Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
| |||||
Show CompareAs Source # | |||||
type SubstArg CompareAs Source # | |||||
Defined in Agda.TypeChecking.Substitute | |||||
type Rep CompareAs Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep CompareAs = D1 ('MetaData "CompareAs" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "AsTermsOf" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)) :+: (C1 ('MetaCons "AsSizes" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "AsTypes" 'PrefixI 'False) (U1 :: Type -> Type))) |
data CompareDirection Source #
An extension of Comparison
to >=
.
Instances
Pretty CompareDirection Source # | |
Defined in Agda.TypeChecking.Monad.Base Methods pretty :: CompareDirection -> Doc Source # prettyPrec :: Int -> CompareDirection -> Doc Source # prettyList :: [CompareDirection] -> Doc Source # | |
Show CompareDirection Source # | |
Defined in Agda.TypeChecking.Monad.Base Methods showsPrec :: Int -> CompareDirection -> ShowS # show :: CompareDirection -> String # showList :: [CompareDirection] -> ShowS # | |
Eq CompareDirection Source # | |
Defined in Agda.TypeChecking.Monad.Base Methods (==) :: CompareDirection -> CompareDirection -> Bool # (/=) :: CompareDirection -> CompareDirection -> Bool # |
data CompilerPragma Source #
The backends are responsible for parsing their own pragmas.
Constructors
CompilerPragma Range String |
Instances
HasRange CompilerPragma Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods getRange :: CompilerPragma -> Range Source # | |||||
KillRange CompiledRepresentation Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods | |||||
EmbPrj CompilerPragma Source # | |||||
NFData CompilerPragma Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: CompilerPragma -> () # | |||||
Generic CompilerPragma Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
Methods from :: CompilerPragma -> Rep CompilerPragma x # to :: Rep CompilerPragma x -> CompilerPragma # | |||||
Show CompilerPragma Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods showsPrec :: Int -> CompilerPragma -> ShowS # show :: CompilerPragma -> String # showList :: [CompilerPragma] -> ShowS # | |||||
Eq CompilerPragma Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods (==) :: CompilerPragma -> CompilerPragma -> Bool # (/=) :: CompilerPragma -> CompilerPragma -> Bool # | |||||
type Rep CompilerPragma Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep CompilerPragma = D1 ('MetaData "CompilerPragma" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "CompilerPragma" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String))) |
data Constraint Source #
Constructors
ValueCmp Comparison CompareAs Term Term | |
ValueCmpOnFace Comparison Term Type Term Term | |
ElimCmp [Polarity] [IsForced] Type Term [Elim] [Elim] | |
SortCmp Comparison Sort Sort | |
LevelCmp Comparison Level Level | |
HasBiggerSort Sort | |
HasPTSRule (Dom Type) (Abs Sort) | |
CheckDataSort QName Sort | Check that the sort |
CheckMetaInst MetaId | |
CheckType Type | |
UnBlock MetaId | Meta created for a term blocked by a postponed type checking problem or unsolved
constraints. The |
IsEmpty Range Type | The range is the one of the absurd pattern. |
CheckSizeLtSat Term | Check that the |
FindInstance MetaId (Maybe [Candidate]) | the first argument is the instance argument and the second one is the list of candidates (or Nothing if we haven’t determined the list of candidates yet) |
ResolveInstanceHead QName | Resolve the head symbol of the type that the given instance targets |
CheckFunDef DefInfo QName [Clause] TCErr | Last argument is the error causing us to postpone. |
UnquoteTactic Term Term Type | First argument is computation and the others are hole and goal type |
CheckLockedVars Term Type (Arg Term) Type |
|
UsableAtModality WhyCheckModality (Maybe Sort) Modality Term | Is the term usable at the given modality?
This check should run if the |
Instances
TermLike Constraint Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods traverseTermM :: Monad m => (Term -> m Term) -> Constraint -> m Constraint Source # foldTerm :: Monoid m => (Term -> m) -> Constraint -> m Source # | |||||
AllMetas Constraint Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
HasRange Constraint Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods getRange :: Constraint -> Range Source # | |||||
Reify Constraint Source # | |||||
Defined in Agda.Interaction.BasicOps Associated Types
Methods reify :: MonadReify m => Constraint -> m (ReifiesTo Constraint) Source # reifyWhen :: MonadReify m => Bool -> Constraint -> m (ReifiesTo Constraint) Source # | |||||
Free Constraint Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
MentionsMeta Constraint Source # | |||||
Defined in Agda.TypeChecking.MetaVars.Mention Methods mentionsMetas :: HashSet MetaId -> Constraint -> Bool Source # | |||||
PrettyTCM Constraint Source # | |||||
Defined in Agda.TypeChecking.Pretty.Constraint Methods prettyTCM :: MonadPretty m => Constraint -> m Doc Source # | |||||
Instantiate Constraint Source # | |||||
Defined in Agda.TypeChecking.Reduce Methods | |||||
InstantiateFull Constraint Source # | |||||
Defined in Agda.TypeChecking.Reduce Methods instantiateFull' :: Constraint -> ReduceM Constraint Source # | |||||
Normalise Constraint Source # | |||||
Defined in Agda.TypeChecking.Reduce Methods normalise' :: Constraint -> ReduceM Constraint Source # | |||||
Reduce Constraint Source # | |||||
Defined in Agda.TypeChecking.Reduce Methods reduce' :: Constraint -> ReduceM Constraint Source # reduceB' :: Constraint -> ReduceM (Blocked Constraint) Source # | |||||
Simplify Constraint Source # | |||||
Defined in Agda.TypeChecking.Reduce Methods simplify' :: Constraint -> ReduceM Constraint | |||||
Subst Constraint Source # | |||||
Defined in Agda.TypeChecking.Substitute Associated Types
Methods applySubst :: Substitution' (SubstArg Constraint) -> Constraint -> Constraint Source # | |||||
NFData Constraint Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: Constraint -> () # | |||||
Generic Constraint Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
| |||||
Show Constraint Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods showsPrec :: Int -> Constraint -> ShowS # show :: Constraint -> String # showList :: [Constraint] -> ShowS # | |||||
type ReifiesTo Constraint Source # | |||||
Defined in Agda.Interaction.BasicOps | |||||
type SubstArg Constraint Source # | |||||
Defined in Agda.TypeChecking.Substitute | |||||
type Rep Constraint Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep Constraint = D1 ('MetaData "Constraint" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) ((((C1 ('MetaCons "ValueCmp" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 CompareAs)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term))) :+: C1 ('MetaCons "ValueCmpOnFace" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term))))) :+: (C1 ('MetaCons "ElimCmp" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Polarity]) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [IsForced]) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type))) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Elim]) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Elim])))) :+: C1 ('MetaCons "SortCmp" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Sort) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Sort))))) :+: ((C1 ('MetaCons "LevelCmp" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Level) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Level))) :+: C1 ('MetaCons "HasBiggerSort" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Sort))) :+: (C1 ('MetaCons "HasPTSRule" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Dom Type)) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Abs Sort))) :+: (C1 ('MetaCons "CheckDataSort" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Sort)) :+: C1 ('MetaCons "CheckMetaInst" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 MetaId)))))) :+: (((C1 ('MetaCons "CheckType" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)) :+: C1 ('MetaCons "UnBlock" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 MetaId))) :+: (C1 ('MetaCons "IsEmpty" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)) :+: (C1 ('MetaCons "CheckSizeLtSat" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)) :+: C1 ('MetaCons "FindInstance" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 MetaId) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe [Candidate])))))) :+: ((C1 ('MetaCons "ResolveInstanceHead" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "CheckFunDef" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 DefInfo) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Clause]) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 TCErr)))) :+: (C1 ('MetaCons "UnquoteTactic" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type))) :+: (C1 ('MetaCons "CheckLockedVars" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Arg Term)) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type))) :+: C1 ('MetaCons "UsableAtModality" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 WhyCheckModality) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe Sort))) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Modality) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)))))))) |
type Constraints = [ProblemConstraint] Source #
data ConstructorData Source #
Constructors
ConstructorData | |
Fields
|
Instances
Pretty ConstructorData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods pretty :: ConstructorData -> Doc Source # prettyPrec :: Int -> ConstructorData -> Doc Source # prettyList :: [ConstructorData] -> Doc Source # | |||||
NFData ConstructorData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: ConstructorData -> () # | |||||
Generic ConstructorData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
Methods from :: ConstructorData -> Rep ConstructorData x # to :: Rep ConstructorData x -> ConstructorData # | |||||
Show ConstructorData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods showsPrec :: Int -> ConstructorData -> ShowS # show :: ConstructorData -> String # showList :: [ConstructorData] -> ShowS # | |||||
type Rep ConstructorData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep ConstructorData = D1 ('MetaData "ConstructorData" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "ConstructorData" 'PrefixI 'True) (((S1 ('MetaSel ('Just "_conPars") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Int) :*: S1 ('MetaSel ('Just "_conArity") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Int)) :*: (S1 ('MetaSel ('Just "_conSrcCon") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ConHead) :*: (S1 ('MetaSel ('Just "_conData") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Just "_conAbstr") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 IsAbstract)))) :*: ((S1 ('MetaSel ('Just "_conComp") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 CompKit) :*: (S1 ('MetaSel ('Just "_conProj") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe [QName])) :*: S1 ('MetaSel ('Just "_conForced") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [IsForced]))) :*: (S1 ('MetaSel ('Just "_conErased") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe [Bool])) :*: (S1 ('MetaSel ('Just "_conErasure") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Bool) :*: S1 ('MetaSel ('Just "_conInline") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Bool)))))) |
data ConstructorDisambiguationData Source #
Information we have constructored in the middle of disambiguating a constructor.
Constructors
ConstructorDisambiguationData | |
Fields
|
Instances
NFData ConstructorDisambiguationData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: ConstructorDisambiguationData -> () # | |||||
Generic ConstructorDisambiguationData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
| |||||
type Rep ConstructorDisambiguationData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep ConstructorDisambiguationData = D1 ('MetaData "ConstructorDisambiguationData" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "ConstructorDisambiguationData" 'PrefixI 'True) ((S1 ('MetaSel ('Just "bcdConName") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Just "bcdCandidates") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 (QName, Type, ConHead)))) :*: (S1 ('MetaSel ('Just "bcdArguments") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Args) :*: S1 ('MetaSel ('Just "bcdType") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)))) |
type CurrentTypeCheckLog = [(TypeCheckAction, PostScopeState)] Source #
Like CachedTypeCheckLog
, but storing the log for an ongoing type
checking of a module. Stored in reverse order (last performed action
first).
data DataOrRecSigData Source #
Constructors
DataOrRecSigData | |
Fields
|
Instances
Pretty DataOrRecSigData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods pretty :: DataOrRecSigData -> Doc Source # prettyPrec :: Int -> DataOrRecSigData -> Doc Source # prettyList :: [DataOrRecSigData] -> Doc Source # | |||||
NFData DataOrRecSigData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: DataOrRecSigData -> () # | |||||
Generic DataOrRecSigData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
Methods from :: DataOrRecSigData -> Rep DataOrRecSigData x # to :: Rep DataOrRecSigData x -> DataOrRecSigData # | |||||
Show DataOrRecSigData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods showsPrec :: Int -> DataOrRecSigData -> ShowS # show :: DataOrRecSigData -> String # showList :: [DataOrRecSigData] -> ShowS # | |||||
type Rep DataOrRecSigData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep DataOrRecSigData = D1 ('MetaData "DataOrRecSigData" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "DataOrRecSigData" 'PrefixI 'True) (S1 ('MetaSel ('Just "_datarecPars") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Int))) |
data DatatypeData Source #
Constructors
DatatypeData | |
Fields
|
Instances
Pretty DatatypeData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods pretty :: DatatypeData -> Doc Source # prettyPrec :: Int -> DatatypeData -> Doc Source # prettyList :: [DatatypeData] -> Doc Source # | |||||
NFData DatatypeData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: DatatypeData -> () # | |||||
Generic DatatypeData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
| |||||
Show DatatypeData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods showsPrec :: Int -> DatatypeData -> ShowS # show :: DatatypeData -> String # showList :: [DatatypeData] -> ShowS # | |||||
type Rep DatatypeData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep DatatypeData = D1 ('MetaData "DatatypeData" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "DatatypeData" 'PrefixI 'True) (((S1 ('MetaSel ('Just "_dataPars") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Nat) :*: S1 ('MetaSel ('Just "_dataIxs") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Nat)) :*: (S1 ('MetaSel ('Just "_dataClause") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe Clause)) :*: (S1 ('MetaSel ('Just "_dataCons") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [QName]) :*: S1 ('MetaSel ('Just "_dataSort") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Sort)))) :*: ((S1 ('MetaSel ('Just "_dataMutual") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe [QName])) :*: S1 ('MetaSel ('Just "_dataAbstr") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 IsAbstract)) :*: (S1 ('MetaSel ('Just "_dataPathCons") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [QName]) :*: (S1 ('MetaSel ('Just "_dataTranspIx") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe QName)) :*: S1 ('MetaSel ('Just "_dataTransp") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe QName))))))) |
data Definition Source #
Constructors
Defn | |
Fields
|
Instances
LensArgInfo Definition Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods getArgInfo :: Definition -> ArgInfo Source # setArgInfo :: ArgInfo -> Definition -> Definition Source # mapArgInfo :: (ArgInfo -> ArgInfo) -> Definition -> Definition Source # | |||||
LensModalPolarity Definition Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods getModalPolarity :: Definition -> PolarityModality Source # setModalPolarity :: PolarityModality -> Definition -> Definition Source # mapModalPolarity :: (PolarityModality -> PolarityModality) -> Definition -> Definition Source # | |||||
LensModality Definition Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods getModality :: Definition -> Modality Source # setModality :: Modality -> Definition -> Definition Source # mapModality :: (Modality -> Modality) -> Definition -> Definition Source # | |||||
LensQuantity Definition Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods getQuantity :: Definition -> Quantity Source # setQuantity :: Quantity -> Definition -> Definition Source # mapQuantity :: (Quantity -> Quantity) -> Definition -> Definition Source # | |||||
LensRelevance Definition Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods getRelevance :: Definition -> Relevance Source # setRelevance :: Relevance -> Definition -> Definition Source # mapRelevance :: (Relevance -> Relevance) -> Definition -> Definition Source # | |||||
Pretty Definition Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods pretty :: Definition -> Doc Source # prettyPrec :: Int -> Definition -> Doc Source # prettyList :: [Definition] -> Doc Source # | |||||
NamesIn Definition Source # | |||||
Defined in Agda.Syntax.Internal.Names Methods namesAndMetasIn' :: Monoid m => (Either QName MetaId -> m) -> Definition -> m Source # | |||||
KillRange Definition Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods | |||||
KillRange Definitions Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods | |||||
InstantiateFull Definition Source # | |||||
Defined in Agda.TypeChecking.Reduce Methods instantiateFull' :: Definition -> ReduceM Definition Source # | |||||
EmbPrj Definition Source # | |||||
Abstract Definition Source # | |||||
Defined in Agda.TypeChecking.Substitute Methods abstract :: Telescope -> Definition -> Definition Source # | |||||
Apply Definition Source # | |||||
Defined in Agda.TypeChecking.Substitute Methods apply :: Definition -> Args -> Definition Source # applyE :: Definition -> Elims -> Definition Source # | |||||
NFData Definition Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: Definition -> () # | |||||
Generic Definition Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
| |||||
Show Definition Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods showsPrec :: Int -> Definition -> ShowS # show :: Definition -> String # showList :: [Definition] -> ShowS # | |||||
type Rep Definition Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep Definition = D1 ('MetaData "Definition" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "Defn" 'PrefixI 'True) ((((S1 ('MetaSel ('Just "defArgInfo") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ArgInfo) :*: S1 ('MetaSel ('Just "defName") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :*: (S1 ('MetaSel ('Just "defType") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type) :*: S1 ('MetaSel ('Just "defPolarity") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Polarity]))) :*: ((S1 ('MetaSel ('Just "defArgOccurrences") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Occurrence]) :*: S1 ('MetaSel ('Just "defGeneralizedParams") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Maybe Name])) :*: (S1 ('MetaSel ('Just "defDisplay") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [LocalDisplayForm]) :*: (S1 ('MetaSel ('Just "defMutual") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 MutualId) :*: S1 ('MetaSel ('Just "defCompiledRep") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 CompiledRepresentation))))) :*: (((S1 ('MetaSel ('Just "defInstance") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe InstanceInfo)) :*: S1 ('MetaSel ('Just "defCopy") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool)) :*: (S1 ('MetaSel ('Just "defMatchable") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Set QName)) :*: S1 ('MetaSel ('Just "defNoCompilation") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool))) :*: ((S1 ('MetaSel ('Just "defInjective") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool) :*: S1 ('MetaSel ('Just "defCopatternLHS") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool)) :*: (S1 ('MetaSel ('Just "defBlocked") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Blocked_) :*: (S1 ('MetaSel ('Just "defLanguage") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Language) :*: S1 ('MetaSel ('Just "theDef") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Defn))))))) |
type Definitions = HashMap QName Definition Source #
Constructors
AxiomDefn AxiomData | Postulate. |
DataOrRecSigDefn DataOrRecSigData | Data or record type signature that doesn't yet have a definition. |
GeneralizableVar | Generalizable variable (introduced in |
Fields
| |
AbstractDefn Defn | Returned by |
FunctionDefn FunctionData | |
DatatypeDefn DatatypeData | |
RecordDefn RecordData | |
ConstructorDefn ConstructorData | |
PrimitiveDefn PrimitiveData | Primitive or builtin functions. |
PrimitiveSortDefn PrimitiveSortData |
Instances
Pretty Defn Source # | |||||
NamesIn Defn Source # | |||||
Defined in Agda.Syntax.Internal.Names | |||||
KillRange Defn Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods | |||||
Occurs Defn Source # | |||||
InstantiateFull Defn Source # | |||||
Defined in Agda.TypeChecking.Reduce | |||||
EmbPrj Defn Source # | |||||
Abstract Defn Source # | |||||
Apply Defn Source # | |||||
NFData Defn Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Generic Defn Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
| |||||
Show Defn Source # | |||||
type Rep Defn Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep Defn = D1 ('MetaData "Defn" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (((C1 ('MetaCons "AxiomDefn" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 AxiomData)) :+: C1 ('MetaCons "DataOrRecSigDefn" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 DataOrRecSigData))) :+: (C1 ('MetaCons "GeneralizableVar" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 NumGeneralizableArgs)) :+: (C1 ('MetaCons "AbstractDefn" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Defn)) :+: C1 ('MetaCons "FunctionDefn" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 FunctionData))))) :+: ((C1 ('MetaCons "DatatypeDefn" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 DatatypeData)) :+: C1 ('MetaCons "RecordDefn" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 RecordData))) :+: (C1 ('MetaCons "ConstructorDefn" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ConstructorData)) :+: (C1 ('MetaCons "PrimitiveDefn" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 PrimitiveData)) :+: C1 ('MetaCons "PrimitiveSortDefn" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 PrimitiveSortData)))))) |
data DisambiguatedName Source #
Name disambiguation for the sake of highlighting.
Constructors
DisambiguatedName NameKind QName |
Instances
NFData DisambiguatedName Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: DisambiguatedName -> () # | |||||
Generic DisambiguatedName Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
Methods from :: DisambiguatedName -> Rep DisambiguatedName x # to :: Rep DisambiguatedName x -> DisambiguatedName # | |||||
type Rep DisambiguatedName Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep DisambiguatedName = D1 ('MetaData "DisambiguatedName" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "DisambiguatedName" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 NameKind) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName))) |
data DisplayForm Source #
A DisplayForm
is in essence a rewrite rule q ts --> dt
for a defined symbol (could be a
constructor as well) q
. The right hand side is a DisplayTerm
which is used to reify
to a
more readable Syntax
.
The patterns ts
are just terms, but the first dfPatternVars
variables are pattern variables
that matches any term.
Constructors
Display | |
Fields
|
Instances
Pretty DisplayForm Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods pretty :: DisplayForm -> Doc Source # prettyPrec :: Int -> DisplayForm -> Doc Source # prettyList :: [DisplayForm] -> Doc Source # | |||||
NamesIn DisplayForm Source # | |||||
Defined in Agda.Syntax.Internal.Names Methods namesAndMetasIn' :: Monoid m => (Either QName MetaId -> m) -> DisplayForm -> m Source # | |||||
KillRange DisplayForm Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods | |||||
Free DisplayForm Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
ChaseDisplayForms DisplayForm Source # | |||||
Defined in Agda.TypeChecking.Monad.Signature Methods chaseDisplayForms :: DisplayForm -> Set QName -> TCM (Set QName) Source # | |||||
InstantiateFull DisplayForm Source # | |||||
Defined in Agda.TypeChecking.Reduce Methods instantiateFull' :: DisplayForm -> ReduceM DisplayForm Source # | |||||
Normalise DisplayForm Source # | |||||
Defined in Agda.TypeChecking.Reduce Methods | |||||
Simplify DisplayForm Source # | |||||
Defined in Agda.TypeChecking.Reduce Methods | |||||
EmbPrj DisplayForm Source # | |||||
Subst DisplayForm Source # | |||||
Defined in Agda.TypeChecking.Substitute Associated Types
Methods applySubst :: Substitution' (SubstArg DisplayForm) -> DisplayForm -> DisplayForm Source # | |||||
NFData DisplayForm Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: DisplayForm -> () # | |||||
Generic DisplayForm Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
| |||||
Show DisplayForm Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods showsPrec :: Int -> DisplayForm -> ShowS # show :: DisplayForm -> String # showList :: [DisplayForm] -> ShowS # | |||||
type SubstArg DisplayForm Source # | |||||
Defined in Agda.TypeChecking.Substitute | |||||
type Rep DisplayForm Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep DisplayForm = D1 ('MetaData "DisplayForm" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "Display" 'PrefixI 'True) (S1 ('MetaSel ('Just "dfPatternVars") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Nat) :*: (S1 ('MetaSel ('Just "dfPats") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Elims) :*: S1 ('MetaSel ('Just "dfRHS") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 DisplayTerm)))) |
type DisplayForms = HashMap QName [LocalDisplayForm] Source #
data DisplayTerm Source #
Constructors
DWithApp DisplayTerm (List1 DisplayTerm) Elims |
|
DCon ConHead ConInfo [Arg DisplayTerm] |
|
DDef QName [Elim' DisplayTerm] |
|
DDot' Term Elims |
|
DTerm' Term Elims |
|
Instances
Pretty DisplayTerm Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods pretty :: DisplayTerm -> Doc Source # prettyPrec :: Int -> DisplayTerm -> Doc Source # prettyList :: [DisplayTerm] -> Doc Source # | |||||
NamesIn DisplayTerm Source # | |||||
Defined in Agda.Syntax.Internal.Names Methods namesAndMetasIn' :: Monoid m => (Either QName MetaId -> m) -> DisplayTerm -> m Source # | |||||
KillRange DisplayTerm Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods | |||||
Reify DisplayTerm Source # | |||||
Defined in Agda.Syntax.Translation.InternalToAbstract Associated Types
Methods reify :: MonadReify m => DisplayTerm -> m (ReifiesTo DisplayTerm) Source # reifyWhen :: MonadReify m => Bool -> DisplayTerm -> m (ReifiesTo DisplayTerm) Source # | |||||
Free DisplayTerm Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
ChaseDisplayForms DisplayTerm Source # | |||||
Defined in Agda.TypeChecking.Monad.Signature Methods chaseDisplayForms :: DisplayTerm -> Set QName -> TCM (Set QName) Source # | |||||
PrettyTCM DisplayTerm Source # | |||||
Defined in Agda.TypeChecking.Pretty Methods prettyTCM :: MonadPretty m => DisplayTerm -> m Doc Source # | |||||
InstantiateFull DisplayTerm Source # | |||||
Defined in Agda.TypeChecking.Reduce Methods instantiateFull' :: DisplayTerm -> ReduceM DisplayTerm Source # | |||||
EmbPrj DisplayTerm Source # | |||||
Apply DisplayTerm Source # | |||||
Defined in Agda.TypeChecking.Substitute Methods apply :: DisplayTerm -> Args -> DisplayTerm Source # applyE :: DisplayTerm -> Elims -> DisplayTerm Source # | |||||
Subst DisplayTerm Source # | |||||
Defined in Agda.TypeChecking.Substitute Associated Types
Methods applySubst :: Substitution' (SubstArg DisplayTerm) -> DisplayTerm -> DisplayTerm Source # | |||||
NFData DisplayTerm Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: DisplayTerm -> () # | |||||
Generic DisplayTerm Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
| |||||
Show DisplayTerm Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods showsPrec :: Int -> DisplayTerm -> ShowS # show :: DisplayTerm -> String # showList :: [DisplayTerm] -> ShowS # | |||||
PrettyTCM (Elim' DisplayTerm) Source # | |||||
Defined in Agda.TypeChecking.Pretty Methods prettyTCM :: MonadPretty m => Elim' DisplayTerm -> m Doc Source # | |||||
type ReifiesTo DisplayTerm Source # | |||||
Defined in Agda.Syntax.Translation.InternalToAbstract | |||||
type SubstArg DisplayTerm Source # | |||||
Defined in Agda.TypeChecking.Substitute | |||||
type Rep DisplayTerm Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep DisplayTerm = D1 ('MetaData "DisplayTerm" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) ((C1 ('MetaCons "DWithApp" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 DisplayTerm) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 DisplayTerm)) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Elims))) :+: C1 ('MetaCons "DCon" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ConHead) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ConInfo) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Arg DisplayTerm])))) :+: (C1 ('MetaCons "DDef" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Elim' DisplayTerm])) :+: (C1 ('MetaCons "DDot'" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Elims)) :+: C1 ('MetaCons "DTerm'" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Elims))))) |
data DoGeneralize Source #
Constructors
YesGeneralizeVar | Generalize because it is a generalizable variable. |
YesGeneralizeMeta | Generalize because it is a metavariable and we're currently checking the type of a generalizable variable (this should get the default modality). |
NoGeneralize | Don't generalize. |
Instances
KillRange DoGeneralize Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods | |||||
EmbPrj DoGeneralize Source # | |||||
NFData DoGeneralize Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: DoGeneralize -> () # | |||||
Generic DoGeneralize Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
| |||||
Show DoGeneralize Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods showsPrec :: Int -> DoGeneralize -> ShowS # show :: DoGeneralize -> String # showList :: [DoGeneralize] -> ShowS # | |||||
Eq DoGeneralize Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Ord DoGeneralize Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods compare :: DoGeneralize -> DoGeneralize -> Ordering # (<) :: DoGeneralize -> DoGeneralize -> Bool # (<=) :: DoGeneralize -> DoGeneralize -> Bool # (>) :: DoGeneralize -> DoGeneralize -> Bool # (>=) :: DoGeneralize -> DoGeneralize -> Bool # max :: DoGeneralize -> DoGeneralize -> DoGeneralize # min :: DoGeneralize -> DoGeneralize -> DoGeneralize # | |||||
type Rep DoGeneralize Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep DoGeneralize = D1 ('MetaData "DoGeneralize" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "YesGeneralizeVar" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "YesGeneralizeMeta" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "NoGeneralize" 'PrefixI 'False) (U1 :: Type -> Type))) |
data EtaEquality Source #
Should a record type admit eta-equality?
Constructors
Specified | User specifed 'eta-equality' or 'no-eta-equality'. |
Fields
| |
Inferred | Positivity checker inferred whether eta is safe. |
Fields
|
Instances
CopatternMatchingAllowed EtaEquality Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods | |||||
PatternMatchingAllowed EtaEquality Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods | |||||
KillRange EtaEquality Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods | |||||
EmbPrj EtaEquality Source # | |||||
NFData EtaEquality Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: EtaEquality -> () # | |||||
Generic EtaEquality Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
| |||||
Show EtaEquality Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods showsPrec :: Int -> EtaEquality -> ShowS # show :: EtaEquality -> String # showList :: [EtaEquality] -> ShowS # | |||||
Eq EtaEquality Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
type Rep EtaEquality Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep EtaEquality = D1 ('MetaData "EtaEquality" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "Specified" 'PrefixI 'True) (S1 ('MetaSel ('Just "theEtaEquality") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 HasEta)) :+: C1 ('MetaCons "Inferred" 'PrefixI 'True) (S1 ('MetaSel ('Just "theEtaEquality") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 HasEta))) |
Error when trying to call an external executable during reflection.
Constructors
ExeNotTrusted ExeName ExeMap | The given executable is not listed as trusted. |
ExeNotFound ExeName FilePath | The given executable could not be found under the given path. |
ExeNotExecutable ExeName FilePath | The given file path does not have executable permissions. |
Instances
PrettyTCM ExecError Source # | |||||
Defined in Agda.TypeChecking.Errors | |||||
NFData ExecError Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Generic ExecError Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
| |||||
Show ExecError Source # | |||||
type Rep ExecError Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep ExecError = D1 ('MetaData "ExecError" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "ExeNotTrusted" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ExeName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ExeMap)) :+: (C1 ('MetaCons "ExeNotFound" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ExeName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 FilePath)) :+: C1 ('MetaCons "ExeNotExecutable" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ExeName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 FilePath)))) |
data ExpandHidden Source #
Constructors
ExpandLast | Add implicit arguments in the end until type is no longer hidden |
DontExpandLast | Do not append implicit arguments. |
ReallyDontExpandLast | Makes |
Instances
NFData ExpandHidden Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: ExpandHidden -> () # | |||||
Generic ExpandHidden Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
| |||||
Eq ExpandHidden Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
type Rep ExpandHidden Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep ExpandHidden = D1 ('MetaData "ExpandHidden" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "ExpandLast" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "DontExpandLast" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "ReallyDontExpandLast" 'PrefixI 'False) (U1 :: Type -> Type))) |
data ExtLamInfo Source #
Additional information for extended lambdas.
Constructors
ExtLamInfo | |
Fields
|
Instances
NamesIn ExtLamInfo Source # | |||||
Defined in Agda.Syntax.Internal.Names Methods namesAndMetasIn' :: Monoid m => (Either QName MetaId -> m) -> ExtLamInfo -> m Source # | |||||
KillRange ExtLamInfo Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods | |||||
InstantiateFull ExtLamInfo Source # | |||||
Defined in Agda.TypeChecking.Reduce Methods instantiateFull' :: ExtLamInfo -> ReduceM ExtLamInfo Source # | |||||
EmbPrj ExtLamInfo Source # | |||||
Apply ExtLamInfo Source # | |||||
Defined in Agda.TypeChecking.Substitute Methods apply :: ExtLamInfo -> Args -> ExtLamInfo Source # applyE :: ExtLamInfo -> Elims -> ExtLamInfo Source # | |||||
NFData ExtLamInfo Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: ExtLamInfo -> () # | |||||
Generic ExtLamInfo Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
| |||||
Show ExtLamInfo Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods showsPrec :: Int -> ExtLamInfo -> ShowS # show :: ExtLamInfo -> String # showList :: [ExtLamInfo] -> ShowS # | |||||
type Rep ExtLamInfo Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep ExtLamInfo = D1 ('MetaData "ExtLamInfo" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "ExtLamInfo" 'PrefixI 'True) (S1 ('MetaSel ('Just "extLamModule") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ModuleName) :*: (S1 ('MetaSel ('Just "extLamAbsurd") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool) :*: S1 ('MetaSel ('Just "extLamSys") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (Maybe System))))) |
data ForeignCode Source #
Constructors
ForeignCode Range String |
Instances
EmbPrj ForeignCode Source # | |||||
NFData ForeignCode Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: ForeignCode -> () # | |||||
Generic ForeignCode Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
| |||||
Show ForeignCode Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods showsPrec :: Int -> ForeignCode -> ShowS # show :: ForeignCode -> String # showList :: [ForeignCode] -> ShowS # | |||||
type Rep ForeignCode Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep ForeignCode = D1 ('MetaData "ForeignCode" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "ForeignCode" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String))) |
newtype ForeignCodeStack Source #
Foreign code fragments are stored in reversed order to support efficient appending: head points to the latest pragma in module.
Constructors
ForeignCodeStack | |
Fields |
Instances
EmbPrj ForeignCodeStack Source # | |||||
NFData ForeignCodeStack Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: ForeignCodeStack -> () # | |||||
Generic ForeignCodeStack Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
Methods from :: ForeignCodeStack -> Rep ForeignCodeStack x # to :: Rep ForeignCodeStack x -> ForeignCodeStack # | |||||
Show ForeignCodeStack Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods showsPrec :: Int -> ForeignCodeStack -> ShowS # show :: ForeignCodeStack -> String # showList :: [ForeignCodeStack] -> ShowS # | |||||
type Rep ForeignCodeStack Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep ForeignCodeStack = D1 ('MetaData "ForeignCodeStack" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'True) (C1 ('MetaCons "ForeignCodeStack" 'PrefixI 'True) (S1 ('MetaSel ('Just "getForeignCodeStack") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [ForeignCode]))) |
class FreshName a where Source #
Create a fresh name from a
.
Methods
freshName_ :: MonadFresh NameId m => a -> m Name Source #
Instances
FreshName Name Source # | |
Defined in Agda.TypeChecking.Monad.Base Methods freshName_ :: MonadFresh NameId m => Name -> m Name Source # | |
FreshName Range Source # | |
Defined in Agda.TypeChecking.Monad.Base Methods freshName_ :: MonadFresh NameId m => Range -> m Name Source # | |
FreshName String Source # | |
Defined in Agda.TypeChecking.Monad.Base Methods freshName_ :: MonadFresh NameId m => String -> m Name Source # | |
FreshName () Source # | |
Defined in Agda.TypeChecking.Monad.Base Methods freshName_ :: MonadFresh NameId m => () -> m Name Source # | |
FreshName (Range, String) Source # | |
Defined in Agda.TypeChecking.Monad.Base Methods freshName_ :: MonadFresh NameId m => (Range, String) -> m Name Source # |
Frozen meta variable cannot be instantiated by unification. This serves to prevent the completion of a definition by its use outside of the current block. (See issues 118, 288, 399).
Constructors
Frozen | Do not instantiate. |
Instantiable |
data FunctionData Source #
Constructors
FunctionData | |
Fields
|
Instances
Pretty FunctionData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods pretty :: FunctionData -> Doc Source # prettyPrec :: Int -> FunctionData -> Doc Source # prettyList :: [FunctionData] -> Doc Source # | |||||
NFData FunctionData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: FunctionData -> () # | |||||
Generic FunctionData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
| |||||
Show FunctionData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods showsPrec :: Int -> FunctionData -> ShowS # show :: FunctionData -> String # showList :: [FunctionData] -> ShowS # | |||||
type Rep FunctionData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep FunctionData = D1 ('MetaData "FunctionData" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "FunctionData" 'PrefixI 'True) (((S1 ('MetaSel ('Just "_funClauses") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Clause]) :*: (S1 ('MetaSel ('Just "_funCompiled") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe CompiledClauses)) :*: S1 ('MetaSel ('Just "_funSplitTree") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe SplitTree)))) :*: ((S1 ('MetaSel ('Just "_funTreeless") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe Compiled)) :*: S1 ('MetaSel ('Just "_funCovering") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Clause])) :*: (S1 ('MetaSel ('Just "_funInv") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 FunctionInverse) :*: S1 ('MetaSel ('Just "_funMutual") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe [QName]))))) :*: ((S1 ('MetaSel ('Just "_funProjection") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Either ProjectionLikenessMissing Projection)) :*: (S1 ('MetaSel ('Just "_funFlags") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (SmallSet FunctionFlag)) :*: S1 ('MetaSel ('Just "_funTerminates") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe Bool)))) :*: ((S1 ('MetaSel ('Just "_funExtLam") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe ExtLamInfo)) :*: S1 ('MetaSel ('Just "_funWith") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe QName))) :*: (S1 ('MetaSel ('Just "_funIsKanOp") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe QName)) :*: S1 ('MetaSel ('Just "_funOpaque") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 IsOpaque)))))) |
data FunctionFlag Source #
Constructors
FunStatic | Should calls to this function be normalised at compile-time? |
FunInline | Should calls to this function be inlined by the compiler? |
FunMacro | Is this function a macro? |
FunFirstOrder | Is this function |
FunErasure | Was |
FunAbstract | Is the function abstract? |
FunProj | Is this function a descendant of a field (typically, a projection)? |
Instances
KillRange FunctionFlag Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods | |||||
EmbPrj FunctionFlag Source # | |||||
SmallSetElement FunctionFlag Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
NFData FunctionFlag Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: FunctionFlag -> () # | |||||
Bounded FunctionFlag Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Enum FunctionFlag Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods succ :: FunctionFlag -> FunctionFlag # pred :: FunctionFlag -> FunctionFlag # toEnum :: Int -> FunctionFlag # fromEnum :: FunctionFlag -> Int # enumFrom :: FunctionFlag -> [FunctionFlag] # enumFromThen :: FunctionFlag -> FunctionFlag -> [FunctionFlag] # enumFromTo :: FunctionFlag -> FunctionFlag -> [FunctionFlag] # enumFromThenTo :: FunctionFlag -> FunctionFlag -> FunctionFlag -> [FunctionFlag] # | |||||
Generic FunctionFlag Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
| |||||
Ix FunctionFlag Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods range :: (FunctionFlag, FunctionFlag) -> [FunctionFlag] # index :: (FunctionFlag, FunctionFlag) -> FunctionFlag -> Int # unsafeIndex :: (FunctionFlag, FunctionFlag) -> FunctionFlag -> Int # inRange :: (FunctionFlag, FunctionFlag) -> FunctionFlag -> Bool # rangeSize :: (FunctionFlag, FunctionFlag) -> Int # unsafeRangeSize :: (FunctionFlag, FunctionFlag) -> Int # | |||||
Show FunctionFlag Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods showsPrec :: Int -> FunctionFlag -> ShowS # show :: FunctionFlag -> String # showList :: [FunctionFlag] -> ShowS # | |||||
Eq FunctionFlag Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Ord FunctionFlag Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods compare :: FunctionFlag -> FunctionFlag -> Ordering # (<) :: FunctionFlag -> FunctionFlag -> Bool # (<=) :: FunctionFlag -> FunctionFlag -> Bool # (>) :: FunctionFlag -> FunctionFlag -> Bool # (>=) :: FunctionFlag -> FunctionFlag -> Bool # max :: FunctionFlag -> FunctionFlag -> FunctionFlag # min :: FunctionFlag -> FunctionFlag -> FunctionFlag # | |||||
KillRange (SmallSet FunctionFlag) Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods | |||||
type Rep FunctionFlag Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep FunctionFlag = D1 ('MetaData "FunctionFlag" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) ((C1 ('MetaCons "FunStatic" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "FunInline" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "FunMacro" 'PrefixI 'False) (U1 :: Type -> Type))) :+: ((C1 ('MetaCons "FunFirstOrder" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "FunErasure" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "FunAbstract" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "FunProj" 'PrefixI 'False) (U1 :: Type -> Type)))) |
type FunctionInverse = FunctionInverse' Clause Source #
data FunctionInverse' c Source #
Constructors
NotInjective | |
Inverse (InversionMap c) |
Instances
DropArgs FunctionInverse Source # | |||||
Defined in Agda.TypeChecking.DropArgs Methods dropArgs :: Int -> FunctionInverse -> FunctionInverse Source # | |||||
InstantiateFull FunctionInverse Source # | |||||
Defined in Agda.TypeChecking.Reduce Methods instantiateFull' :: FunctionInverse -> ReduceM FunctionInverse Source # | |||||
Abstract FunctionInverse Source # | |||||
Defined in Agda.TypeChecking.Substitute Methods abstract :: Telescope -> FunctionInverse -> FunctionInverse Source # | |||||
Apply FunctionInverse Source # | |||||
Defined in Agda.TypeChecking.Substitute Methods apply :: FunctionInverse -> Args -> FunctionInverse Source # applyE :: FunctionInverse -> Elims -> FunctionInverse Source # | |||||
Functor FunctionInverse' Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods fmap :: (a -> b) -> FunctionInverse' a -> FunctionInverse' b # (<$) :: a -> FunctionInverse' b -> FunctionInverse' a # | |||||
Pretty c => Pretty (FunctionInverse' c) Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods pretty :: FunctionInverse' c -> Doc Source # prettyPrec :: Int -> FunctionInverse' c -> Doc Source # prettyList :: [FunctionInverse' c] -> Doc Source # | |||||
NamesIn a => NamesIn (FunctionInverse' a) Source # | |||||
Defined in Agda.Syntax.Internal.Names Methods namesAndMetasIn' :: Monoid m => (Either QName MetaId -> m) -> FunctionInverse' a -> m Source # | |||||
KillRange c => KillRange (FunctionInverse' c) Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods killRange :: KillRangeT (FunctionInverse' c) Source # | |||||
EmbPrj a => EmbPrj (FunctionInverse' a) Source # | |||||
NFData c => NFData (FunctionInverse' c) Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: FunctionInverse' c -> () # | |||||
Generic (FunctionInverse' c) Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
Methods from :: FunctionInverse' c -> Rep (FunctionInverse' c) x # to :: Rep (FunctionInverse' c) x -> FunctionInverse' c # | |||||
Show c => Show (FunctionInverse' c) Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods showsPrec :: Int -> FunctionInverse' c -> ShowS # show :: FunctionInverse' c -> String # showList :: [FunctionInverse' c] -> ShowS # | |||||
type Rep (FunctionInverse' c) Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep (FunctionInverse' c) = D1 ('MetaData "FunctionInverse'" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "NotInjective" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "Inverse" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (InversionMap c)))) |
data GHCBackendError Source #
Errors raised by the GHC backend.
Constructors
ConstructorCountMismatch QName [QName] [String] | The number of Haskell constructors ( |
NotAHaskellType Term WhyNotAHaskellType | GHC backend fails to represent given Agda type in Haskell. |
WrongTypeOfMain QName Type | The type of |
Instances
PrettyTCM GHCBackendError Source # | |||||
Defined in Agda.TypeChecking.Errors Methods prettyTCM :: MonadPretty m => GHCBackendError -> m Doc Source # | |||||
NFData GHCBackendError Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: GHCBackendError -> () # | |||||
Generic GHCBackendError Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
Methods from :: GHCBackendError -> Rep GHCBackendError x # to :: Rep GHCBackendError x -> GHCBackendError # | |||||
Show GHCBackendError Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods showsPrec :: Int -> GHCBackendError -> ShowS # show :: GHCBackendError -> String # showList :: [GHCBackendError] -> ShowS # | |||||
type Rep GHCBackendError Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep GHCBackendError = D1 ('MetaData "GHCBackendError" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "ConstructorCountMismatch" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [QName]) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [String]))) :+: (C1 ('MetaCons "NotAHaskellType" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 WhyNotAHaskellType)) :+: C1 ('MetaCons "WrongTypeOfMain" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)))) |
data GeneralizedValue Source #
The value of a generalizable variable. This is created to be a generalizable meta before checking the type to be generalized.
Constructors
GeneralizedValue | |
Fields
|
Instances
NFData GeneralizedValue Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: GeneralizedValue -> () # | |||||
Generic GeneralizedValue Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
Methods from :: GeneralizedValue -> Rep GeneralizedValue x # to :: Rep GeneralizedValue x -> GeneralizedValue # | |||||
Show GeneralizedValue Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods showsPrec :: Int -> GeneralizedValue -> ShowS # show :: GeneralizedValue -> String # showList :: [GeneralizedValue] -> ShowS # | |||||
type Rep GeneralizedValue Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep GeneralizedValue = D1 ('MetaData "GeneralizedValue" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "GeneralizedValue" 'PrefixI 'True) (S1 ('MetaSel ('Just "genvalCheckpoint") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 CheckpointId) :*: (S1 ('MetaSel ('Just "genvalTerm") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Just "genvalType") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)))) |
class Enum i => HasFresh i where Source #
Minimal complete definition
Instances
HasFresh InteractionId Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
HasFresh MetaId Source # | |
HasFresh NameId Source # | |
HasFresh OpaqueId Source # | |
HasFresh ProblemId Source # | |
HasFresh CheckpointId Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
HasFresh MutualId Source # | |
HasFresh Int Source # | |
type IPBoundary = IPBoundary' Term Source #
newtype IPBoundary' t Source #
Constructors
IPBoundary | |
Fields
|
Instances
Functor IPBoundary' Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods fmap :: (a -> b) -> IPBoundary' a -> IPBoundary' b # (<$) :: a -> IPBoundary' b -> IPBoundary' a # | |||||
Foldable IPBoundary' Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods fold :: Monoid m => IPBoundary' m -> m # foldMap :: Monoid m => (a -> m) -> IPBoundary' a -> m # foldMap' :: Monoid m => (a -> m) -> IPBoundary' a -> m # foldr :: (a -> b -> b) -> b -> IPBoundary' a -> b # foldr' :: (a -> b -> b) -> b -> IPBoundary' a -> b # foldl :: (b -> a -> b) -> b -> IPBoundary' a -> b # foldl' :: (b -> a -> b) -> b -> IPBoundary' a -> b # foldr1 :: (a -> a -> a) -> IPBoundary' a -> a # foldl1 :: (a -> a -> a) -> IPBoundary' a -> a # toList :: IPBoundary' a -> [a] # null :: IPBoundary' a -> Bool # length :: IPBoundary' a -> Int # elem :: Eq a => a -> IPBoundary' a -> Bool # maximum :: Ord a => IPBoundary' a -> a # minimum :: Ord a => IPBoundary' a -> a # sum :: Num a => IPBoundary' a -> a # product :: Num a => IPBoundary' a -> a # | |||||
Traversable IPBoundary' Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods traverse :: Applicative f => (a -> f b) -> IPBoundary' a -> f (IPBoundary' b) # sequenceA :: Applicative f => IPBoundary' (f a) -> f (IPBoundary' a) # mapM :: Monad m => (a -> m b) -> IPBoundary' a -> m (IPBoundary' b) # sequence :: Monad m => IPBoundary' (m a) -> m (IPBoundary' a) # | |||||
ToConcrete a => ToConcrete (IPBoundary' a) Source # | |||||
Defined in Agda.Interaction.BasicOps Associated Types
Methods toConcrete :: MonadToConcrete m => IPBoundary' a -> m (ConOfAbs (IPBoundary' a)) Source # bindToConcrete :: MonadToConcrete m => IPBoundary' a -> (ConOfAbs (IPBoundary' a) -> m b) -> m b Source # | |||||
Reify a => Reify (IPBoundary' a) Source # | |||||
Defined in Agda.Interaction.BasicOps Associated Types
Methods reify :: MonadReify m => IPBoundary' a -> m (ReifiesTo (IPBoundary' a)) Source # reifyWhen :: MonadReify m => Bool -> IPBoundary' a -> m (ReifiesTo (IPBoundary' a)) Source # | |||||
Instantiate t => Instantiate (IPBoundary' t) Source # | |||||
Defined in Agda.TypeChecking.Reduce Methods instantiate' :: IPBoundary' t -> ReduceM (IPBoundary' t) Source # | |||||
InstantiateFull t => InstantiateFull (IPBoundary' t) Source # | |||||
Defined in Agda.TypeChecking.Reduce Methods instantiateFull' :: IPBoundary' t -> ReduceM (IPBoundary' t) Source # | |||||
Normalise t => Normalise (IPBoundary' t) Source # | |||||
Defined in Agda.TypeChecking.Reduce Methods normalise' :: IPBoundary' t -> ReduceM (IPBoundary' t) Source # | |||||
Reduce t => Reduce (IPBoundary' t) Source # | |||||
Defined in Agda.TypeChecking.Reduce Methods reduce' :: IPBoundary' t -> ReduceM (IPBoundary' t) Source # reduceB' :: IPBoundary' t -> ReduceM (Blocked (IPBoundary' t)) Source # | |||||
Simplify t => Simplify (IPBoundary' t) Source # | |||||
Defined in Agda.TypeChecking.Reduce Methods simplify' :: IPBoundary' t -> ReduceM (IPBoundary' t) | |||||
NFData t => NFData (IPBoundary' t) Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: IPBoundary' t -> () # | |||||
Generic (IPBoundary' t) Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
Methods from :: IPBoundary' t -> Rep (IPBoundary' t) x # to :: Rep (IPBoundary' t) x -> IPBoundary' t # | |||||
Show t => Show (IPBoundary' t) Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods showsPrec :: Int -> IPBoundary' t -> ShowS # show :: IPBoundary' t -> String # showList :: [IPBoundary' t] -> ShowS # | |||||
type ConOfAbs (IPBoundary' a) Source # | |||||
Defined in Agda.Interaction.BasicOps | |||||
type ReifiesTo (IPBoundary' a) Source # | |||||
Defined in Agda.Interaction.BasicOps | |||||
type Rep (IPBoundary' t) Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep (IPBoundary' t) = D1 ('MetaData "IPBoundary'" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'True) (C1 ('MetaCons "IPBoundary" 'PrefixI 'True) (S1 ('MetaSel ('Just "getBoundary") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Map (IntMap Bool) t)))) |
Which clause is an interaction point located in?
Constructors
IPClause | |
Fields
| |
IPNoClause | The interaction point is not in the rhs of a clause. |
Instances
NFData IPClause Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Generic IPClause Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
| |||||
Eq IPClause Source # | |||||
type Rep IPClause Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep IPClause = D1 ('MetaData "IPClause" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "IPClause" 'PrefixI 'True) ((S1 ('MetaSel ('Just "ipcQName") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: (S1 ('MetaSel ('Just "ipcClauseNo") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Int) :*: S1 ('MetaSel ('Just "ipcType") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type))) :*: (S1 ('MetaSel ('Just "ipcWithSub") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe Substitution)) :*: (S1 ('MetaSel ('Just "ipcClause") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 SpineClause) :*: S1 ('MetaSel ('Just "ipcClosure") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Closure ()))))) :+: C1 ('MetaCons "IPNoClause" 'PrefixI 'False) (U1 :: Type -> Type)) |
data IllegalRewriteRuleReason Source #
Constructors
LHSNotDefinitionOrConstructor | |
VariablesNotBoundByLHS IntSet | |
VariablesBoundMoreThanOnce IntSet | |
LHSReduces Term Term | |
HeadSymbolIsProjectionLikeFunction QName | |
HeadSymbolIsTypeConstructor QName | |
HeadSymbolContainsMetas QName | |
ConstructorParametersNotGeneral ConHead Args | |
ContainsUnsolvedMetaVariables (Set1 MetaId) | |
BlockedOnProblems (Set1 ProblemId) | |
RequiresDefinitions (Set1 QName) | |
DoesNotTargetRewriteRelation | |
BeforeFunctionDefinition | |
BeforeMutualFunctionDefinition QName | |
DuplicateRewriteRule |
Instances
EmbPrj IllegalRewriteRuleReason Source # | |||||
Defined in Agda.TypeChecking.Serialise.Instances.Errors | |||||
NFData IllegalRewriteRuleReason Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: IllegalRewriteRuleReason -> () # | |||||
Generic IllegalRewriteRuleReason Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
Methods from :: IllegalRewriteRuleReason -> Rep IllegalRewriteRuleReason x # to :: Rep IllegalRewriteRuleReason x -> IllegalRewriteRuleReason # | |||||
Show IllegalRewriteRuleReason Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods showsPrec :: Int -> IllegalRewriteRuleReason -> ShowS # show :: IllegalRewriteRuleReason -> String # showList :: [IllegalRewriteRuleReason] -> ShowS # | |||||
type Rep IllegalRewriteRuleReason Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep IllegalRewriteRuleReason = D1 ('MetaData "IllegalRewriteRuleReason" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (((C1 ('MetaCons "LHSNotDefinitionOrConstructor" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "VariablesNotBoundByLHS" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 IntSet)) :+: C1 ('MetaCons "VariablesBoundMoreThanOnce" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 IntSet)))) :+: ((C1 ('MetaCons "LHSReduces" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)) :+: C1 ('MetaCons "HeadSymbolIsProjectionLikeFunction" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName))) :+: (C1 ('MetaCons "HeadSymbolIsTypeConstructor" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "HeadSymbolContainsMetas" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName))))) :+: (((C1 ('MetaCons "ConstructorParametersNotGeneral" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ConHead) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Args)) :+: C1 ('MetaCons "ContainsUnsolvedMetaVariables" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Set1 MetaId)))) :+: (C1 ('MetaCons "BlockedOnProblems" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Set1 ProblemId))) :+: C1 ('MetaCons "RequiresDefinitions" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Set1 QName))))) :+: ((C1 ('MetaCons "DoesNotTargetRewriteRelation" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BeforeFunctionDefinition" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "BeforeMutualFunctionDefinition" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "DuplicateRewriteRule" 'PrefixI 'False) (U1 :: Type -> Type))))) |
type ImportedModules = Set TopLevelModuleName Source #
data IncorrectTypeForRewriteRelationReason Source #
Constructors
ShouldAcceptAtLeastTwoArguments | |
FinalTwoArgumentsNotVisible | |
TypeDoesNotEndInSort Type Telescope |
Instances
NFData IncorrectTypeForRewriteRelationReason Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: IncorrectTypeForRewriteRelationReason -> () # | |||||
Generic IncorrectTypeForRewriteRelationReason Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
| |||||
Show IncorrectTypeForRewriteRelationReason Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods showsPrec :: Int -> IncorrectTypeForRewriteRelationReason -> ShowS # show :: IncorrectTypeForRewriteRelationReason -> String # showList :: [IncorrectTypeForRewriteRelationReason] -> ShowS # | |||||
type Rep IncorrectTypeForRewriteRelationReason Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep IncorrectTypeForRewriteRelationReason = D1 ('MetaData "IncorrectTypeForRewriteRelationReason" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "ShouldAcceptAtLeastTwoArguments" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "FinalTwoArgumentsNotVisible" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "TypeDoesNotEndInSort" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Telescope)))) |
data InductionAndEta Source #
Constructors
InductionAndEta | |
Fields |
Instances
NFData InductionAndEta Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: InductionAndEta -> () # | |||||
Generic InductionAndEta Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
Methods from :: InductionAndEta -> Rep InductionAndEta x # to :: Rep InductionAndEta x -> InductionAndEta # | |||||
Show InductionAndEta Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods showsPrec :: Int -> InductionAndEta -> ShowS # show :: InductionAndEta -> String # showList :: [InductionAndEta] -> ShowS # | |||||
type Rep InductionAndEta Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep InductionAndEta = D1 ('MetaData "InductionAndEta" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "InductionAndEta" 'PrefixI 'True) (S1 ('MetaSel ('Just "recordInduction") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe Induction)) :*: S1 ('MetaSel ('Just "recordEtaEquality") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 EtaEquality))) |
data InstanceInfo Source #
Information about an instance
definition.
Constructors
InstanceInfo | |
Fields
|
Instances
KillRange InstanceInfo Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods | |||||
EmbPrj InstanceInfo Source # | |||||
NFData InstanceInfo Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: InstanceInfo -> () # | |||||
Generic InstanceInfo Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
| |||||
Show InstanceInfo Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods showsPrec :: Int -> InstanceInfo -> ShowS # show :: InstanceInfo -> String # showList :: [InstanceInfo] -> ShowS # | |||||
type Rep InstanceInfo Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep InstanceInfo = D1 ('MetaData "InstanceInfo" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "InstanceInfo" 'PrefixI 'True) (S1 ('MetaSel ('Just "instanceClass") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Just "instanceOverlap") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 OverlapMode))) |
data InstanceTable Source #
Records information about the instances in the signature. Does not deal with local instances.
Constructors
InstanceTable | |
Fields
|
Instances
KillRange InstanceTable Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods | |||||
EmbPrj InstanceTable Source # | |||||
NFData InstanceTable Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: InstanceTable -> () # | |||||
Monoid InstanceTable Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods mempty :: InstanceTable # mappend :: InstanceTable -> InstanceTable -> InstanceTable # mconcat :: [InstanceTable] -> InstanceTable # | |||||
Semigroup InstanceTable Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods (<>) :: InstanceTable -> InstanceTable -> InstanceTable # sconcat :: NonEmpty InstanceTable -> InstanceTable # stimes :: Integral b => b -> InstanceTable -> InstanceTable # | |||||
Generic InstanceTable Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
| |||||
Show InstanceTable Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods showsPrec :: Int -> InstanceTable -> ShowS # show :: InstanceTable -> String # showList :: [InstanceTable] -> ShowS # | |||||
type Rep InstanceTable Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep InstanceTable = D1 ('MetaData "InstanceTable" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "InstanceTable" 'PrefixI 'True) (S1 ('MetaSel ('Just "_itableTree") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (DiscrimTree QName)) :*: S1 ('MetaSel ('Just "_itableCounts") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Map QName Int)))) |
data Instantiation Source #
Meta-variable instantiations.
Constructors
Instantiation | |
Instances
InstantiateFull Instantiation Source # | |||||
Defined in Agda.TypeChecking.Reduce Methods instantiateFull' :: Instantiation -> ReduceM Instantiation Source # | |||||
EmbPrj Instantiation Source # | |||||
NFData Instantiation Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: Instantiation -> () # | |||||
Generic Instantiation Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
| |||||
Show Instantiation Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods showsPrec :: Int -> Instantiation -> ShowS # show :: Instantiation -> String # showList :: [Instantiation] -> ShowS # | |||||
type Rep Instantiation Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep Instantiation = D1 ('MetaData "Instantiation" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "Instantiation" 'PrefixI 'True) (S1 ('MetaSel ('Just "instTel") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Arg String]) :*: S1 ('MetaSel ('Just "instBody") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term))) |
data InteractionError Source #
Errors raised in --interaction
mode.
Constructors
CannotGive Expr | Failure of the |
CannotRefine String | Failure of the |
CaseSplitError Doc | Failure of the |
ExpectedIdentifier Expr | Expected the given expression to be an identifier. |
ExpectedApplication | Expected an argument of the form |
NoActionForInteractionPoint InteractionId | Interaction point has not been reached during type checking. |
NoSuchInteractionPoint InteractionId |
|
UnexpectedWhere |
|
Instances
PrettyTCM InteractionError Source # | |||||
Defined in Agda.TypeChecking.Errors Methods prettyTCM :: MonadPretty m => InteractionError -> m Doc Source # | |||||
NFData InteractionError Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: InteractionError -> () # | |||||
Generic InteractionError Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
Methods from :: InteractionError -> Rep InteractionError x # to :: Rep InteractionError x -> InteractionError # | |||||
Show InteractionError Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods showsPrec :: Int -> InteractionError -> ShowS # show :: InteractionError -> String # showList :: [InteractionError] -> ShowS # | |||||
type Rep InteractionError Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep InteractionError = D1 ('MetaData "InteractionError" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (((C1 ('MetaCons "CannotGive" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr)) :+: C1 ('MetaCons "CannotRefine" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String))) :+: (C1 ('MetaCons "CaseSplitError" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Doc)) :+: C1 ('MetaCons "ExpectedIdentifier" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr)))) :+: ((C1 ('MetaCons "ExpectedApplication" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "NoActionForInteractionPoint" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 InteractionId))) :+: (C1 ('MetaCons "NoSuchInteractionPoint" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 InteractionId)) :+: C1 ('MetaCons "UnexpectedWhere" 'PrefixI 'False) (U1 :: Type -> Type)))) |
type InteractionOutputCallback = Response_boot TCErr TCWarning WarningsAndNonFatalErrors -> TCM () Source #
Callback fuction to call when there is a response to give to the interactive frontend.
Note that the response is given in pieces and incrementally, so the user can have timely response even during long computations.
Typical InteractionOutputCallback
functions:
- Convert the response into a
String
representation and print it on standard output (suitable for inter-process communication). - Put the response into a mutable variable stored in the
closure of the
InteractionOutputCallback
function. (suitable for intra-process communication).
data InteractionPoint Source #
Interaction points are created by the scope checker who sets the range. The meta variable is created by the type checker and then hooked up to the interaction point.
Constructors
InteractionPoint | |
Fields
|
Instances
HasTag InteractionPoint Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
Methods tag :: InteractionPoint -> Maybe (Tag InteractionPoint) Source # | |||||
NFData InteractionPoint Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: InteractionPoint -> () # | |||||
NFData InteractionPoints Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: InteractionPoints -> () # | |||||
Generic InteractionPoint Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
Methods from :: InteractionPoint -> Rep InteractionPoint x # to :: Rep InteractionPoint x -> InteractionPoint # | |||||
Eq InteractionPoint Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods (==) :: InteractionPoint -> InteractionPoint -> Bool # (/=) :: InteractionPoint -> InteractionPoint -> Bool # | |||||
type Tag InteractionPoint Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
type Rep InteractionPoint Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep InteractionPoint = D1 ('MetaData "InteractionPoint" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "InteractionPoint" 'PrefixI 'True) ((S1 ('MetaSel ('Just "ipRange") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range) :*: S1 ('MetaSel ('Just "ipMeta") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe MetaId))) :*: (S1 ('MetaSel ('Just "ipSolved") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool) :*: (S1 ('MetaSel ('Just "ipClause") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 IPClause) :*: S1 ('MetaSel ('Just "ipBoundary") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 IPBoundary))))) |
type InteractionPoints = BiMap InteractionId InteractionPoint Source #
Data structure managing the interaction points.
We never remove interaction points from this map, only set their
ipSolved
to True
. (Issue #2368)
Constructors
Interface | |
Fields
|
Instances
Pretty Interface Source # | |||||
InstantiateFull Interface Source # | |||||
Defined in Agda.TypeChecking.Reduce | |||||
EmbPrj Interface Source # | |||||
NFData Interface Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Generic Interface Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
| |||||
Show Interface Source # | |||||
type Rep Interface Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep Interface = D1 ('MetaData "Interface" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "Interface" 'PrefixI 'True) ((((S1 ('MetaSel ('Just "iSourceHash") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Hash) :*: (S1 ('MetaSel ('Just "iSource") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Text) :*: S1 ('MetaSel ('Just "iFileType") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 FileType))) :*: (S1 ('MetaSel ('Just "iImportedModules") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [(TopLevelModuleName, Hash)]) :*: (S1 ('MetaSel ('Just "iModuleName") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ModuleName) :*: S1 ('MetaSel ('Just "iTopLevelModuleName") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 TopLevelModuleName)))) :*: ((S1 ('MetaSel ('Just "iScope") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Map ModuleName Scope)) :*: (S1 ('MetaSel ('Just "iInsideScope") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ScopeInfo) :*: S1 ('MetaSel ('Just "iSignature") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Signature))) :*: (S1 ('MetaSel ('Just "iMetaBindings") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 RemoteMetaStore) :*: (S1 ('MetaSel ('Just "iDisplayForms") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 DisplayForms) :*: S1 ('MetaSel ('Just "iUserWarnings") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 UserWarnings))))) :*: (((S1 ('MetaSel ('Just "iImportWarning") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe Text)) :*: (S1 ('MetaSel ('Just "iBuiltin") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (BuiltinThings' (PrimitiveId, QName))) :*: S1 ('MetaSel ('Just "iForeignCode") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Map BackendName ForeignCodeStack)))) :*: (S1 ('MetaSel ('Just "iHighlighting") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 HighlightingInfo) :*: (S1 ('MetaSel ('Just "iDefaultPragmaOptions") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [OptionsPragma]) :*: S1 ('MetaSel ('Just "iFilePragmaOptions") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [OptionsPragma])))) :*: ((S1 ('MetaSel ('Just "iOptionsUsed") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 PragmaOptions) :*: (S1 ('MetaSel ('Just "iPatternSyns") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 PatternSynDefns) :*: S1 ('MetaSel ('Just "iWarnings") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Set TCWarning)))) :*: (S1 ('MetaSel ('Just "iPartialDefs") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Set QName)) :*: (S1 ('MetaSel ('Just "iOpaqueBlocks") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Map OpaqueId OpaqueBlock)) :*: S1 ('MetaSel ('Just "iOpaqueNames") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Map QName OpaqueId)))))))) |
data InvalidFileNameReason Source #
Extra information for InvalidFileName
error.
Instances
NFData InvalidFileNameReason Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: InvalidFileNameReason -> () # | |||||
Generic InvalidFileNameReason Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
Methods from :: InvalidFileNameReason -> Rep InvalidFileNameReason x # to :: Rep InvalidFileNameReason x -> InvalidFileNameReason # | |||||
Show InvalidFileNameReason Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods showsPrec :: Int -> InvalidFileNameReason -> ShowS # show :: InvalidFileNameReason -> String # showList :: [InvalidFileNameReason] -> ShowS # | |||||
type Rep InvalidFileNameReason Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep InvalidFileNameReason = D1 ('MetaData "InvalidFileNameReason" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "DoesNotCorrespondToValidModuleName" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "RootNameModuleNotAQualifiedModuleName" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Text))) |
type InversionMap c = Map TermHead [c] Source #
data IsAmbiguous Source #
Boolean flag whether a name is ambiguous.
Constructors
YesAmbiguous AmbiguousQName | |
NotAmbiguous |
Instances
EmbPrj IsAmbiguous Source # | |||||
Defined in Agda.TypeChecking.Serialise.Instances.Errors | |||||
NFData IsAmbiguous Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: IsAmbiguous -> () # | |||||
Generic IsAmbiguous Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
| |||||
Show IsAmbiguous Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods showsPrec :: Int -> IsAmbiguous -> ShowS # show :: IsAmbiguous -> String # showList :: [IsAmbiguous] -> ShowS # | |||||
type Rep IsAmbiguous Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep IsAmbiguous = D1 ('MetaData "IsAmbiguous" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "YesAmbiguous" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 AmbiguousQName)) :+: C1 ('MetaCons "NotAmbiguous" 'PrefixI 'False) (U1 :: Type -> Type)) |
Information about whether an argument is forced by the type of a function.
Instances
KillRange IsForced Source # | |
Defined in Agda.TypeChecking.Monad.Base Methods | |
PrettyTCM IsForced Source # | |
Defined in Agda.TypeChecking.Pretty | |
ChooseFlex IsForced Source # | |
Defined in Agda.TypeChecking.Rules.LHS.Problem Methods chooseFlex :: IsForced -> IsForced -> FlexChoice Source # | |
EmbPrj IsForced Source # | |
NFData IsForced Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
Generic IsForced Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
Show IsForced Source # | |
Eq IsForced Source # | |
type Rep IsForced Source # | |
Three cases: 1. not reduced, 2. reduced, but blocked, 3. reduced, not blocked.
Constructors
NotReduced | |
Reduced (Blocked ()) |
data JSBackendError Source #
Errors raised by the JS backend.
Constructors
BadCompilePragma |
Instances
PrettyTCM JSBackendError Source # | |||||
Defined in Agda.TypeChecking.Errors Methods prettyTCM :: MonadPretty m => JSBackendError -> m Doc Source # | |||||
NFData JSBackendError Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: JSBackendError -> () # | |||||
Generic JSBackendError Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
Methods from :: JSBackendError -> Rep JSBackendError x # to :: Rep JSBackendError x -> JSBackendError # | |||||
Show JSBackendError Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods showsPrec :: Int -> JSBackendError -> ShowS # show :: JSBackendError -> String # showList :: [JSBackendError] -> ShowS # | |||||
type Rep JSBackendError Source # | |||||
Parametrized since it is used without MetaId when creating a new meta.
Constructors
HasType | |
Fields
| |
IsSort | |
Instances
Pretty a => Pretty (Judgement a) Source # | |||||
PrettyTCM a => PrettyTCM (Judgement a) Source # | |||||
Defined in Agda.TypeChecking.Pretty | |||||
InstantiateFull (Judgement MetaId) Source # | |||||
Defined in Agda.TypeChecking.Reduce | |||||
EmbPrj a => EmbPrj (Judgement a) Source # | |||||
NFData a => NFData (Judgement a) Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Generic (Judgement a) Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
| |||||
Show a => Show (Judgement a) Source # | |||||
type Rep (Judgement a) Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep (Judgement a) = D1 ('MetaData "Judgement" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "HasType" 'PrefixI 'True) (S1 ('MetaSel ('Just "jMetaId") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 a) :*: (S1 ('MetaSel ('Just "jComparison") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: S1 ('MetaSel ('Just "jMetaType") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type))) :+: C1 ('MetaCons "IsSort" 'PrefixI 'True) (S1 ('MetaSel ('Just "jMetaId") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 a) :*: S1 ('MetaSel ('Just "jMetaType") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type))) |
data LHSOrPatSyn Source #
Distinguish error message when parsing lhs or pattern synonym, resp.
Instances
EmbPrj LHSOrPatSyn Source # | |||||
Defined in Agda.TypeChecking.Serialise.Instances.Errors | |||||
NFData LHSOrPatSyn Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: LHSOrPatSyn -> () # | |||||
Bounded LHSOrPatSyn Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Enum LHSOrPatSyn Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods succ :: LHSOrPatSyn -> LHSOrPatSyn # pred :: LHSOrPatSyn -> LHSOrPatSyn # toEnum :: Int -> LHSOrPatSyn # fromEnum :: LHSOrPatSyn -> Int # enumFrom :: LHSOrPatSyn -> [LHSOrPatSyn] # enumFromThen :: LHSOrPatSyn -> LHSOrPatSyn -> [LHSOrPatSyn] # enumFromTo :: LHSOrPatSyn -> LHSOrPatSyn -> [LHSOrPatSyn] # enumFromThenTo :: LHSOrPatSyn -> LHSOrPatSyn -> LHSOrPatSyn -> [LHSOrPatSyn] # | |||||
Generic LHSOrPatSyn Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
| |||||
Show LHSOrPatSyn Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods showsPrec :: Int -> LHSOrPatSyn -> ShowS # show :: LHSOrPatSyn -> String # showList :: [LHSOrPatSyn] -> ShowS # | |||||
Eq LHSOrPatSyn Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
type Rep LHSOrPatSyn Source # | |||||
class LensClosure b a | b -> a where Source #
Methods
lensClosure :: Lens' b (Closure a) Source #
Instances
LensClosure MetaInfo Range Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
LensClosure MetaVariable Range Source # | |
Defined in Agda.TypeChecking.Monad.Base Methods | |
LensClosure (Closure a) a Source # | |
Defined in Agda.TypeChecking.Monad.Base |
data LetBinding Source #
Instances
InstantiateFull LetBinding Source # | |||||
Defined in Agda.TypeChecking.Reduce Methods instantiateFull' :: LetBinding -> ReduceM LetBinding Source # | |||||
Subst LetBinding Source # | |||||
Defined in Agda.TypeChecking.Substitute Associated Types
Methods applySubst :: Substitution' (SubstArg LetBinding) -> LetBinding -> LetBinding Source # | |||||
NFData LetBinding Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: LetBinding -> () # | |||||
Generic LetBinding Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
| |||||
Show LetBinding Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods showsPrec :: Int -> LetBinding -> ShowS # show :: LetBinding -> String # showList :: [LetBinding] -> ShowS # | |||||
type SubstArg LetBinding Source # | |||||
Defined in Agda.TypeChecking.Substitute | |||||
type Rep LetBinding Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep LetBinding = D1 ('MetaData "LetBinding" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "LetBinding" 'PrefixI 'True) (S1 ('MetaSel ('Just "letOrigin") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Origin) :*: (S1 ('MetaSel ('Just "letTerm") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Just "letType") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Dom Type))))) |
type LetBindings = Map Name (Open LetBinding) Source #
Constructors
EtaExpand MetaId | |
CheckConstraint Nat ProblemConstraint |
Instances
NFData Listener Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Generic Listener Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
| |||||
Eq Listener Source # | |||||
Ord Listener Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
type Rep Listener Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep Listener = D1 ('MetaData "Listener" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "EtaExpand" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 MetaId)) :+: C1 ('MetaCons "CheckConstraint" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Nat) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ProblemConstraint))) |
data LoadedFileCache Source #
Constructors
LoadedFileCache | |
Fields |
Instances
NFData LoadedFileCache Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: LoadedFileCache -> () # | |||||
Generic LoadedFileCache Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
Methods from :: LoadedFileCache -> Rep LoadedFileCache x # to :: Rep LoadedFileCache x -> LoadedFileCache # | |||||
type Rep LoadedFileCache Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep LoadedFileCache = D1 ('MetaData "LoadedFileCache" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "LoadedFileCache" 'PrefixI 'True) (S1 ('MetaSel ('Just "lfcCached") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 CachedTypeCheckLog) :*: S1 ('MetaSel ('Just "lfcCurrent") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 CurrentTypeCheckLog))) |
type LocalDisplayForm = Open DisplayForm Source #
type LocalMetaStore = Map MetaId MetaVariable Source #
Used for meta-variables from the current module.
data MaybeReduced a Source #
Constructors
MaybeRed | |
Fields
|
Instances
Functor MaybeReduced Source # | |
Defined in Agda.TypeChecking.Monad.Base Methods fmap :: (a -> b) -> MaybeReduced a -> MaybeReduced b # (<$) :: a -> MaybeReduced b -> MaybeReduced a # | |
IsProjElim e => IsProjElim (MaybeReduced e) Source # | |
Defined in Agda.TypeChecking.Monad.Base Methods isProjElim :: MaybeReduced e -> Maybe (ProjOrigin, QName) Source # | |
PrettyTCM a => PrettyTCM (MaybeReduced a) Source # | |
Defined in Agda.TypeChecking.Pretty Methods prettyTCM :: MonadPretty m => MaybeReduced a -> m Doc Source # |
type MaybeReducedArgs = [MaybeReduced (Arg Term)] Source #
type MaybeReducedElims = [MaybeReduced Elim] Source #
MetaInfo
is cloned from one meta to the next during pruning.
Constructors
MetaInfo | |
Fields
|
Instances
LensIsAbstract MetaInfo Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods | |||||
LensModalPolarity MetaInfo Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods getModalPolarity :: MetaInfo -> PolarityModality Source # setModalPolarity :: PolarityModality -> MetaInfo -> MetaInfo Source # mapModalPolarity :: (PolarityModality -> PolarityModality) -> MetaInfo -> MetaInfo Source # | |||||
LensModality MetaInfo Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
LensQuantity MetaInfo Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
LensRelevance MetaInfo Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
HasRange MetaInfo Source # | |||||
SetRange MetaInfo Source # | |||||
NFData MetaInfo Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Generic MetaInfo Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
| |||||
LensClosure MetaInfo Range Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
type Rep MetaInfo Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep MetaInfo = D1 ('MetaData "MetaInfo" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "MetaInfo" 'PrefixI 'True) ((S1 ('MetaSel ('Just "miClosRange") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Closure Range)) :*: S1 ('MetaSel ('Just "miModality") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Modality)) :*: (S1 ('MetaSel ('Just "miMetaOccursCheck") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 RunMetaOccursCheck) :*: (S1 ('MetaSel ('Just "miNameSuggestion") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 MetaNameSuggestion) :*: S1 ('MetaSel ('Just "miGeneralizable") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Arg DoGeneralize)))))) |
data MetaInstantiation Source #
Solution status of meta.
Constructors
InstV Instantiation | Solved by |
OpenMeta MetaKind | Unsolved (open to solutions). |
BlockedConst Term | Solved, but solution blocked by unsolved constraints. |
PostponedTypeCheckingProblem (Closure TypeCheckingProblem) | Meta stands for value of the expression that is still to be type checked. |
Instances
Pretty MetaInstantiation Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods pretty :: MetaInstantiation -> Doc Source # prettyPrec :: Int -> MetaInstantiation -> Doc Source # prettyList :: [MetaInstantiation] -> Doc Source # | |||||
NFData MetaInstantiation Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: MetaInstantiation -> () # | |||||
Generic MetaInstantiation Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
Methods from :: MetaInstantiation -> Rep MetaInstantiation x # to :: Rep MetaInstantiation x -> MetaInstantiation # | |||||
type Rep MetaInstantiation Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep MetaInstantiation = D1 ('MetaData "MetaInstantiation" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) ((C1 ('MetaCons "InstV" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Instantiation)) :+: C1 ('MetaCons "OpenMeta" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 MetaKind))) :+: (C1 ('MetaCons "BlockedConst" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)) :+: C1 ('MetaCons "PostponedTypeCheckingProblem" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Closure TypeCheckingProblem))))) |
newtype MetaPriority Source #
Meta variable priority: When we have an equation between meta-variables, which one should be instantiated?
Higher value means higher priority to be instantiated.
Constructors
MetaPriority Int |
Instances
NFData MetaPriority Source # | |
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: MetaPriority -> () # | |
Show MetaPriority Source # | |
Defined in Agda.TypeChecking.Monad.Base Methods showsPrec :: Int -> MetaPriority -> ShowS # show :: MetaPriority -> String # showList :: [MetaPriority] -> ShowS # | |
Eq MetaPriority Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
Ord MetaPriority Source # | |
Defined in Agda.TypeChecking.Monad.Base Methods compare :: MetaPriority -> MetaPriority -> Ordering # (<) :: MetaPriority -> MetaPriority -> Bool # (<=) :: MetaPriority -> MetaPriority -> Bool # (>) :: MetaPriority -> MetaPriority -> Bool # (>=) :: MetaPriority -> MetaPriority -> Bool # max :: MetaPriority -> MetaPriority -> MetaPriority # min :: MetaPriority -> MetaPriority -> MetaPriority # |
data MetaVariable Source #
Information about local meta-variables.
Constructors
MetaVar | |
Fields
|
Instances
LensModality MetaVariable Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods getModality :: MetaVariable -> Modality Source # setModality :: Modality -> MetaVariable -> MetaVariable Source # mapModality :: (Modality -> Modality) -> MetaVariable -> MetaVariable Source # | |||||
LensQuantity MetaVariable Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods getQuantity :: MetaVariable -> Quantity Source # setQuantity :: Quantity -> MetaVariable -> MetaVariable Source # mapQuantity :: (Quantity -> Quantity) -> MetaVariable -> MetaVariable Source # | |||||
LensRelevance MetaVariable Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods getRelevance :: MetaVariable -> Relevance Source # setRelevance :: Relevance -> MetaVariable -> MetaVariable Source # mapRelevance :: (Relevance -> Relevance) -> MetaVariable -> MetaVariable Source # | |||||
HasRange MetaVariable Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods getRange :: MetaVariable -> Range Source # | |||||
SetRange MetaVariable Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods setRange :: Range -> MetaVariable -> MetaVariable Source # | |||||
NFData MetaVariable Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: MetaVariable -> () # | |||||
Generic MetaVariable Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
| |||||
LensClosure MetaVariable Range Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods | |||||
type Rep MetaVariable Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep MetaVariable = D1 ('MetaData "MetaVariable" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "MetaVar" 'PrefixI 'True) (((S1 ('MetaSel ('Just "mvInfo") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 MetaInfo) :*: S1 ('MetaSel ('Just "mvPriority") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 MetaPriority)) :*: (S1 ('MetaSel ('Just "mvPermutation") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Permutation) :*: S1 ('MetaSel ('Just "mvJudgement") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Judgement MetaId)))) :*: ((S1 ('MetaSel ('Just "mvInstantiation") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 MetaInstantiation) :*: S1 ('MetaSel ('Just "mvListeners") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Set Listener))) :*: (S1 ('MetaSel ('Just "mvFrozen") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Frozen) :*: S1 ('MetaSel ('Just "mvTwin") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe MetaId)))))) |
data MissingTypeSignatureInfo Source #
Extra information for MissingTypeSignature
error.
Constructors
MissingDataSignature Name | The |
MissingRecordSignature Name | The |
MissingFunctionSignature LHS | The function lhs misses a type signature. |
Instances
PrettyTCM MissingTypeSignatureInfo Source # | |||||
Defined in Agda.TypeChecking.Errors Methods prettyTCM :: MonadPretty m => MissingTypeSignatureInfo -> m Doc Source # | |||||
NFData MissingTypeSignatureInfo Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: MissingTypeSignatureInfo -> () # | |||||
Generic MissingTypeSignatureInfo Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
Methods from :: MissingTypeSignatureInfo -> Rep MissingTypeSignatureInfo x # to :: Rep MissingTypeSignatureInfo x -> MissingTypeSignatureInfo # | |||||
Show MissingTypeSignatureInfo Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods showsPrec :: Int -> MissingTypeSignatureInfo -> ShowS # show :: MissingTypeSignatureInfo -> String # showList :: [MissingTypeSignatureInfo] -> ShowS # | |||||
type Rep MissingTypeSignatureInfo Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep MissingTypeSignatureInfo = D1 ('MetaData "MissingTypeSignatureInfo" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "MissingDataSignature" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Name)) :+: (C1 ('MetaCons "MissingRecordSignature" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Name)) :+: C1 ('MetaCons "MissingFunctionSignature" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 LHS)))) |
data ModuleCheckMode Source #
Distinguishes between type-checked and scope-checked interfaces
when stored in the map of VisitedModules
.
Constructors
ModuleScopeChecked | |
ModuleTypeChecked |
Instances
NFData ModuleCheckMode Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: ModuleCheckMode -> () # | |||||
Bounded ModuleCheckMode Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Enum ModuleCheckMode Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods succ :: ModuleCheckMode -> ModuleCheckMode # pred :: ModuleCheckMode -> ModuleCheckMode # toEnum :: Int -> ModuleCheckMode # fromEnum :: ModuleCheckMode -> Int # enumFrom :: ModuleCheckMode -> [ModuleCheckMode] # enumFromThen :: ModuleCheckMode -> ModuleCheckMode -> [ModuleCheckMode] # enumFromTo :: ModuleCheckMode -> ModuleCheckMode -> [ModuleCheckMode] # enumFromThenTo :: ModuleCheckMode -> ModuleCheckMode -> ModuleCheckMode -> [ModuleCheckMode] # | |||||
Generic ModuleCheckMode Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
Methods from :: ModuleCheckMode -> Rep ModuleCheckMode x # to :: Rep ModuleCheckMode x -> ModuleCheckMode # | |||||
Show ModuleCheckMode Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods showsPrec :: Int -> ModuleCheckMode -> ShowS # show :: ModuleCheckMode -> String # showList :: [ModuleCheckMode] -> ShowS # | |||||
Eq ModuleCheckMode Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods (==) :: ModuleCheckMode -> ModuleCheckMode -> Bool # (/=) :: ModuleCheckMode -> ModuleCheckMode -> Bool # | |||||
Ord ModuleCheckMode Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods compare :: ModuleCheckMode -> ModuleCheckMode -> Ordering # (<) :: ModuleCheckMode -> ModuleCheckMode -> Bool # (<=) :: ModuleCheckMode -> ModuleCheckMode -> Bool # (>) :: ModuleCheckMode -> ModuleCheckMode -> Bool # (>=) :: ModuleCheckMode -> ModuleCheckMode -> Bool # max :: ModuleCheckMode -> ModuleCheckMode -> ModuleCheckMode # min :: ModuleCheckMode -> ModuleCheckMode -> ModuleCheckMode # | |||||
type Rep ModuleCheckMode Source # | |||||
Defined in Agda.TypeChecking.Monad.Base |
data ModuleInfo Source #
Constructors
ModuleInfo | |
Fields
|
Instances
NFData ModuleInfo Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: ModuleInfo -> () # | |||||
Generic ModuleInfo Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
| |||||
type Rep ModuleInfo Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep ModuleInfo = D1 ('MetaData "ModuleInfo" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "ModuleInfo" 'PrefixI 'True) ((S1 ('MetaSel ('Just "miInterface") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Interface) :*: S1 ('MetaSel ('Just "miWarnings") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Set TCWarning))) :*: (S1 ('MetaSel ('Just "miPrimitive") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool) :*: S1 ('MetaSel ('Just "miMode") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ModuleCheckMode)))) |
class Monad m => MonadBlock (m :: Type -> Type) where Source #
Minimal complete definition
Methods
patternViolation :: Blocker -> m a Source #
`patternViolation b` aborts the current computation
default patternViolation :: forall (t :: (Type -> Type) -> Type -> Type) (n :: Type -> Type) a. (MonadTrans t, MonadBlock n, m ~ t n) => Blocker -> m a Source #
catchPatternErr :: (Blocker -> m a) -> m a -> m a Source #
`catchPatternErr handle m` runs m, handling pattern violations
with handle
(doesn't roll back the state)
Instances
MonadBlock TCM Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
MonadBlock NLM Source # | |
Defined in Agda.TypeChecking.Rewriting.NonLinMatch | |
Monad m => MonadBlock (PureConversionT m) Source # | |
Defined in Agda.TypeChecking.Conversion.Pure Methods patternViolation :: Blocker -> PureConversionT m a Source # catchPatternErr :: (Blocker -> PureConversionT m a) -> PureConversionT m a -> PureConversionT m a Source # | |
Monad m => MonadBlock (BlockT m) Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
MonadBlock m => MonadBlock (MaybeT m) Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
Monad m => MonadBlock (ExceptT TCErr m) Source # | |
MonadBlock m => MonadBlock (ReaderT e m) Source # | |
Defined in Agda.TypeChecking.Monad.Base |
class Monad m => MonadFresh i (m :: Type -> Type) where Source #
Minimal complete definition
Nothing
Methods
Instances
HasFresh i => MonadFresh i TCM Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
Monad m => MonadFresh NameId (PureConversionT m) Source # | |
Defined in Agda.TypeChecking.Conversion.Pure Methods fresh :: PureConversionT m NameId Source # | |
Monad m => MonadFresh ProblemId (PureConversionT m) Source # | |
Defined in Agda.TypeChecking.Conversion.Pure Methods fresh :: PureConversionT m ProblemId Source # | |
Monad m => MonadFresh Int (PureConversionT m) Source # | |
Defined in Agda.TypeChecking.Conversion.Pure Methods fresh :: PureConversionT m Int Source # | |
MonadFresh i m => MonadFresh i (ListT m) Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
MonadFresh i m => MonadFresh i (MaybeT m) Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
MonadFresh i m => MonadFresh i (ExceptT e m) Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
MonadFresh i m => MonadFresh i (IdentityT m) Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
MonadFresh i m => MonadFresh i (ReaderT r m) Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
MonadFresh i m => MonadFresh i (StateT s m) Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
(MonadFresh i m, Monoid w) => MonadFresh i (WriterT w m) Source # | |
Defined in Agda.TypeChecking.Monad.Base |
class (Applicative m, MonadTCEnv m, ReadTCState m, HasOptions m) => MonadReduce (m :: Type -> Type) where Source #
Minimal complete definition
Nothing
Methods
liftReduce :: ReduceM a -> m a Source #
default liftReduce :: forall (t :: (Type -> Type) -> Type -> Type) (n :: Type -> Type) a. (MonadTrans t, MonadReduce n, t n ~ m) => ReduceM a -> m a Source #
Instances
MonadReduce TerM Source # | |
Defined in Agda.Termination.Monad Methods liftReduce :: ReduceM a -> TerM a Source # | |
MonadReduce ReduceM Source # | |
Defined in Agda.TypeChecking.Monad.Base Methods liftReduce :: ReduceM a -> ReduceM a Source # | |
MonadReduce NLM Source # | |
Defined in Agda.TypeChecking.Rewriting.NonLinMatch Methods liftReduce :: ReduceM a -> NLM a Source # | |
MonadReduce m => MonadReduce (PureConversionT m) Source # | |
Defined in Agda.TypeChecking.Conversion.Pure Methods liftReduce :: ReduceM a -> PureConversionT m a Source # | |
MonadReduce m => MonadReduce (BlockT m) Source # | |
Defined in Agda.TypeChecking.Monad.Base Methods liftReduce :: ReduceM a -> BlockT m a Source # | |
MonadIO m => MonadReduce (TCMT m) Source # | |
Defined in Agda.TypeChecking.Monad.Base Methods liftReduce :: ReduceM a -> TCMT m a Source # | |
MonadReduce m => MonadReduce (NamesT m) Source # | |
Defined in Agda.TypeChecking.Names Methods liftReduce :: ReduceM a -> NamesT m a Source # | |
MonadReduce m => MonadReduce (ListT m) Source # | |
Defined in Agda.TypeChecking.Monad.Base Methods liftReduce :: ReduceM a -> ListT m a Source # | |
MonadReduce m => MonadReduce (ChangeT m) Source # | |
Defined in Agda.TypeChecking.Monad.Base Methods liftReduce :: ReduceM a -> ChangeT m a Source # | |
MonadReduce m => MonadReduce (MaybeT m) Source # | |
Defined in Agda.TypeChecking.Monad.Base Methods liftReduce :: ReduceM a -> MaybeT m a Source # | |
MonadReduce m => MonadReduce (ExceptT err m) Source # | |
Defined in Agda.TypeChecking.Monad.Base Methods liftReduce :: ReduceM a -> ExceptT err m a Source # | |
MonadReduce m => MonadReduce (IdentityT m) Source # | |
Defined in Agda.TypeChecking.Monad.Base Methods liftReduce :: ReduceM a -> IdentityT m a Source # | |
MonadReduce m => MonadReduce (ReaderT r m) Source # | |
Defined in Agda.TypeChecking.Monad.Base Methods liftReduce :: ReduceM a -> ReaderT r m a Source # | |
MonadReduce m => MonadReduce (StateT w m) Source # | |
Defined in Agda.TypeChecking.Monad.Base Methods liftReduce :: ReduceM a -> StateT w m a Source # | |
(Monoid w, MonadReduce m) => MonadReduce (WriterT w m) Source # | |
Defined in Agda.TypeChecking.Monad.Base Methods liftReduce :: ReduceM a -> WriterT w m a Source # |
class Monad m => MonadStConcreteNames (m :: Type -> Type) where Source #
A monad that has read and write access to the stConcreteNames part of the TCState. Basically, this is a synonym for `MonadState ConcreteNames m` (which cannot be used directly because of the limitations of Haskell's typeclass system).
Minimal complete definition
Methods
runStConcreteNames :: StateT ConcreteNames m a -> m a Source #
useConcreteNames :: m ConcreteNames Source #
modifyConcreteNames :: (ConcreteNames -> ConcreteNames) -> m () Source #
Instances
MonadStConcreteNames TCM Source # | |
Defined in Agda.TypeChecking.Monad.Base Methods runStConcreteNames :: StateT ConcreteNames TCM a -> TCM a Source # useConcreteNames :: TCM ConcreteNames Source # modifyConcreteNames :: (ConcreteNames -> ConcreteNames) -> TCM () Source # | |
ReadTCState m => MonadStConcreteNames (PureConversionT m) Source # | |
Defined in Agda.TypeChecking.Conversion.Pure Methods runStConcreteNames :: StateT ConcreteNames (PureConversionT m) a -> PureConversionT m a Source # useConcreteNames :: PureConversionT m ConcreteNames Source # modifyConcreteNames :: (ConcreteNames -> ConcreteNames) -> PureConversionT m () Source # | |
MonadStConcreteNames m => MonadStConcreteNames (MaybeT m) Source # | |
Defined in Agda.TypeChecking.Monad.Base Methods runStConcreteNames :: StateT ConcreteNames (MaybeT m) a -> MaybeT m a Source # useConcreteNames :: MaybeT m ConcreteNames Source # modifyConcreteNames :: (ConcreteNames -> ConcreteNames) -> MaybeT m () Source # | |
MonadStConcreteNames m => MonadStConcreteNames (ExceptT e m) Source # | The concrete names get lost in case of an exception. |
Defined in Agda.TypeChecking.Monad.Base Methods runStConcreteNames :: StateT ConcreteNames (ExceptT e m) a -> ExceptT e m a Source # useConcreteNames :: ExceptT e m ConcreteNames Source # modifyConcreteNames :: (ConcreteNames -> ConcreteNames) -> ExceptT e m () Source # | |
MonadStConcreteNames m => MonadStConcreteNames (IdentityT m) Source # | |
Defined in Agda.TypeChecking.Monad.Base Methods runStConcreteNames :: StateT ConcreteNames (IdentityT m) a -> IdentityT m a Source # useConcreteNames :: IdentityT m ConcreteNames Source # modifyConcreteNames :: (ConcreteNames -> ConcreteNames) -> IdentityT m () Source # | |
MonadStConcreteNames m => MonadStConcreteNames (ReaderT r m) Source # | |
Defined in Agda.TypeChecking.Monad.Base Methods runStConcreteNames :: StateT ConcreteNames (ReaderT r m) a -> ReaderT r m a Source # useConcreteNames :: ReaderT r m ConcreteNames Source # modifyConcreteNames :: (ConcreteNames -> ConcreteNames) -> ReaderT r m () Source # | |
MonadStConcreteNames m => MonadStConcreteNames (StateT s m) Source # | |
Defined in Agda.TypeChecking.Monad.Base Methods runStConcreteNames :: StateT ConcreteNames (StateT s m) a -> StateT s m a Source # useConcreteNames :: StateT s m ConcreteNames Source # modifyConcreteNames :: (ConcreteNames -> ConcreteNames) -> StateT s m () Source # | |
(MonadStConcreteNames m, Monoid w) => MonadStConcreteNames (WriterT w m) Source # | |
Defined in Agda.TypeChecking.Monad.Base Methods runStConcreteNames :: StateT ConcreteNames (WriterT w m) a -> WriterT w m a Source # useConcreteNames :: WriterT w m ConcreteNames Source # modifyConcreteNames :: (ConcreteNames -> ConcreteNames) -> WriterT w m () Source # |
class Monad m => MonadTCEnv (m :: Type -> Type) where Source #
MonadTCEnv
made into its own dedicated service class.
This allows us to use MonadReader
for ReaderT
extensions of TCM
.
Minimal complete definition
Nothing
Methods
Instances
MonadTCEnv IM Source # | |
MonadTCEnv TerM Source # | |
MonadTCEnv ReduceM Source # | |
MonadTCEnv NLM Source # | |
MonadTCEnv m => MonadTCEnv (PureConversionT m) Source # | |
Defined in Agda.TypeChecking.Conversion.Pure Methods askTC :: PureConversionT m TCEnv Source # localTC :: (TCEnv -> TCEnv) -> PureConversionT m a -> PureConversionT m a Source # | |
MonadTCEnv m => MonadTCEnv (BlockT m) Source # | |
MonadIO m => MonadTCEnv (TCMT m) Source # | |
MonadTCEnv m => MonadTCEnv (NamesT m) Source # | |
MonadTCEnv m => MonadTCEnv (ListT m) Source # | |
MonadTCEnv m => MonadTCEnv (ChangeT m) Source # | |
MonadTCEnv m => MonadTCEnv (MaybeT m) Source # | |
MonadTCEnv m => MonadTCEnv (ExceptT err m) Source # | |
MonadTCEnv m => MonadTCEnv (IdentityT m) Source # | |
MonadTCEnv m => MonadTCEnv (ReaderT r m) Source # | |
MonadTCEnv m => MonadTCEnv (StateT s m) Source # | |
(Monoid w, MonadTCEnv m) => MonadTCEnv (WriterT w m) Source # | |
type MonadTCError (m :: Type -> Type) = (MonadTCEnv m, ReadTCState m, MonadError TCErr m) Source #
The constraints needed for typeError
and similar.
class (Applicative tcm, MonadIO tcm, MonadTCEnv tcm, MonadTCState tcm, HasOptions tcm) => MonadTCM (tcm :: Type -> Type) where Source #
Embedding a TCM computation.
Minimal complete definition
Nothing
Methods
Instances
MonadTCM IM Source # | |
MonadTCM TerM Source # | |
MonadTCM m => MonadTCM (BlockT m) Source # | |
MonadIO m => MonadTCM (TCMT m) Source # | |
MonadTCM m => MonadTCM (NamesT m) Source # | |
MonadTCM tcm => MonadTCM (ListT tcm) Source # | |
MonadTCM tcm => MonadTCM (ChangeT tcm) Source # | |
MonadTCM tcm => MonadTCM (MaybeT tcm) Source # | |
MonadTCM tcm => MonadTCM (ExceptT err tcm) Source # | |
MonadTCM tcm => MonadTCM (IdentityT tcm) Source # | |
MonadTCM tcm => MonadTCM (ReaderT r tcm) Source # | |
MonadTCM tcm => MonadTCM (StateT s tcm) Source # | |
(Monoid w, MonadTCM tcm) => MonadTCM (WriterT w tcm) Source # | |
class Monad m => MonadTCState (m :: Type -> Type) where Source #
MonadTCState
made into its own dedicated service class.
This allows us to use MonadState
for StateT
extensions of TCM
.
Minimal complete definition
Nothing
Methods
default getTC :: forall (t :: (Type -> Type) -> Type -> Type) (n :: Type -> Type). (MonadTrans t, MonadTCState n, t n ~ m) => m TCState Source #
Instances
MonadTCState IM Source # | |
MonadTCState TerM Source # | |
MonadTCState m => MonadTCState (BlockT m) Source # | |
MonadIO m => MonadTCState (TCMT m) Source # | |
MonadTCState m => MonadTCState (NamesT m) Source # | |
MonadTCState m => MonadTCState (ListT m) Source # | |
MonadTCState m => MonadTCState (ChangeT m) Source # | |
MonadTCState m => MonadTCState (MaybeT m) Source # | |
MonadTCState m => MonadTCState (ExceptT err m) Source # | |
MonadTCState m => MonadTCState (IdentityT m) Source # | |
MonadTCState m => MonadTCState (ReaderT r m) Source # | |
MonadTCState m => MonadTCState (StateT s m) Source # | |
(Monoid w, MonadTCState m) => MonadTCState (WriterT w m) Source # | |
data MutualBlock Source #
A mutual block of names in the signature.
Constructors
MutualBlock | |
Fields
|
Instances
Null MutualBlock Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
NFData MutualBlock Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: MutualBlock -> () # | |||||
Generic MutualBlock Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
| |||||
Show MutualBlock Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods showsPrec :: Int -> MutualBlock -> ShowS # show :: MutualBlock -> String # showList :: [MutualBlock] -> ShowS # | |||||
Eq MutualBlock Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
type Rep MutualBlock Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep MutualBlock = D1 ('MetaData "MutualBlock" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "MutualBlock" 'PrefixI 'True) (S1 ('MetaSel ('Just "mutualInfo") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 MutualInfo) :*: S1 ('MetaSel ('Just "mutualNames") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Set QName)))) |
type MutualBlocks = IntMap MutualBlock Source #
Map MutualId
to MutualBlock
.
Instances
Pretty MutualId Source # | |
KillRange MutualId Source # | |
Defined in Agda.TypeChecking.Monad.Base Methods | |
HasFresh MutualId Source # | |
EmbPrj MutualId Source # | |
NFData MutualId Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
Enum MutualId Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
Num MutualId Source # | |
Show MutualId Source # | |
Eq MutualId Source # | |
Ord MutualId Source # | |
Defined in Agda.TypeChecking.Monad.Base |
Constructors
PUniv Univ NLPat | |
PInf Univ Integer | |
PSizeUniv | |
PLockUniv | |
PLevelUniv | |
PIntervalUniv |
Instances
TermLike NLPSort Source # | |||||
AllMetas NLPSort Source # | |||||
NamesIn NLPSort Source # | |||||
Defined in Agda.Syntax.Internal.Names | |||||
KillRange NLPSort Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods | |||||
Free NLPSort Source # | |||||
PrettyTCM NLPSort Source # | |||||
Defined in Agda.TypeChecking.Pretty | |||||
InstantiateFull NLPSort Source # | |||||
Defined in Agda.TypeChecking.Reduce | |||||
GetMatchables NLPSort Source # | |||||
Defined in Agda.TypeChecking.Rewriting.NonLinPattern Methods getMatchables :: NLPSort -> [QName] Source # | |||||
NLPatVars NLPSort Source # | |||||
EmbPrj NLPSort Source # | |||||
Subst NLPSort Source # | |||||
Defined in Agda.TypeChecking.Substitute Associated Types
Methods applySubst :: Substitution' (SubstArg NLPSort) -> NLPSort -> NLPSort Source # | |||||
NFData NLPSort Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Generic NLPSort Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
| |||||
Show NLPSort Source # | |||||
Match NLPSort Sort Source # | |||||
NLPatToTerm NLPSort Sort Source # | |||||
Defined in Agda.TypeChecking.Rewriting.NonLinPattern | |||||
PatternFrom Sort NLPSort Source # | |||||
type SubstArg NLPSort Source # | |||||
Defined in Agda.TypeChecking.Substitute | |||||
type Rep NLPSort Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep NLPSort = D1 ('MetaData "NLPSort" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) ((C1 ('MetaCons "PUniv" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Univ) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 NLPat)) :+: (C1 ('MetaCons "PInf" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Univ) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Integer)) :+: C1 ('MetaCons "PSizeUniv" 'PrefixI 'False) (U1 :: Type -> Type))) :+: (C1 ('MetaCons "PLockUniv" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "PLevelUniv" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "PIntervalUniv" 'PrefixI 'False) (U1 :: Type -> Type)))) |
Constructors
NLPType | |
Fields
|
Instances
TermLike NLPType Source # | |||||
AllMetas NLPType Source # | |||||
NamesIn NLPType Source # | |||||
Defined in Agda.Syntax.Internal.Names | |||||
KillRange NLPType Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods | |||||
Free NLPType Source # | |||||
PrettyTCM NLPType Source # | |||||
Defined in Agda.TypeChecking.Pretty | |||||
InstantiateFull NLPType Source # | |||||
Defined in Agda.TypeChecking.Reduce | |||||
GetMatchables NLPType Source # | |||||
Defined in Agda.TypeChecking.Rewriting.NonLinPattern Methods getMatchables :: NLPType -> [QName] Source # | |||||
NLPatVars NLPType Source # | |||||
EmbPrj NLPType Source # | |||||
Subst NLPType Source # | |||||
Defined in Agda.TypeChecking.Substitute Associated Types
Methods applySubst :: Substitution' (SubstArg NLPType) -> NLPType -> NLPType Source # | |||||
NFData NLPType Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Generic NLPType Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
| |||||
Show NLPType Source # | |||||
Match NLPType Type Source # | |||||
NLPatToTerm NLPType Type Source # | |||||
Defined in Agda.TypeChecking.Rewriting.NonLinPattern | |||||
PatternFrom Type NLPType Source # | |||||
type SubstArg NLPType Source # | |||||
Defined in Agda.TypeChecking.Substitute | |||||
type Rep NLPType Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep NLPType = D1 ('MetaData "NLPType" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "NLPType" 'PrefixI 'True) (S1 ('MetaSel ('Just "nlpTypeSort") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 NLPSort) :*: S1 ('MetaSel ('Just "nlpTypeUnEl") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 NLPat))) |
Non-linear (non-constructor) first-order pattern.
Constructors
PVar !Int [Arg Int] | Matches anything (modulo non-linearity) that only contains bound variables that occur in the given arguments. |
PDef QName PElims | Matches |
PLam ArgInfo (Abs NLPat) | Matches |
PPi (Dom NLPType) (Abs NLPType) | Matches |
PSort NLPSort | Matches a sort of the given shape. |
PBoundVar !Int PElims | Matches |
PTerm Term | Matches the term modulo β (ideally βη). |
Instances
TermLike NLPat Source # | |||||
AllMetas NLPat Source # | |||||
NamesIn NLPat Source # | |||||
Defined in Agda.Syntax.Internal.Names | |||||
KillRange NLPat Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods | |||||
Free NLPat Source # | Only computes free variables that are not bound (see | ||||
PrettyTCM NLPat Source # | |||||
Defined in Agda.TypeChecking.Pretty | |||||
InstantiateFull NLPat Source # | |||||
Defined in Agda.TypeChecking.Reduce | |||||
GetMatchables NLPat Source # | |||||
Defined in Agda.TypeChecking.Rewriting.NonLinPattern Methods getMatchables :: NLPat -> [QName] Source # | |||||
NLPatVars NLPat Source # | |||||
EmbPrj NLPat Source # | |||||
Subst NLPat Source # | |||||
Defined in Agda.TypeChecking.Substitute Associated Types
Methods applySubst :: Substitution' (SubstArg NLPat) -> NLPat -> NLPat Source # | |||||
DeBruijn NLPat Source # | |||||
Defined in Agda.TypeChecking.Substitute | |||||
NFData NLPat Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Generic NLPat Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
| |||||
Show NLPat Source # | |||||
Match NLPat Level Source # | |||||
Match NLPat Term Source # | |||||
NLPatToTerm NLPat Level Source # | |||||
Defined in Agda.TypeChecking.Rewriting.NonLinPattern | |||||
NLPatToTerm NLPat Term Source # | |||||
Defined in Agda.TypeChecking.Rewriting.NonLinPattern | |||||
PatternFrom Level NLPat Source # | |||||
PatternFrom Term NLPat Source # | |||||
PatternFrom Elims [Elim' NLPat] Source # | |||||
PrettyTCM (Type' NLPat) Source # | |||||
Defined in Agda.TypeChecking.Pretty | |||||
PrettyTCM (Elim' NLPat) Source # | |||||
Defined in Agda.TypeChecking.Pretty | |||||
Match [Elim' NLPat] Elims Source # | |||||
ToNLPat (Arg DeBruijnPattern) (Elim' NLPat) Source # | |||||
Defined in Agda.TypeChecking.Rewriting.Clause | |||||
ToNLPat (NamedArg DeBruijnPattern) (Elim' NLPat) Source # | |||||
Defined in Agda.TypeChecking.Rewriting.Clause | |||||
type TypeOf NLPat Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
type SubstArg NLPat Source # | |||||
Defined in Agda.TypeChecking.Substitute | |||||
type Rep NLPat Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep NLPat = D1 ('MetaData "NLPat" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) ((C1 ('MetaCons "PVar" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Int) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Arg Int])) :+: (C1 ('MetaCons "PDef" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 PElims)) :+: C1 ('MetaCons "PLam" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ArgInfo) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Abs NLPat))))) :+: ((C1 ('MetaCons "PPi" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Dom NLPType)) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Abs NLPType))) :+: C1 ('MetaCons "PSort" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 NLPSort))) :+: (C1 ('MetaCons "PBoundVar" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'SourceUnpack 'SourceStrict 'DecidedStrict) (Rec0 Int) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 PElims)) :+: C1 ('MetaCons "PTerm" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term))))) | |||||
type TypeOf [Elim' NLPat] Source # | |||||
data NegativeUnification Source #
Constructors
UnifyConflict Telescope Term Term | |
UnifyCycle Telescope Int Term |
Instances
PrettyTCM NegativeUnification Source # | |||||
Defined in Agda.TypeChecking.Errors Methods prettyTCM :: MonadPretty m => NegativeUnification -> m Doc Source # | |||||
NFData NegativeUnification Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: NegativeUnification -> () # | |||||
Generic NegativeUnification Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
Methods from :: NegativeUnification -> Rep NegativeUnification x # to :: Rep NegativeUnification x -> NegativeUnification # | |||||
Show NegativeUnification Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods showsPrec :: Int -> NegativeUnification -> ShowS # show :: NegativeUnification -> String # showList :: [NegativeUnification] -> ShowS # | |||||
type Rep NegativeUnification Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep NegativeUnification = D1 ('MetaData "NegativeUnification" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "UnifyConflict" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Telescope) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term))) :+: C1 ('MetaCons "UnifyCycle" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Telescope) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Int) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)))) |
data NumGeneralizableArgs Source #
Constructors
NoGeneralizableArgs | |
SomeGeneralizableArgs !Int | When lambda-lifting new args are generalizable if
|
Instances
KillRange NumGeneralizableArgs Source # | |
Defined in Agda.TypeChecking.Monad.Base Methods | |
EmbPrj NumGeneralizableArgs Source # | |
Abstract NumGeneralizableArgs Source # | |
Defined in Agda.TypeChecking.Substitute Methods abstract :: Telescope -> NumGeneralizableArgs -> NumGeneralizableArgs Source # | |
Apply NumGeneralizableArgs Source # | |
Defined in Agda.TypeChecking.Substitute Methods apply :: NumGeneralizableArgs -> Args -> NumGeneralizableArgs Source # applyE :: NumGeneralizableArgs -> Elims -> NumGeneralizableArgs Source # | |
NFData NumGeneralizableArgs Source # | |
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: NumGeneralizableArgs -> () # | |
Show NumGeneralizableArgs Source # | |
Defined in Agda.TypeChecking.Monad.Base Methods showsPrec :: Int -> NumGeneralizableArgs -> ShowS # show :: NumGeneralizableArgs -> String # showList :: [NumGeneralizableArgs] -> ShowS # |
data OpaqueBlock Source #
A block of opaque definitions.
Constructors
OpaqueBlock | |
Fields
|
Instances
Pretty OpaqueBlock Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods pretty :: OpaqueBlock -> Doc Source # prettyPrec :: Int -> OpaqueBlock -> Doc Source # prettyList :: [OpaqueBlock] -> Doc Source # | |||||
EmbPrj OpaqueBlock Source # | |||||
NFData OpaqueBlock Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: OpaqueBlock -> () # | |||||
Generic OpaqueBlock Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
| |||||
Show OpaqueBlock Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods showsPrec :: Int -> OpaqueBlock -> ShowS # show :: OpaqueBlock -> String # showList :: [OpaqueBlock] -> ShowS # | |||||
Eq OpaqueBlock Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Hashable OpaqueBlock Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
type Rep OpaqueBlock Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep OpaqueBlock = D1 ('MetaData "OpaqueBlock" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "OpaqueBlock" 'PrefixI 'True) ((S1 ('MetaSel ('Just "opaqueId") 'SourceUnpack 'SourceStrict 'DecidedStrict) (Rec0 OpaqueId) :*: S1 ('MetaSel ('Just "opaqueUnfolding") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (HashSet QName))) :*: (S1 ('MetaSel ('Just "opaqueDecls") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (HashSet QName)) :*: (S1 ('MetaSel ('Just "opaqueParent") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe OpaqueId)) :*: S1 ('MetaSel ('Just "opaqueRange") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range))))) |
A thing tagged with the context it came from. Also keeps the substitution from previous checkpoints. This lets us handle the case when an open thing was created in a context that we have since exited. Remember which module it's from to make sure we don't get confused by checkpoints from other files.
Constructors
OpenThing | |
Instances
Decoration Open Source # | |||||
Functor Open Source # | |||||
Foldable Open Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods fold :: Monoid m => Open m -> m # foldMap :: Monoid m => (a -> m) -> Open a -> m # foldMap' :: Monoid m => (a -> m) -> Open a -> m # foldr :: (a -> b -> b) -> b -> Open a -> b # foldr' :: (a -> b -> b) -> b -> Open a -> b # foldl :: (b -> a -> b) -> b -> Open a -> b # foldl' :: (b -> a -> b) -> b -> Open a -> b # foldr1 :: (a -> a -> a) -> Open a -> a # foldl1 :: (a -> a -> a) -> Open a -> a # elem :: Eq a => a -> Open a -> Bool # maximum :: Ord a => Open a -> a # | |||||
Traversable Open Source # | |||||
Pretty a => Pretty (Open a) Source # | |||||
NamesIn a => NamesIn (Open a) Source # | |||||
Defined in Agda.Syntax.Internal.Names | |||||
KillRange a => KillRange (Open a) Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods killRange :: KillRangeT (Open a) Source # | |||||
ChaseDisplayForms a => ChaseDisplayForms (Open a) Source # | |||||
Defined in Agda.TypeChecking.Monad.Signature | |||||
InstantiateFull t => InstantiateFull (Open t) Source # | |||||
Defined in Agda.TypeChecking.Reduce | |||||
EmbPrj a => EmbPrj (Open a) Source # | |||||
NFData a => NFData (Open a) Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Generic (Open a) Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
| |||||
Show a => Show (Open a) Source # | |||||
type Rep (Open a) Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep (Open a) = D1 ('MetaData "Open" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "OpenThing" 'PrefixI 'True) ((S1 ('MetaSel ('Just "openThingCheckpoint") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 CheckpointId) :*: S1 ('MetaSel ('Just "openThingCheckpointMap") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Map CheckpointId Substitution))) :*: (S1 ('MetaSel ('Just "openThingModule") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ModuleNameHash) :*: S1 ('MetaSel ('Just "openThing") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 a)))) |
data Overapplied Source #
Flag to indicate whether the meta is overapplied in the constraint. A meta is overapplied if it has more arguments than the size of the telescope in its creation environment (as stored in MetaInfo).
Constructors
Overapplied | |
NotOverapplied |
Instances
NFData Overapplied Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: Overapplied -> () # | |||||
Generic Overapplied Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
| |||||
Show Overapplied Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods showsPrec :: Int -> Overapplied -> ShowS # show :: Overapplied -> String # showList :: [Overapplied] -> ShowS # | |||||
Eq Overapplied Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
type Rep Overapplied Source # | |||||
data PersistentTCState Source #
A part of the state which is not reverted when an error is thrown or the state is reset.
Constructors
PersistentTCSt | |
Fields
|
Instances
LensCommandLineOptions PersistentTCState Source # | |||||
Defined in Agda.Interaction.Options.Lenses | |||||
LensIncludePaths PersistentTCState Source # | |||||
Defined in Agda.Interaction.Options.Lenses Methods getIncludePaths :: PersistentTCState -> [FilePath] Source # setIncludePaths :: [FilePath] -> PersistentTCState -> PersistentTCState Source # mapIncludePaths :: ([FilePath] -> [FilePath]) -> PersistentTCState -> PersistentTCState Source # getAbsoluteIncludePaths :: PersistentTCState -> [AbsolutePath] Source # setAbsoluteIncludePaths :: [AbsolutePath] -> PersistentTCState -> PersistentTCState Source # mapAbsoluteIncludePaths :: ([AbsolutePath] -> [AbsolutePath]) -> PersistentTCState -> PersistentTCState Source # | |||||
LensPersistentVerbosity PersistentTCState Source # | |||||
Defined in Agda.Interaction.Options.Lenses | |||||
LensSafeMode PersistentTCState Source # | |||||
Defined in Agda.Interaction.Options.Lenses Methods getSafeMode :: PersistentTCState -> SafeMode Source # setSafeMode :: SafeMode -> PersistentTCState -> PersistentTCState Source # mapSafeMode :: (SafeMode -> SafeMode) -> PersistentTCState -> PersistentTCState Source # | |||||
NFData PersistentTCState Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: PersistentTCState -> () # | |||||
Generic PersistentTCState Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
Methods from :: PersistentTCState -> Rep PersistentTCState x # to :: Rep PersistentTCState x -> PersistentTCState # | |||||
type Rep PersistentTCState Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep PersistentTCState = D1 ('MetaData "PersistentTCState" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "PersistentTCSt" 'PrefixI 'True) ((S1 ('MetaSel ('Just "stPersistentSession") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 SessionTCState) :*: (S1 ('MetaSel ('Just "stDecodedModules") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 DecodedModules) :*: S1 ('MetaSel ('Just "stPersistentTopLevelModuleNames") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (BiMap RawTopLevelModuleName ModuleNameHash)))) :*: ((S1 ('MetaSel ('Just "stPersistentOptions") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 CommandLineOptions) :*: S1 ('MetaSel ('Just "stInteractionOutputCallback") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 InteractionOutputCallback)) :*: (S1 ('MetaSel ('Just "stAccumStatistics") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Statistics) :*: S1 ('MetaSel ('Just "stPersistLoadedFileCache") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (Maybe LoadedFileCache)))))) |
data PostScopeState Source #
Constructors
PostScopeState | |
Fields
|
Instances
NFData PostScopeState Source # | This instance could be optimised, some things are guaranteed to be forced. | ||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: PostScopeState -> () # | |||||
Generic PostScopeState Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
Methods from :: PostScopeState -> Rep PostScopeState x # to :: Rep PostScopeState x -> PostScopeState # | |||||
type Rep PostScopeState Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep PostScopeState = D1 ('MetaData "PostScopeState" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "PostScopeState" 'PrefixI 'True) (((((S1 ('MetaSel ('Just "stPostSyntaxInfo") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 HighlightingInfo) :*: S1 ('MetaSel ('Just "stPostDisambiguatedNames") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 DisambiguatedNames)) :*: (S1 ('MetaSel ('Just "stPostOpenMetaStore") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 LocalMetaStore) :*: S1 ('MetaSel ('Just "stPostSolvedMetaStore") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 LocalMetaStore))) :*: ((S1 ('MetaSel ('Just "stPostInteractionPoints") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 InteractionPoints) :*: S1 ('MetaSel ('Just "stPostAwakeConstraints") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Constraints)) :*: (S1 ('MetaSel ('Just "stPostSleepingConstraints") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Constraints) :*: (S1 ('MetaSel ('Just "stPostDirty") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Bool) :*: S1 ('MetaSel ('Just "stPostOccursCheckDefs") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (Set QName)))))) :*: (((S1 ('MetaSel ('Just "stPostSignature") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Signature) :*: S1 ('MetaSel ('Just "stPostModuleCheckpoints") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (Map ModuleName CheckpointId))) :*: (S1 ('MetaSel ('Just "stPostImportsDisplayForms") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 DisplayForms) :*: (S1 ('MetaSel ('Just "stPostForeignCode") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 BackendForeignCode) :*: S1 ('MetaSel ('Just "stPostCurrentModule") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (Maybe (ModuleName, TopLevelModuleName)))))) :*: ((S1 ('MetaSel ('Just "stPostPendingInstances") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (Set QName)) :*: S1 ('MetaSel ('Just "stPostTemporaryInstances") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (Set QName))) :*: (S1 ('MetaSel ('Just "stPostConcreteNames") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 ConcreteNames) :*: (S1 ('MetaSel ('Just "stPostUsedNames") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 UsedNames) :*: S1 ('MetaSel ('Just "stPostShadowingNames") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 ShadowingNames)))))) :*: ((((S1 ('MetaSel ('Just "stPostStatistics") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Statistics) :*: S1 ('MetaSel ('Just "stPostTCWarnings") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (Set TCWarning))) :*: (S1 ('MetaSel ('Just "stPostMutualBlocks") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 MutualBlocks) :*: S1 ('MetaSel ('Just "stPostLocalBuiltins") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 BuiltinThings))) :*: ((S1 ('MetaSel ('Just "stPostFreshMetaId") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 MetaId) :*: S1 ('MetaSel ('Just "stPostFreshMutualId") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 MutualId)) :*: (S1 ('MetaSel ('Just "stPostFreshProblemId") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 ProblemId) :*: (S1 ('MetaSel ('Just "stPostFreshCheckpointId") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 CheckpointId) :*: S1 ('MetaSel ('Just "stPostFreshInt") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Int))))) :*: (((S1 ('MetaSel ('Just "stPostFreshNameId") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 NameId) :*: S1 ('MetaSel ('Just "stPostFreshOpaqueId") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 OpaqueId)) :*: (S1 ('MetaSel ('Just "stPostAreWeCaching") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Bool) :*: (S1 ('MetaSel ('Just "stPostPostponeInstanceSearch") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Bool) :*: S1 ('MetaSel ('Just "stPostConsideringInstance") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Bool)))) :*: ((S1 ('MetaSel ('Just "stPostInstantiateBlocking") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Bool) :*: S1 ('MetaSel ('Just "stPostLocalPartialDefs") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (Set QName))) :*: (S1 ('MetaSel ('Just "stPostOpaqueBlocks") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Map OpaqueId OpaqueBlock)) :*: (S1 ('MetaSel ('Just "stPostOpaqueIds") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Map QName OpaqueId)) :*: S1 ('MetaSel ('Just "stPostInstanceHack") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Bool)))))))) |
data PreScopeState Source #
Constructors
PreScopeState | |
Fields
|
Instances
LensPragmaOptions PreScopeState Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods getPragmaOptions :: PreScopeState -> PragmaOptions Source # setPragmaOptions :: PragmaOptions -> PreScopeState -> PreScopeState Source # mapPragmaOptions :: (PragmaOptions -> PragmaOptions) -> PreScopeState -> PreScopeState Source # lensPragmaOptions :: Lens' PreScopeState PragmaOptions Source # | |||||
NFData PreScopeState Source # | This instance could be optimised, some things are guaranteed to be forced. | ||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: PreScopeState -> () # | |||||
Generic PreScopeState Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
| |||||
type Rep PreScopeState Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep PreScopeState = D1 ('MetaData "PreScopeState" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "PreScopeState" 'PrefixI 'True) ((((S1 ('MetaSel ('Just "stPreTokens") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 HighlightingInfo) :*: S1 ('MetaSel ('Just "stPreImports") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Signature)) :*: (S1 ('MetaSel ('Just "stPreImportedModules") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 ImportedModules) :*: (S1 ('MetaSel ('Just "stPreModuleToSourceId") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 ModuleToSourceId) :*: S1 ('MetaSel ('Just "stPreVisitedModules") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 VisitedModules)))) :*: ((S1 ('MetaSel ('Just "stPreScope") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 ScopeInfo) :*: S1 ('MetaSel ('Just "stPrePatternSyns") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 PatternSynDefns)) :*: (S1 ('MetaSel ('Just "stPrePatternSynImports") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 PatternSynDefns) :*: (S1 ('MetaSel ('Just "stPreGeneralizedVars") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (Maybe (Set QName))) :*: S1 ('MetaSel ('Just "stPrePragmaOptions") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 PragmaOptions))))) :*: (((S1 ('MetaSel ('Just "stPreImportedBuiltins") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 BuiltinThings) :*: S1 ('MetaSel ('Just "stPreImportedDisplayForms") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 DisplayForms)) :*: (S1 ('MetaSel ('Just "stPreFreshInteractionId") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 InteractionId) :*: (S1 ('MetaSel ('Just "stPreImportedUserWarnings") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 UserWarnings) :*: S1 ('MetaSel ('Just "stPreLocalUserWarnings") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 UserWarnings)))) :*: ((S1 ('MetaSel ('Just "stPreWarningOnImport") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (Maybe Text)) :*: (S1 ('MetaSel ('Just "stPreImportedPartialDefs") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (Set QName)) :*: S1 ('MetaSel ('Just "stPreLibCache") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 LibCache))) :*: (S1 ('MetaSel ('Just "stPreImportedMetaStore") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 RemoteMetaStore) :*: (S1 ('MetaSel ('Just "stPreCopiedNames") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (HashMap QName QName)) :*: S1 ('MetaSel ('Just "stPreNameCopies") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (HashMap QName (HashSet QName))))))))) |
Constructors
PrimFun | |
Fields
|
Instances
NamesIn PrimFun Source # | Note that the | ||||
Defined in Agda.Syntax.Internal.Names | |||||
Abstract PrimFun Source # | |||||
Apply PrimFun Source # | |||||
NFData PrimFun Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Generic PrimFun Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
| |||||
type Rep PrimFun Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep PrimFun = D1 ('MetaData "PrimFun" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "PrimFun" 'PrefixI 'True) ((S1 ('MetaSel ('Just "primFunName") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Just "primFunArity") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Arity)) :*: (S1 ('MetaSel ('Just "primFunArgOccurrences") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Occurrence]) :*: S1 ('MetaSel ('Just "primFunImplementation") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ([Arg Term] -> Int -> ReduceM (Reduced MaybeReducedArgs Term)))))) |
data PrimitiveData Source #
Constructors
PrimitiveData | |
Fields
|
Instances
Pretty PrimitiveData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods pretty :: PrimitiveData -> Doc Source # prettyPrec :: Int -> PrimitiveData -> Doc Source # prettyList :: [PrimitiveData] -> Doc Source # | |||||
NFData PrimitiveData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: PrimitiveData -> () # | |||||
Generic PrimitiveData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
| |||||
Show PrimitiveData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods showsPrec :: Int -> PrimitiveData -> ShowS # show :: PrimitiveData -> String # showList :: [PrimitiveData] -> ShowS # | |||||
type Rep PrimitiveData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep PrimitiveData = D1 ('MetaData "PrimitiveData" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "PrimitiveData" 'PrefixI 'True) ((S1 ('MetaSel ('Just "_primAbstr") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 IsAbstract) :*: (S1 ('MetaSel ('Just "_primName") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 PrimitiveId) :*: S1 ('MetaSel ('Just "_primClauses") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Clause]))) :*: (S1 ('MetaSel ('Just "_primInv") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 FunctionInverse) :*: (S1 ('MetaSel ('Just "_primCompiled") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe CompiledClauses)) :*: S1 ('MetaSel ('Just "_primOpaque") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 IsOpaque))))) |
data PrimitiveSortData Source #
Constructors
PrimitiveSortData | |
Fields |
Instances
Pretty PrimitiveSortData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods pretty :: PrimitiveSortData -> Doc Source # prettyPrec :: Int -> PrimitiveSortData -> Doc Source # prettyList :: [PrimitiveSortData] -> Doc Source # | |||||
NFData PrimitiveSortData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: PrimitiveSortData -> () # | |||||
Generic PrimitiveSortData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
Methods from :: PrimitiveSortData -> Rep PrimitiveSortData x # to :: Rep PrimitiveSortData x -> PrimitiveSortData # | |||||
Show PrimitiveSortData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods showsPrec :: Int -> PrimitiveSortData -> ShowS # show :: PrimitiveSortData -> String # showList :: [PrimitiveSortData] -> ShowS # | |||||
type Rep PrimitiveSortData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep PrimitiveSortData = D1 ('MetaData "PrimitiveSortData" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "PrimitiveSortData" 'PrefixI 'True) (S1 ('MetaSel ('Just "_primSortName") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 BuiltinSort) :*: S1 ('MetaSel ('Just "_primSortSort") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Sort))) |
data PrincipalArgTypeMetas Source #
Constructors
PrincipalArgTypeMetas | |
Fields
|
Instances
NFData PrincipalArgTypeMetas Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: PrincipalArgTypeMetas -> () # | |||||
Generic PrincipalArgTypeMetas Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
Methods from :: PrincipalArgTypeMetas -> Rep PrincipalArgTypeMetas x # to :: Rep PrincipalArgTypeMetas x -> PrincipalArgTypeMetas # | |||||
type Rep PrincipalArgTypeMetas Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep PrincipalArgTypeMetas = D1 ('MetaData "PrincipalArgTypeMetas" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "PrincipalArgTypeMetas" 'PrefixI 'True) (S1 ('MetaSel ('Just "patmMetas") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Args) :*: S1 ('MetaSel ('Just "patmRemainder") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type))) |
data ProblemConstraint Source #
Constructors
PConstr | |
Fields |
Instances
HasRange ProblemConstraint Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods getRange :: ProblemConstraint -> Range Source # | |||||
Reify ProblemConstraint Source # | |||||
Defined in Agda.Interaction.BasicOps Associated Types
Methods reify :: MonadReify m => ProblemConstraint -> m (ReifiesTo ProblemConstraint) Source # reifyWhen :: MonadReify m => Bool -> ProblemConstraint -> m (ReifiesTo ProblemConstraint) Source # | |||||
MentionsMeta ProblemConstraint Source # | |||||
Defined in Agda.TypeChecking.MetaVars.Mention Methods mentionsMetas :: HashSet MetaId -> ProblemConstraint -> Bool Source # | |||||
PrettyTCM ProblemConstraint Source # | |||||
Defined in Agda.TypeChecking.Pretty.Constraint Methods prettyTCM :: MonadPretty m => ProblemConstraint -> m Doc Source # | |||||
InstantiateFull ProblemConstraint Source # | |||||
Defined in Agda.TypeChecking.Reduce Methods instantiateFull' :: ProblemConstraint -> ReduceM ProblemConstraint Source # | |||||
Normalise ProblemConstraint Source # | |||||
Defined in Agda.TypeChecking.Reduce Methods normalise' :: ProblemConstraint -> ReduceM ProblemConstraint Source # | |||||
Simplify ProblemConstraint Source # | |||||
Defined in Agda.TypeChecking.Reduce Methods | |||||
NFData ProblemConstraint Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: ProblemConstraint -> () # | |||||
Generic ProblemConstraint Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
Methods from :: ProblemConstraint -> Rep ProblemConstraint x # to :: Rep ProblemConstraint x -> ProblemConstraint # | |||||
Show ProblemConstraint Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods showsPrec :: Int -> ProblemConstraint -> ShowS # show :: ProblemConstraint -> String # showList :: [ProblemConstraint] -> ShowS # | |||||
type ReifiesTo ProblemConstraint Source # | |||||
Defined in Agda.Interaction.BasicOps | |||||
type Rep ProblemConstraint Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep ProblemConstraint = D1 ('MetaData "ProblemConstraint" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "PConstr" 'PrefixI 'True) (S1 ('MetaSel ('Just "constraintProblems") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Set ProblemId)) :*: (S1 ('MetaSel ('Just "constraintUnblocker") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Blocker) :*: S1 ('MetaSel ('Just "theConstraint") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Closure Constraint))))) |
Abstractions to build projection function (dropping parameters).
Constructors
ProjLams | |
Fields
|
Instances
Pretty ProjLams Source # | |||||
KillRange ProjLams Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods | |||||
EmbPrj ProjLams Source # | |||||
Abstract ProjLams Source # | |||||
Apply ProjLams Source # | |||||
Null ProjLams Source # | |||||
NFData ProjLams Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Generic ProjLams Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
| |||||
Show ProjLams Source # | |||||
type Rep ProjLams Source # | |||||
Defined in Agda.TypeChecking.Monad.Base |
data Projection Source #
Additional information for projection Function
s.
Constructors
Projection | |
Fields
|
Instances
Pretty Projection Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods pretty :: Projection -> Doc Source # prettyPrec :: Int -> Projection -> Doc Source # prettyList :: [Projection] -> Doc Source # | |||||
KillRange Projection Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods | |||||
EmbPrj Projection Source # | |||||
Abstract Projection Source # | |||||
Defined in Agda.TypeChecking.Substitute Methods abstract :: Telescope -> Projection -> Projection Source # | |||||
Apply Projection Source # | |||||
Defined in Agda.TypeChecking.Substitute Methods apply :: Projection -> Args -> Projection Source # applyE :: Projection -> Elims -> Projection Source # | |||||
NFData Projection Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: Projection -> () # | |||||
Generic Projection Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
| |||||
Show Projection Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods showsPrec :: Int -> Projection -> ShowS # show :: Projection -> String # showList :: [Projection] -> ShowS # | |||||
type Rep Projection Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep Projection = D1 ('MetaData "Projection" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "Projection" 'PrefixI 'True) ((S1 ('MetaSel ('Just "projProper") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe QName)) :*: S1 ('MetaSel ('Just "projOrig") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :*: (S1 ('MetaSel ('Just "projFromType") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Arg QName)) :*: (S1 ('MetaSel ('Just "projIndex") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Int) :*: S1 ('MetaSel ('Just "projLams") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ProjLams))))) |
data ProjectionLikenessMissing Source #
Indicates the reason behind a function having not been marked projection-like.
Constructors
MaybeProjection | Projection-likeness analysis has not run on this function yet. It may do so in the future. |
NeverProjection | The user has requested that this function be not be marked projection-like. The analysis may already have run on this function, but the results have been discarded, and it will not be run again. |
Instances
Pretty ProjectionLikenessMissing Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods pretty :: ProjectionLikenessMissing -> Doc Source # prettyPrec :: Int -> ProjectionLikenessMissing -> Doc Source # prettyList :: [ProjectionLikenessMissing] -> Doc Source # | |||||
KillRange ProjectionLikenessMissing Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods | |||||
EmbPrj ProjectionLikenessMissing Source # | |||||
NFData ProjectionLikenessMissing Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: ProjectionLikenessMissing -> () # | |||||
Bounded ProjectionLikenessMissing Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Enum ProjectionLikenessMissing Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods succ :: ProjectionLikenessMissing -> ProjectionLikenessMissing # pred :: ProjectionLikenessMissing -> ProjectionLikenessMissing # toEnum :: Int -> ProjectionLikenessMissing # fromEnum :: ProjectionLikenessMissing -> Int # enumFrom :: ProjectionLikenessMissing -> [ProjectionLikenessMissing] # enumFromThen :: ProjectionLikenessMissing -> ProjectionLikenessMissing -> [ProjectionLikenessMissing] # enumFromTo :: ProjectionLikenessMissing -> ProjectionLikenessMissing -> [ProjectionLikenessMissing] # enumFromThenTo :: ProjectionLikenessMissing -> ProjectionLikenessMissing -> ProjectionLikenessMissing -> [ProjectionLikenessMissing] # | |||||
Generic ProjectionLikenessMissing Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
Methods from :: ProjectionLikenessMissing -> Rep ProjectionLikenessMissing x # to :: Rep ProjectionLikenessMissing x -> ProjectionLikenessMissing # | |||||
Show ProjectionLikenessMissing Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods showsPrec :: Int -> ProjectionLikenessMissing -> ShowS # show :: ProjectionLikenessMissing -> String # showList :: [ProjectionLikenessMissing] -> ShowS # | |||||
type Rep ProjectionLikenessMissing Source # | |||||
Defined in Agda.TypeChecking.Monad.Base |
class Monad m => ReadTCState (m :: Type -> Type) where Source #
Minimal complete definition
Nothing
Methods
getTCState :: m TCState Source #
default getTCState :: forall (t :: (Type -> Type) -> Type -> Type) (n :: Type -> Type). (MonadTrans t, ReadTCState n, t n ~ m) => m TCState Source #
locallyTCState :: Lens' TCState a -> (a -> a) -> m b -> m b Source #
default locallyTCState :: forall (t :: (Type -> Type) -> Type -> Type) (n :: Type -> Type) a b. (MonadTransControl t, ReadTCState n, t n ~ m) => Lens' TCState a -> (a -> a) -> m b -> m b Source #
withTCState :: (TCState -> TCState) -> m a -> m a Source #
Instances
ReadTCState IM Source # | |
Defined in Agda.Interaction.Monad | |
ReadTCState TerM Source # | |
Defined in Agda.Termination.Monad | |
ReadTCState ReduceM Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
ReadTCState NLM Source # | |
Defined in Agda.TypeChecking.Rewriting.NonLinMatch | |
ReadTCState m => ReadTCState (PureConversionT m) Source # | |
Defined in Agda.TypeChecking.Conversion.Pure Methods getTCState :: PureConversionT m TCState Source # locallyTCState :: Lens' TCState a -> (a -> a) -> PureConversionT m b -> PureConversionT m b Source # withTCState :: (TCState -> TCState) -> PureConversionT m a -> PureConversionT m a Source # | |
ReadTCState m => ReadTCState (BlockT m) Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
MonadIO m => ReadTCState (TCMT m) Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
ReadTCState m => ReadTCState (NamesT m) Source # | |
Defined in Agda.TypeChecking.Names | |
ReadTCState m => ReadTCState (ListT m) Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
ReadTCState m => ReadTCState (ChangeT m) Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
ReadTCState m => ReadTCState (MaybeT m) Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
ReadTCState m => ReadTCState (ExceptT err m) Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
ReadTCState m => ReadTCState (IdentityT m) Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
ReadTCState m => ReadTCState (ReaderT r m) Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
ReadTCState m => ReadTCState (StateT s m) Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
(Monoid w, ReadTCState m) => ReadTCState (WriterT w m) Source # | |
Defined in Agda.TypeChecking.Monad.Base |
data RecordData Source #
Constructors
RecordData | |
Fields
|
Instances
Pretty RecordData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods pretty :: RecordData -> Doc Source # prettyPrec :: Int -> RecordData -> Doc Source # prettyList :: [RecordData] -> Doc Source # | |||||
NFData RecordData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: RecordData -> () # | |||||
Generic RecordData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
| |||||
Show RecordData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods showsPrec :: Int -> RecordData -> ShowS # show :: RecordData -> String # showList :: [RecordData] -> ShowS # | |||||
type Rep RecordData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep RecordData = D1 ('MetaData "RecordData" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "RecordData" 'PrefixI 'True) (((S1 ('MetaSel ('Just "_recPars") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Nat) :*: (S1 ('MetaSel ('Just "_recClause") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe Clause)) :*: S1 ('MetaSel ('Just "_recConHead") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ConHead))) :*: (S1 ('MetaSel ('Just "_recNamedCon") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool) :*: (S1 ('MetaSel ('Just "_recFields") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Dom QName]) :*: S1 ('MetaSel ('Just "_recTel") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Telescope)))) :*: ((S1 ('MetaSel ('Just "_recMutual") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe [QName])) :*: (S1 ('MetaSel ('Just "_recEtaEquality'") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 EtaEquality) :*: S1 ('MetaSel ('Just "_recPatternMatching") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 PatternOrCopattern))) :*: ((S1 ('MetaSel ('Just "_recInduction") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe Induction)) :*: S1 ('MetaSel ('Just "_recTerminates") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe Bool))) :*: (S1 ('MetaSel ('Just "_recAbstr") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 IsAbstract) :*: S1 ('MetaSel ('Just "_recComp") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 CompKit)))))) |
data ReduceDefs Source #
Constructors
OnlyReduceDefs (Set QName) | |
DontReduceDefs (Set QName) |
Instances
NFData ReduceDefs Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: ReduceDefs -> () # | |||||
Generic ReduceDefs Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
| |||||
type Rep ReduceDefs Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep ReduceDefs = D1 ('MetaData "ReduceDefs" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "OnlyReduceDefs" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Set QName))) :+: C1 ('MetaCons "DontReduceDefs" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Set QName)))) |
Environment of the reduce monad.
Constructors
ReduceEnv | |
Instances
HasOptions ReduceM Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
MonadReduce ReduceM Source # | |
Defined in Agda.TypeChecking.Monad.Base Methods liftReduce :: ReduceM a -> ReduceM a Source # | |
MonadTCEnv ReduceM Source # | |
ReadTCState ReduceM Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
HasBuiltins ReduceM Source # | |
Defined in Agda.TypeChecking.Reduce.Monad Methods getBuiltinThing :: SomeBuiltin -> ReduceM (Maybe (Builtin PrimFun)) Source # | |
MonadAddContext ReduceM Source # | |
Defined in Agda.TypeChecking.Reduce.Monad Methods addCtx :: Name -> Dom Type -> ReduceM a -> ReduceM a Source # addLetBinding' :: Origin -> Name -> Term -> Dom Type -> ReduceM a -> ReduceM a Source # updateContext :: Substitution -> (Context -> Context) -> ReduceM a -> ReduceM a Source # withFreshName :: Range -> ArgName -> (Name -> ReduceM a) -> ReduceM a Source # | |
MonadDebug ReduceM Source # | |
Defined in Agda.TypeChecking.Reduce.Monad Methods formatDebugMessage :: VerboseKey -> VerboseLevel -> TCM Doc -> ReduceM String Source # traceDebugMessage :: VerboseKey -> VerboseLevel -> String -> ReduceM a -> ReduceM a Source # verboseBracket :: VerboseKey -> VerboseLevel -> String -> ReduceM a -> ReduceM a Source # getVerbosity :: ReduceM Verbosity Source # getProfileOptions :: ReduceM ProfileOptions Source # isDebugPrinting :: ReduceM Bool Source # nowDebugPrinting :: ReduceM a -> ReduceM a Source # | |
PureTCM ReduceM Source # | |
Defined in Agda.TypeChecking.Reduce.Monad | |
HasConstInfo ReduceM Source # | |
Defined in Agda.TypeChecking.Reduce.Monad Methods getConstInfo :: QName -> ReduceM Definition Source # getConstInfo' :: QName -> ReduceM (Either SigError Definition) Source # getRewriteRulesFor :: QName -> ReduceM RewriteRules Source # | |
Applicative ReduceM Source # | |
Functor ReduceM Source # | |
Monad ReduceM Source # | |
MonadFail ReduceM Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
(Semigroup a, Monoid a) => Monoid (OccM a) Source # | |
Constructors
NoReduction no | |
YesReduction Simplification yes |
type RemoteMetaStore = HashMap MetaId RemoteMetaVariable Source #
Used for meta-variables from other modules (and in Interface
s).
data RemoteMetaVariable Source #
Information about remote meta-variables.
Remote meta-variables are meta-variables originating in other modules. These meta-variables are always instantiated. We do not retain all the information about a local meta-variable when creating an interface:
- The
mvPriority
field is not needed, because the meta-variable cannot be instantiated. - The
mvFrozen
field is not needed, because there is no point in freezing instantiated meta-variables. - The
mvListeners
field is not needed, because no meta-variable should be listening to this one. - The
mvTwin
field is not needed, because the meta-variable has already been instantiated. - The
mvPermutation
is currently removed, but could be retained if it turns out to be useful for something. - The only part of the
mvInfo
field that is kept is themiModality
field. ThemiMetaOccursCheck
andmiGeneralizable
fields are omitted, because the meta-variable has already been instantiated. TheRange
that is part of themiClosRange
field and themiNameSuggestion
field are omitted because instantiated meta-variables are typically not presented to users. Finally theClosure
part of themiClosRange
field is omitted because it can be large (at least if we ignore potential sharing).
Constructors
RemoteMetaVariable | |
Fields |
Instances
LensModalPolarity RemoteMetaVariable Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
LensModality RemoteMetaVariable Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods getModality :: RemoteMetaVariable -> Modality Source # setModality :: Modality -> RemoteMetaVariable -> RemoteMetaVariable Source # mapModality :: (Modality -> Modality) -> RemoteMetaVariable -> RemoteMetaVariable Source # | |||||
LensQuantity RemoteMetaVariable Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods getQuantity :: RemoteMetaVariable -> Quantity Source # setQuantity :: Quantity -> RemoteMetaVariable -> RemoteMetaVariable Source # mapQuantity :: (Quantity -> Quantity) -> RemoteMetaVariable -> RemoteMetaVariable Source # | |||||
LensRelevance RemoteMetaVariable Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods getRelevance :: RemoteMetaVariable -> Relevance Source # setRelevance :: Relevance -> RemoteMetaVariable -> RemoteMetaVariable Source # mapRelevance :: (Relevance -> Relevance) -> RemoteMetaVariable -> RemoteMetaVariable Source # | |||||
InstantiateFull RemoteMetaVariable Source # | |||||
Defined in Agda.TypeChecking.Reduce Methods instantiateFull' :: RemoteMetaVariable -> ReduceM RemoteMetaVariable Source # | |||||
EmbPrj RemoteMetaVariable Source # | |||||
NFData RemoteMetaVariable Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: RemoteMetaVariable -> () # | |||||
Generic RemoteMetaVariable Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
Methods from :: RemoteMetaVariable -> Rep RemoteMetaVariable x # to :: Rep RemoteMetaVariable x -> RemoteMetaVariable # | |||||
Show RemoteMetaVariable Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods showsPrec :: Int -> RemoteMetaVariable -> ShowS # show :: RemoteMetaVariable -> String # showList :: [RemoteMetaVariable] -> ShowS # | |||||
type Rep RemoteMetaVariable Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep RemoteMetaVariable = D1 ('MetaData "RemoteMetaVariable" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "RemoteMetaVariable" 'PrefixI 'True) (S1 ('MetaSel ('Just "rmvInstantiation") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Instantiation) :*: (S1 ('MetaSel ('Just "rmvModality") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Modality) :*: S1 ('MetaSel ('Just "rmvJudgement") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Judgement MetaId))))) |
data RewriteRule Source #
Rewrite rules can be added independently from function clauses.
Constructors
RewriteRule | |
Fields
|
Instances
NamesIn RewriteRule Source # | |||||
Defined in Agda.Syntax.Internal.Names Methods namesAndMetasIn' :: Monoid m => (Either QName MetaId -> m) -> RewriteRule -> m Source # | |||||
KillRange RewriteRule Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods | |||||
KillRange RewriteRuleMap Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods | |||||
PrettyTCM RewriteRule Source # | |||||
Defined in Agda.TypeChecking.Pretty Methods prettyTCM :: MonadPretty m => RewriteRule -> m Doc Source # | |||||
InstantiateFull RewriteRule Source # | |||||
Defined in Agda.TypeChecking.Reduce Methods instantiateFull' :: RewriteRule -> ReduceM RewriteRule Source # | |||||
GetMatchables RewriteRule Source # | |||||
Defined in Agda.TypeChecking.Rewriting.NonLinPattern Methods getMatchables :: RewriteRule -> [QName] Source # | |||||
EmbPrj RewriteRule Source # | |||||
Abstract RewriteRule Source # |
| ||||
Defined in Agda.TypeChecking.Substitute Methods abstract :: Telescope -> RewriteRule -> RewriteRule Source # | |||||
Apply RewriteRule Source # | |||||
Defined in Agda.TypeChecking.Substitute Methods apply :: RewriteRule -> Args -> RewriteRule Source # applyE :: RewriteRule -> Elims -> RewriteRule Source # | |||||
Subst RewriteRule Source # | |||||
Defined in Agda.TypeChecking.Substitute Associated Types
Methods applySubst :: Substitution' (SubstArg RewriteRule) -> RewriteRule -> RewriteRule Source # | |||||
NFData RewriteRule Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: RewriteRule -> () # | |||||
Generic RewriteRule Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
| |||||
Show RewriteRule Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods showsPrec :: Int -> RewriteRule -> ShowS # show :: RewriteRule -> String # showList :: [RewriteRule] -> ShowS # | |||||
type SubstArg RewriteRule Source # | |||||
Defined in Agda.TypeChecking.Substitute | |||||
type Rep RewriteRule Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep RewriteRule = D1 ('MetaData "RewriteRule" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "RewriteRule" 'PrefixI 'True) ((S1 ('MetaSel ('Just "rewName") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: (S1 ('MetaSel ('Just "rewContext") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Telescope) :*: S1 ('MetaSel ('Just "rewHead") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName))) :*: ((S1 ('MetaSel ('Just "rewPats") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 PElims) :*: S1 ('MetaSel ('Just "rewRHS") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)) :*: (S1 ('MetaSel ('Just "rewType") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type) :*: S1 ('MetaSel ('Just "rewFromClause") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool))))) |
type RewriteRuleMap = HashMap QName RewriteRules Source #
type RewriteRules = [RewriteRule] Source #
data RunMetaOccursCheck Source #
Constructors
RunMetaOccursCheck | |
DontRunMetaOccursCheck |
Instances
NFData RunMetaOccursCheck Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: RunMetaOccursCheck -> () # | |||||
Generic RunMetaOccursCheck Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
Methods from :: RunMetaOccursCheck -> Rep RunMetaOccursCheck x # to :: Rep RunMetaOccursCheck x -> RunMetaOccursCheck # | |||||
Show RunMetaOccursCheck Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods showsPrec :: Int -> RunMetaOccursCheck -> ShowS # show :: RunMetaOccursCheck -> String # showList :: [RunMetaOccursCheck] -> ShowS # | |||||
Eq RunMetaOccursCheck Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods (==) :: RunMetaOccursCheck -> RunMetaOccursCheck -> Bool # (/=) :: RunMetaOccursCheck -> RunMetaOccursCheck -> Bool # | |||||
Ord RunMetaOccursCheck Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods compare :: RunMetaOccursCheck -> RunMetaOccursCheck -> Ordering # (<) :: RunMetaOccursCheck -> RunMetaOccursCheck -> Bool # (<=) :: RunMetaOccursCheck -> RunMetaOccursCheck -> Bool # (>) :: RunMetaOccursCheck -> RunMetaOccursCheck -> Bool # (>=) :: RunMetaOccursCheck -> RunMetaOccursCheck -> Bool # max :: RunMetaOccursCheck -> RunMetaOccursCheck -> RunMetaOccursCheck # min :: RunMetaOccursCheck -> RunMetaOccursCheck -> RunMetaOccursCheck # | |||||
type Rep RunMetaOccursCheck Source # | |||||
Defined in Agda.TypeChecking.Monad.Base |
Constructors
Section | |
Fields |
Instances
Pretty Section Source # | |
NamesIn Section Source # | |
Defined in Agda.Syntax.Internal.Names | |
KillRange Section Source # | |
Defined in Agda.TypeChecking.Monad.Base Methods | |
KillRange Sections Source # | |
Defined in Agda.TypeChecking.Monad.Base Methods | |
InstantiateFull Section Source # | |
Defined in Agda.TypeChecking.Reduce | |
EmbPrj Section Source # | |
NFData Section Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
Show Section Source # | |
Eq Section Source # | |
data SessionTCState Source #
A part of the state that grows monotonically over the whole Agda session. Never reset.
Constructors
SessionTCState | |
Fields
|
Instances
NFData SessionTCState Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: SessionTCState -> () # | |||||
Generic SessionTCState Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
Methods from :: SessionTCState -> Rep SessionTCState x # to :: Rep SessionTCState x -> SessionTCState # | |||||
type Rep SessionTCState Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep SessionTCState = D1 ('MetaData "SessionTCState" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "SessionTCState" 'PrefixI 'True) (S1 ('MetaSel ('Just "stSessionBenchmark") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Benchmark) :*: (S1 ('MetaSel ('Just "stSessionBackends") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Backend]) :*: S1 ('MetaSel ('Just "stSessionFileDict") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 FileDictWithBuiltins)))) |
Constructors
Sig | |
Fields
|
Instances
KillRange Signature Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods | |||||
InstantiateFull Signature Source # | |||||
Defined in Agda.TypeChecking.Reduce | |||||
EmbPrj Signature Source # | |||||
NFData Signature Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Generic Signature Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
| |||||
Show Signature Source # | |||||
type Rep Signature Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep Signature = D1 ('MetaData "Signature" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "Sig" 'PrefixI 'True) ((S1 ('MetaSel ('Just "_sigSections") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Sections) :*: S1 ('MetaSel ('Just "_sigDefinitions") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Definitions)) :*: (S1 ('MetaSel ('Just "_sigRewriteRules") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 RewriteRuleMap) :*: S1 ('MetaSel ('Just "_sigInstances") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 InstanceTable)))) |
data Simplification Source #
Did we encounter a simplifying reduction? In terms of CIC, that would be a iota-reduction. In terms of Agda, this is a constructor or literal pattern that matched. Just beta-reduction (substitution) or delta-reduction (unfolding of definitions) does not count as simplifying?
Constructors
YesSimplification | |
NoSimplification |
Instances
Null Simplification Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
NFData Simplification Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: Simplification -> () # | |||||
Monoid Simplification Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods mappend :: Simplification -> Simplification -> Simplification # mconcat :: [Simplification] -> Simplification # | |||||
Semigroup Simplification Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods (<>) :: Simplification -> Simplification -> Simplification # sconcat :: NonEmpty Simplification -> Simplification # stimes :: Integral b => b -> Simplification -> Simplification # | |||||
Generic Simplification Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
Methods from :: Simplification -> Rep Simplification x # to :: Rep Simplification x -> Simplification # | |||||
Show Simplification Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods showsPrec :: Int -> Simplification -> ShowS # show :: Simplification -> String # showList :: [Simplification] -> ShowS # | |||||
Eq Simplification Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods (==) :: Simplification -> Simplification -> Bool # (/=) :: Simplification -> Simplification -> Bool # | |||||
type Rep Simplification Source # | |||||
Defined in Agda.TypeChecking.Monad.Base |
data SplitError Source #
Error when splitting a pattern variable into possible constructor patterns.
Constructors
NotADatatype (Closure Type) | Neither data type nor record. |
BlockedType Blocker (Closure Type) | Type could not be sufficiently reduced. |
ErasedDatatype ErasedDatatypeReason (Closure Type) | Data type, but in erased position. |
CoinductiveDatatype (Closure Type) | Split on codata not allowed. UNUSED, but keep! -- | NoRecordConstructor Type -- ^ record type, but no constructor |
UnificationStuck | |
Fields
| |
CosplitCatchall | Copattern split with a catchall |
CosplitNoTarget | We do not know the target type of the clause. |
CosplitNoRecordType (Closure Type) | Target type is not a record type. |
CannotCreateMissingClause QName (Telescope, [NamedArg DeBruijnPattern]) Doc (Closure (Abs Type)) | |
GenericSplitError String |
Instances
PrettyTCM SplitError Source # | |||||
Defined in Agda.TypeChecking.Errors Methods prettyTCM :: MonadPretty m => SplitError -> m Doc Source # | |||||
NFData SplitError Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: SplitError -> () # | |||||
Generic SplitError Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
| |||||
Show SplitError Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods showsPrec :: Int -> SplitError -> ShowS # show :: SplitError -> String # showList :: [SplitError] -> ShowS # | |||||
type Rep SplitError Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep SplitError = D1 ('MetaData "SplitError" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (((C1 ('MetaCons "NotADatatype" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Closure Type))) :+: C1 ('MetaCons "BlockedType" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Blocker) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Closure Type)))) :+: (C1 ('MetaCons "ErasedDatatype" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ErasedDatatypeReason) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Closure Type))) :+: (C1 ('MetaCons "CoinductiveDatatype" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Closure Type))) :+: C1 ('MetaCons "UnificationStuck" 'PrefixI 'True) ((S1 ('MetaSel ('Just "cantSplitBlocker") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe Blocker)) :*: (S1 ('MetaSel ('Just "cantSplitConName") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Just "cantSplitTel") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Telescope))) :*: (S1 ('MetaSel ('Just "cantSplitConIdx") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Args) :*: (S1 ('MetaSel ('Just "cantSplitGivenIdx") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Args) :*: S1 ('MetaSel ('Just "cantSplitFailures") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [UnificationFailure]))))))) :+: ((C1 ('MetaCons "CosplitCatchall" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "CosplitNoTarget" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "CosplitNoRecordType" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Closure Type))) :+: (C1 ('MetaCons "CannotCreateMissingClause" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Telescope, [NamedArg DeBruijnPattern]))) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Doc) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Closure (Abs Type))))) :+: C1 ('MetaCons "GenericSplitError" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String)))))) |
An alternative representation of partial elements in a telescope: Γ ⊢ λ Δ. [φ₁ u₁, ... , φₙ uₙ] : Δ → PartialP (∨_ᵢ φᵢ) T see cubicaltt paper (however we do not store the type T).
Constructors
System | |
Fields
|
Instances
NamesIn System Source # | |||||
Defined in Agda.Syntax.Internal.Names | |||||
KillRange System Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods | |||||
InstantiateFull System Source # | |||||
Defined in Agda.TypeChecking.Reduce | |||||
EmbPrj System Source # | |||||
Abstract System Source # | |||||
Apply System Source # | |||||
NFData System Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Generic System Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
| |||||
Show System Source # | |||||
Reify (QNamed System) Source # | |||||
Defined in Agda.Syntax.Translation.InternalToAbstract Associated Types
| |||||
type Rep System Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep System = D1 ('MetaData "System" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "System" 'PrefixI 'True) (S1 ('MetaSel ('Just "systemTel") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Telescope) :*: S1 ('MetaSel ('Just "systemClauses") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [(Face, Term)]))) | |||||
type ReifiesTo (QNamed System) Source # | |||||
Defined in Agda.Syntax.Translation.InternalToAbstract |
Constructors
TCEnv | |
Fields
|
Instances
LensIsAbstract TCEnv Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods | |||||
LensIsOpaque TCEnv Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
LensTCEnv TCEnv Source # | |||||
NFData TCEnv Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Generic TCEnv Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
| |||||
type Rep TCEnv Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep TCEnv = D1 ('MetaData "TCEnv" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "TCEnv" 'PrefixI 'True) (((((S1 ('MetaSel ('Just "envContext") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Context) :*: (S1 ('MetaSel ('Just "envLetBindings") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 LetBindings) :*: S1 ('MetaSel ('Just "envCurrentModule") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ModuleName))) :*: (S1 ('MetaSel ('Just "envCurrentPath") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe FileId)) :*: (S1 ('MetaSel ('Just "envAnonymousModules") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [(ModuleName, Nat)]) :*: S1 ('MetaSel ('Just "envImportPath") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [TopLevelModuleName])))) :*: ((S1 ('MetaSel ('Just "envMutualBlock") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe MutualId)) :*: (S1 ('MetaSel ('Just "envTerminationCheck") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (TerminationCheck ())) :*: S1 ('MetaSel ('Just "envCoverageCheck") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 CoverageCheck))) :*: ((S1 ('MetaSel ('Just "envMakeCase") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool) :*: S1 ('MetaSel ('Just "envSolvingConstraints") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool)) :*: (S1 ('MetaSel ('Just "envCheckingWhere") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 WhereClause_) :*: S1 ('MetaSel ('Just "envWorkingOnTypes") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool))))) :*: (((S1 ('MetaSel ('Just "envAssignMetas") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool) :*: (S1 ('MetaSel ('Just "envActiveProblems") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Set ProblemId)) :*: S1 ('MetaSel ('Just "envUnquoteProblem") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe ProblemId)))) :*: ((S1 ('MetaSel ('Just "envAbstractMode") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 AbstractMode) :*: S1 ('MetaSel ('Just "envRelevance") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Relevance)) :*: (S1 ('MetaSel ('Just "envQuantity") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Quantity) :*: S1 ('MetaSel ('Just "envHardCompileTimeMode") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool)))) :*: ((S1 ('MetaSel ('Just "envSplitOnStrict") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool) :*: (S1 ('MetaSel ('Just "envDisplayFormsEnabled") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool) :*: S1 ('MetaSel ('Just "envFoldLetBindings") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool))) :*: ((S1 ('MetaSel ('Just "envRange") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range) :*: S1 ('MetaSel ('Just "envHighlightingRange") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range)) :*: (S1 ('MetaSel ('Just "envClause") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 IPClause) :*: S1 ('MetaSel ('Just "envCall") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe (Closure Call)))))))) :*: ((((S1 ('MetaSel ('Just "envHighlightingLevel") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 HighlightingLevel) :*: (S1 ('MetaSel ('Just "envHighlightingMethod") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 HighlightingMethod) :*: S1 ('MetaSel ('Just "envExpandLast") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ExpandHidden))) :*: ((S1 ('MetaSel ('Just "envAppDef") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe QName)) :*: S1 ('MetaSel ('Just "envSimplification") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Simplification)) :*: (S1 ('MetaSel ('Just "envAllowedReductions") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 AllowedReductions) :*: S1 ('MetaSel ('Just "envReduceDefs") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ReduceDefs)))) :*: ((S1 ('MetaSel ('Just "envReconstructed") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool) :*: (S1 ('MetaSel ('Just "envInjectivityDepth") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Int) :*: S1 ('MetaSel ('Just "envCompareBlocked") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool))) :*: ((S1 ('MetaSel ('Just "envPrintDomainFreePi") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool) :*: S1 ('MetaSel ('Just "envPrintMetasBare") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool)) :*: (S1 ('MetaSel ('Just "envInsideDotPattern") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool) :*: S1 ('MetaSel ('Just "envUnquoteFlags") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 UnquoteFlags))))) :*: (((S1 ('MetaSel ('Just "envInstanceDepth") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Int) :*: (S1 ('MetaSel ('Just "envIsDebugPrinting") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool) :*: S1 ('MetaSel ('Just "envPrintingPatternLambdas") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [QName]))) :*: ((S1 ('MetaSel ('Just "envCallByNeed") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool) :*: S1 ('MetaSel ('Just "envCurrentCheckpoint") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 CheckpointId)) :*: (S1 ('MetaSel ('Just "envCheckpoints") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Map CheckpointId Substitution)) :*: S1 ('MetaSel ('Just "envGeneralizeMetas") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 DoGeneralize)))) :*: ((S1 ('MetaSel ('Just "envGeneralizedVars") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Map QName GeneralizedValue)) :*: (S1 ('MetaSel ('Just "envActiveBackendName") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe BackendName)) :*: S1 ('MetaSel ('Just "envConflComputingOverlap") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool))) :*: ((S1 ('MetaSel ('Just "envCurrentlyElaborating") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool) :*: S1 ('MetaSel ('Just "envSyntacticEqualityFuel") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (Maybe Int))) :*: (S1 ('MetaSel ('Just "envCurrentOpaqueId") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (Maybe OpaqueId)) :*: S1 ('MetaSel ('Just "envTermCheckReducing") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool)))))))) |
Type-checking errors.
Constructors
TypeError | |
Fields
| |
ParserError ParseError | Error raised by the Happy parser. |
GenericException String | Unspecific error without |
IOException (Maybe TCState) Range IOException | The first argument is the state in which the error was raised. |
PatternErr Blocker | The exception which is usually caught.
Raised for pattern violations during unification ( |
Instances
EncodeTCM DisplayInfo Source # | |||||
Defined in Agda.Interaction.JSONTop | |||||
EncodeTCM Info_Error Source # | |||||
Defined in Agda.Interaction.JSONTop | |||||
EncodeTCM Response Source # | |||||
EncodeTCM TCErr Source # | |||||
HasRange TCErr Source # | |||||
PrettyTCM TCErr Source # | |||||
Defined in Agda.TypeChecking.Errors | |||||
NFData TCErr Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Exception TCErr Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods toException :: TCErr -> SomeException # fromException :: SomeException -> Maybe TCErr # displayException :: TCErr -> String # backtraceDesired :: TCErr -> Bool # | |||||
Show TCErr Source # | |||||
MonadError TCErr IM Source # | |||||
Defined in Agda.Interaction.Monad | |||||
MonadError TCErr TerM Source # | |||||
Defined in Agda.Termination.Monad | |||||
Monad m => MonadError TCErr (PureConversionT m) Source # | |||||
Defined in Agda.TypeChecking.Conversion.Pure Methods throwError :: TCErr -> PureConversionT m a # catchError :: PureConversionT m a -> (TCErr -> PureConversionT m a) -> PureConversionT m a # | |||||
(CatchIO m, MonadIO m) => MonadError TCErr (TCMT m) Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
EncodeTCM (OutputForm Expr Expr) Source # | |||||
Defined in Agda.Interaction.JSONTop | |||||
(Pretty a, Pretty b) => Pretty (OutputConstraint a b) Source # | |||||
Defined in Agda.Interaction.BasicOps Methods pretty :: OutputConstraint a b -> Doc Source # prettyPrec :: Int -> OutputConstraint a b -> Doc Source # prettyList :: [OutputConstraint a b] -> Doc Source # | |||||
(Pretty a, Pretty b) => Pretty (OutputForm a b) Source # | |||||
Defined in Agda.Interaction.BasicOps Methods pretty :: OutputForm a b -> Doc Source # prettyPrec :: Int -> OutputForm a b -> Doc Source # prettyList :: [OutputForm a b] -> Doc Source # | |||||
(ToConcrete a, ToConcrete b) => ToConcrete (OutputConstraint a b) Source # | |||||
Defined in Agda.Interaction.BasicOps Associated Types
Methods toConcrete :: MonadToConcrete m => OutputConstraint a b -> m (ConOfAbs (OutputConstraint a b)) Source # bindToConcrete :: MonadToConcrete m => OutputConstraint a b -> (ConOfAbs (OutputConstraint a b) -> m b0) -> m b0 Source # | |||||
(ToConcrete a, ToConcrete b) => ToConcrete (OutputForm a b) Source # | |||||
Defined in Agda.Interaction.BasicOps Associated Types
Methods toConcrete :: MonadToConcrete m => OutputForm a b -> m (ConOfAbs (OutputForm a b)) Source # bindToConcrete :: MonadToConcrete m => OutputForm a b -> (ConOfAbs (OutputForm a b) -> m b0) -> m b0 Source # | |||||
Monad m => MonadBlock (ExceptT TCErr m) Source # | |||||
(Pretty a, Pretty b) => PrettyTCM (OutputForm a b) Source # | |||||
Defined in Agda.Interaction.BasicOps Methods prettyTCM :: MonadPretty m => OutputForm a b -> m Doc Source # | |||||
type ConOfAbs (OutputConstraint a b) Source # | |||||
Defined in Agda.Interaction.BasicOps | |||||
type ConOfAbs (OutputForm a b) Source # | |||||
Defined in Agda.Interaction.BasicOps |
newtype TCMT (m :: Type -> Type) a Source #
Instances
MonadFixityError ScopeM Source # | |||||
Defined in Agda.Syntax.Scope.Monad Methods throwMultipleFixityDecls :: List1 (Name, Pair Fixity') -> ScopeM a Source # throwMultiplePolarityPragmas :: List1 Name -> ScopeM a Source # warnUnknownNamesInFixityDecl :: Set1 Name -> ScopeM () Source # warnUnknownNamesInPolarityPragmas :: Set1 Name -> ScopeM () Source # warnUnknownFixityInMixfixDecl :: Set1 Name -> ScopeM () Source # warnPolarityPragmasButNotPostulates :: Set1 Name -> ScopeM () Source # warnEmptyPolarityPragma :: Range -> ScopeM () Source # | |||||
MonadBlock TCM Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
MonadStConcreteNames TCM Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods runStConcreteNames :: StateT ConcreteNames TCM a -> TCM a Source # useConcreteNames :: TCM ConcreteNames Source # modifyConcreteNames :: (ConcreteNames -> ConcreteNames) -> TCM () Source # | |||||
MonadConstraint TCM Source # | |||||
Defined in Agda.TypeChecking.Constraints Methods addConstraint :: Blocker -> Constraint -> TCM () Source # addAwakeConstraint :: Blocker -> Constraint -> TCM () Source # solveConstraint :: Constraint -> TCM () Source # solveSomeAwakeConstraints :: (ProblemConstraint -> Bool) -> Bool -> TCM () Source # wakeConstraints :: (ProblemConstraint -> WakeUp) -> TCM () Source # stealConstraints :: ProblemId -> TCM () Source # modifyAwakeConstraints :: (Constraints -> Constraints) -> TCM () Source # modifySleepingConstraints :: (Constraints -> Constraints) -> TCM () Source # | |||||
MonadAddContext TCM Source # | |||||
Defined in Agda.TypeChecking.Monad.Context | |||||
MonadDebug TCM Source # | |||||
Defined in Agda.TypeChecking.Monad.Debug Methods formatDebugMessage :: VerboseKey -> VerboseLevel -> TCM Doc -> TCM String Source # traceDebugMessage :: VerboseKey -> VerboseLevel -> String -> TCM a -> TCM a Source # verboseBracket :: VerboseKey -> VerboseLevel -> String -> TCM a -> TCM a Source # getVerbosity :: TCM Verbosity Source # getProfileOptions :: TCM ProfileOptions Source # isDebugPrinting :: TCM Bool Source # nowDebugPrinting :: TCM a -> TCM a Source # | |||||
MonadInteractionPoints TCM Source # | |||||
Defined in Agda.TypeChecking.Monad.MetaVars Methods freshInteractionId :: TCM InteractionId Source # modifyInteractionPoints :: (InteractionPoints -> InteractionPoints) -> TCM () Source # | |||||
MonadMetaSolver TCM Source # | |||||
Defined in Agda.TypeChecking.MetaVars Methods newMeta' :: MetaInstantiation -> Frozen -> MetaInfo -> MetaPriority -> Permutation -> Judgement a -> TCM MetaId Source # assignV :: CompareDirection -> MetaId -> Args -> Term -> CompareAs -> TCM () Source # assignTerm' :: MetaId -> [Arg ArgName] -> Term -> TCM () Source # etaExpandMeta :: [MetaClass] -> MetaId -> TCM () Source # updateMetaVar :: MetaId -> (MetaVariable -> MetaVariable) -> TCM () Source # speculateMetas :: TCM () -> TCM KeepMetas -> TCM () Source # | |||||
PureTCM TCM Source # | |||||
Defined in Agda.TypeChecking.Monad.Pure | |||||
MonadStatistics TCM Source # | |||||
Defined in Agda.TypeChecking.Monad.Statistics | |||||
MonadTrace TCM Source # | |||||
Defined in Agda.TypeChecking.Monad.Trace Methods traceCall :: Call -> TCM a -> TCM a Source # traceCallM :: TCM Call -> TCM a -> TCM a Source # traceCallCPS :: Call -> ((a -> TCM b) -> TCM b) -> (a -> TCM b) -> TCM b Source # traceClosureCall :: Closure Call -> TCM a -> TCM a Source # printHighlightingInfo :: RemoveTokenBasedHighlighting -> HighlightingInfo -> TCM () Source # | |||||
MonadWarning TCM Source # | |||||
Defined in Agda.TypeChecking.Warnings | |||||
MonadBench TCM Source # | We store benchmark statistics in an IORef. This enables benchmarking pure computation, see Agda.Benchmarking. | ||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
Methods getBenchmark :: TCM (Benchmark (BenchPhase TCM)) Source # putBenchmark :: Benchmark (BenchPhase TCM) -> TCM () Source # modifyBenchmark :: (Benchmark (BenchPhase TCM) -> Benchmark (BenchPhase TCM)) -> TCM () Source # | |||||
CatchImpossible TCM Source # | Like The intended use is to catch internal errors during debug printing. In debug printing, we are not expecting state changes. | ||||
Defined in Agda.TypeChecking.Monad.Base Methods catchImpossible :: TCM a -> (Impossible -> TCM a) -> TCM a Source # catchImpossibleJust :: (Impossible -> Maybe b) -> TCM a -> (b -> TCM a) -> TCM a Source # handleImpossible :: (Impossible -> TCM a) -> TCM a -> TCM a Source # handleImpossibleJust :: (Impossible -> Maybe b) -> (b -> TCM a) -> TCM a -> TCM a Source # | |||||
MonadTrans TCMT Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
HasFresh i => MonadFresh i TCM Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
(CatchIO m, MonadIO m) => MonadError TCErr (TCMT m) Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
MonadIO m => HasOptions (TCMT m) Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods | |||||
MonadIO m => MonadReduce (TCMT m) Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods liftReduce :: ReduceM a -> TCMT m a Source # | |||||
MonadIO m => MonadTCEnv (TCMT m) Source # | |||||
MonadIO m => MonadTCM (TCMT m) Source # | |||||
MonadIO m => MonadTCState (TCMT m) Source # | |||||
MonadIO m => ReadTCState (TCMT m) Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
MonadIO m => HasBuiltins (TCMT m) Source # | |||||
Defined in Agda.TypeChecking.Monad.Builtin Methods getBuiltinThing :: SomeBuiltin -> TCMT m (Maybe (Builtin PrimFun)) Source # | |||||
ReportS (TCM Doc) Source # | |||||
Defined in Agda.TypeChecking.Monad.Debug Methods reportS :: MonadDebug m => VerboseKey -> VerboseLevel -> TCM Doc -> m () Source # | |||||
ReportS [TCM Doc] Source # | |||||
Defined in Agda.TypeChecking.Monad.Debug Methods reportS :: MonadDebug m => VerboseKey -> VerboseLevel -> [TCM Doc] -> m () Source # | |||||
TraceS (TCM Doc) Source # | |||||
Defined in Agda.TypeChecking.Monad.Debug Methods traceS :: MonadDebug m => VerboseKey -> VerboseLevel -> TCM Doc -> m c -> m c Source # | |||||
TraceS [TCM Doc] Source # | |||||
Defined in Agda.TypeChecking.Monad.Debug Methods traceS :: MonadDebug m => VerboseKey -> VerboseLevel -> [TCM Doc] -> m c -> m c Source # | |||||
HasConstInfo (TCMT IO) Source # | |||||
Defined in Agda.TypeChecking.Monad.Signature Methods getConstInfo :: QName -> TCMT IO Definition Source # getConstInfo' :: QName -> TCMT IO (Either SigError Definition) Source # getRewriteRulesFor :: QName -> TCMT IO RewriteRules Source # | |||||
MonadIO m => MonadFileId (TCMT m) Source # | |||||
Defined in Agda.TypeChecking.Monad.State | |||||
Null (TCM Doc) Source # | |||||
Applicative m => Applicative (TCMT m) Source # | |||||
Functor m => Functor (TCMT m) Source # | |||||
Monad m => Monad (TCMT m) Source # | |||||
Semigroup (TCM Doc) Source # | This instance is more specific than a generic instance
| ||||
(CatchIO m, MonadIO m) => MonadFail (TCMT m) Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
MonadIO m => MonadIO (TCMT m) Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
(MonadIO m, Null a) => Null (TCMT m a) Source # | |||||
(MonadIO m, Semigroup a, Monoid a) => Monoid (TCMT m a) Source # | Strict (non-shortcut) monoid. | ||||
(MonadIO m, Semigroup a) => Semigroup (TCMT m a) Source # | Strict (non-shortcut) semigroup. Note that there might be a lazy alternative, e.g.,
for TCM All we might want | ||||
(IsString a, MonadIO m) => IsString (TCMT m a) Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods fromString :: String -> TCMT m a # | |||||
type BenchPhase TCM Source # | |||||
Defined in Agda.TypeChecking.Monad.Base |
Constructors
TCSt | |
Fields
|
Instances
LensCommandLineOptions TCState Source # | |||||
Defined in Agda.Interaction.Options.Lenses Methods getCommandLineOptions :: TCState -> CommandLineOptions Source # setCommandLineOptions :: CommandLineOptions -> TCState -> TCState Source # mapCommandLineOptions :: (CommandLineOptions -> CommandLineOptions) -> TCState -> TCState Source # | |||||
LensIncludePaths TCState Source # | |||||
Defined in Agda.Interaction.Options.Lenses Methods getIncludePaths :: TCState -> [FilePath] Source # setIncludePaths :: [FilePath] -> TCState -> TCState Source # mapIncludePaths :: ([FilePath] -> [FilePath]) -> TCState -> TCState Source # getAbsoluteIncludePaths :: TCState -> [AbsolutePath] Source # setAbsoluteIncludePaths :: [AbsolutePath] -> TCState -> TCState Source # mapAbsoluteIncludePaths :: ([AbsolutePath] -> [AbsolutePath]) -> TCState -> TCState Source # | |||||
LensPersistentVerbosity TCState Source # | |||||
Defined in Agda.Interaction.Options.Lenses Methods getPersistentVerbosity :: TCState -> PersistentVerbosity Source # setPersistentVerbosity :: PersistentVerbosity -> TCState -> TCState Source # mapPersistentVerbosity :: (PersistentVerbosity -> PersistentVerbosity) -> TCState -> TCState Source # | |||||
LensSafeMode TCState Source # | |||||
Defined in Agda.Interaction.Options.Lenses | |||||
LensVerbosity TCState Source # | |||||
Defined in Agda.Interaction.Options.Lenses | |||||
LensPragmaOptions TCState Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods getPragmaOptions :: TCState -> PragmaOptions Source # setPragmaOptions :: PragmaOptions -> TCState -> TCState Source # mapPragmaOptions :: (PragmaOptions -> PragmaOptions) -> TCState -> TCState Source # | |||||
NFData TCState Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Generic TCState Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
| |||||
Show TCState Source # | |||||
type Rep TCState Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep TCState = D1 ('MetaData "TCState" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "TCSt" 'PrefixI 'True) (S1 ('MetaSel ('Just "stPersistentState") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 PersistentTCState) :*: (S1 ('MetaSel ('Just "stPreScopeState") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 PreScopeState) :*: S1 ('MetaSel ('Just "stPostScopeState") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 PostScopeState)))) |
Constructors
TCWarning | |
Fields
|
Instances
EncodeTCM DisplayInfo Source # | |||||
Defined in Agda.Interaction.JSONTop | |||||
EncodeTCM Info_Error Source # | |||||
Defined in Agda.Interaction.JSONTop | |||||
EncodeTCM Response Source # | |||||
EncodeTCM TCWarning Source # | |||||
HasRange TCWarning Source # | |||||
PrettyTCM TCWarning Source # | |||||
Defined in Agda.TypeChecking.Pretty.Warning | |||||
EmbPrj TCWarning Source # | |||||
NFData TCWarning Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Generic TCWarning Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
| |||||
Show TCWarning Source # | |||||
Eq TCWarning Source # | |||||
Ord TCWarning Source # | |||||
type Rep TCWarning Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep TCWarning = D1 ('MetaData "TCWarning" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "TCWarning" 'PrefixI 'True) ((S1 ('MetaSel ('Just "tcWarningLocation") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 CallStack) :*: (S1 ('MetaSel ('Just "tcWarningRange") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range) :*: S1 ('MetaSel ('Just "tcWarning") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Warning))) :*: (S1 ('MetaSel ('Just "tcWarningDoc") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Doc) :*: (S1 ('MetaSel ('Just "tcWarningString") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String) :*: S1 ('MetaSel ('Just "tcWarningCached") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool))))) |
type TempInstanceTable = (InstanceTable, Set QName) Source #
When typechecking something of the following form:
instance x : _ x = y
it's not yet known where to add x
, so we add it to a list of
unresolved instances and we'll deal with it later.
Instances
Pretty TermHead Source # | |||||
KillRange TermHead Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods | |||||
EmbPrj TermHead Source # | |||||
NFData TermHead Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Generic TermHead Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
| |||||
Show TermHead Source # | |||||
Eq TermHead Source # | |||||
Ord TermHead Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
type Rep TermHead Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep TermHead = D1 ('MetaData "TermHead" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) ((C1 ('MetaCons "SortHead" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "PiHead" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "ConsHead" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: (C1 ('MetaCons "VarHead" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Nat)) :+: C1 ('MetaCons "UnknownHead" 'PrefixI 'False) (U1 :: Type -> Type)))) |
data TerminationError Source #
Information about a mutual block which did not pass the termination checker.
Constructors
TerminationError | |
Fields
|
Instances
NFData TerminationError Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: TerminationError -> () # | |||||
Generic TerminationError Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
Methods from :: TerminationError -> Rep TerminationError x # to :: Rep TerminationError x -> TerminationError # | |||||
Show TerminationError Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods showsPrec :: Int -> TerminationError -> ShowS # show :: TerminationError -> String # showList :: [TerminationError] -> ShowS # | |||||
type Rep TerminationError Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep TerminationError = D1 ('MetaData "TerminationError" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "TerminationError" 'PrefixI 'True) (S1 ('MetaSel ('Just "termErrFunctions") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [QName]) :*: (S1 ('MetaSel ('Just "termErrCalls") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [CallInfo]) :*: S1 ('MetaSel ('Just "termErrGuardednessHelps") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 GuardednessHelps)))) |
data TypeCheckAction Source #
A complete log for a module will look like this:
Pragmas
EnterSection
, entering the main module.Decl
/EnterSection
/LeaveSection
, for declarations and nested modulesLeaveSection
, leaving the main module.
Constructors
EnterSection !Erased !ModuleName !Telescope | |
LeaveSection !ModuleName | |
Decl !Declaration | Never a Section or ScopeDecl |
Pragmas !PragmaOptions |
Instances
NFData TypeCheckAction Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: TypeCheckAction -> () # | |||||
Generic TypeCheckAction Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
Methods from :: TypeCheckAction -> Rep TypeCheckAction x # to :: Rep TypeCheckAction x -> TypeCheckAction # | |||||
type Rep TypeCheckAction Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep TypeCheckAction = D1 ('MetaData "TypeCheckAction" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) ((C1 ('MetaCons "EnterSection" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Erased) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 ModuleName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Telescope))) :+: C1 ('MetaCons "LeaveSection" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 ModuleName))) :+: (C1 ('MetaCons "Decl" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Declaration)) :+: C1 ('MetaCons "Pragmas" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 PragmaOptions)))) |
data TypeCheckingProblem Source #
Constructors
CheckExpr Comparison Expr Type | |
CheckArgs Comparison ExpandHidden Expr [NamedArg Expr] Type Type (ArgsCheckState CheckedTarget -> TCM Term) | |
CheckProjAppToKnownPrincipalArg Comparison Expr ProjOrigin (List1 QName) Expr Args Type Int Term Type PrincipalArgTypeMetas | |
CheckLambda Comparison (Arg (List1 (WithHiding Name), Maybe Type)) Expr Type |
|
DisambiguateConstructor ConstructorDisambiguationData (ConHead -> TCM Term) | A stuck constructor disambiguation with the bits to retry it on and the success continuation. |
DoQuoteTerm Comparison Term Type | Quote the given term and check type against |
Instances
PrettyTCM TypeCheckingProblem Source # | |||||
Defined in Agda.TypeChecking.Pretty Methods prettyTCM :: MonadPretty m => TypeCheckingProblem -> m Doc Source # | |||||
NFData TypeCheckingProblem Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: TypeCheckingProblem -> () # | |||||
Generic TypeCheckingProblem Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
Methods from :: TypeCheckingProblem -> Rep TypeCheckingProblem x # to :: Rep TypeCheckingProblem x -> TypeCheckingProblem # | |||||
type Rep TypeCheckingProblem Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep TypeCheckingProblem = D1 ('MetaData "TypeCheckingProblem" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) ((C1 ('MetaCons "CheckExpr" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type))) :+: (C1 ('MetaCons "CheckArgs" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ExpandHidden) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr))) :*: ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [NamedArg Expr]) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (ArgsCheckState CheckedTarget -> TCM Term))))) :+: C1 ('MetaCons "CheckProjAppToKnownPrincipalArg" 'PrefixI 'False) (((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ProjOrigin) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 QName)) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr)))) :*: ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Args) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Int))) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 PrincipalArgTypeMetas))))))) :+: (C1 ('MetaCons "CheckLambda" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Arg (List1 (WithHiding Name), Maybe Type)))) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type))) :+: (C1 ('MetaCons "DisambiguateConstructor" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ConstructorDisambiguationData) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (ConHead -> TCM Term))) :+: C1 ('MetaCons "DoQuoteTerm" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)))))) |
Constructors
InternalError String | |
NotImplemented String | |
NotSupported String | |
CompilationError String | |
SyntaxError String | Essential syntax error thrown after successful parsing.
Description in |
OptionError OptionError | Error thrown by the option parser. |
NicifierError DeclarationException' | Error thrown in the nicifier phase |
DoNotationError String | Error during unsugaring some |
IdiomBracketError String | Error during (operator) parsing and interpreting the contents of idiom brackets.
Error message in |
NoKnownRecordWithSuchFields [Name] | The user has given a record expression with the given fields, but no record type known to type inference has all these fields. The list can be empty. |
ShouldEndInApplicationOfTheDatatype Type | The target of a constructor isn't an application of its
datatype. The |
ConstructorPatternInWrongDatatype QName QName | constructor, datatype |
CantResolveOverloadedConstructorsTargetingSameDatatype QName (List1 QName) | Datatype, constructors. |
ConstructorDoesNotTargetGivenType QName Type | constructor, type |
InvalidDottedExpression |
|
LiteralTooBig | An integer literal that would be too costly to expand to unary. |
NegativeLiteralInPattern | Negative literals are not supported in patterns. |
WrongHidingInLHS | The left hand side of a function definition has a hidden argument where a non-hidden was expected. |
WrongHidingInLambda Type | Expected a non-hidden function and found a hidden lambda. |
WrongHidingInApplication Type | A function is applied to a hidden argument where a non-hidden was expected. |
WrongHidingInProjection QName | |
IllegalHidingInPostfixProjection (NamedArg Expr) | |
WrongNamedArgument (NamedArg Expr) (List1 NamedName) | A function is applied to a hidden named argument it does not have. The list contains names of possible hidden arguments at this point. |
WrongAnnotationInLambda | Wrong user-given (lock/tick) annotation in lambda. |
WrongIrrelevanceInLambda | Wrong user-given relevance annotation in lambda. |
WrongQuantityInLambda | Wrong user-given quantity annotation in lambda. |
WrongCohesionInLambda | Wrong user-given cohesion annotation in lambda. |
WrongPolarityInLambda | Wrong user-given polarity annotation in lambda. |
QuantityMismatch Quantity Quantity | The given quantity does not correspond to the expected quantity. |
HidingMismatch Hiding Hiding | The given hiding does not correspond to the expected hiding. |
RelevanceMismatch Relevance Relevance | The given relevance does not correspond to the expected relevance. |
ForcedConstructorNotInstantiated Pattern | |
IllformedProjectionPatternAbstract Pattern | |
IllformedProjectionPatternConcrete Pattern | |
CannotEliminateWithPattern (Maybe Blocker) (NamedArg Pattern) Type | |
CannotEliminateWithProjection (Arg Type) Bool QName | |
WrongNumberOfConstructorArguments QName Nat Nat | |
ShouldBeEmpty Type [DeBruijnPattern] | Type should be empty. The list gives possible patterns that match, but can be empty. |
ShouldBeASort Type | The given type should have been a sort. |
ShouldBePi Type | The given type should have been a pi. |
ShouldBePath Type | |
ShouldBeRecordType Type | |
ShouldBeRecordPattern DeBruijnPattern | |
CannotApply Expr Type | The given expression is used as a function but its type is not a function type. |
InvalidTypeSort Sort | This sort is not a type expression. |
SplitOnCoinductive | |
SplitOnIrrelevant (Dom Type) | |
SplitOnUnusableCohesion (Dom Type) | |
SplitOnUnusablePolarity (Dom Type) | |
SplitOnNonVariable Term Type | |
SplitOnNonEtaRecord QName | |
SplitOnAbstract QName | |
SplitOnUnchecked QName | |
SplitOnPartial (Dom Type) | |
SplitInProp DataOrRecordE | |
DefinitionIsIrrelevant QName | |
DefinitionIsErased QName | |
ProjectionIsIrrelevant QName | |
VariableIsIrrelevant Name | |
VariableIsErased Name | |
VariableIsOfUnusableCohesion Name Cohesion | |
LambdaIsErased | |
RecordIsErased | |
InvalidModalTelescopeUse Term Modality Modality Definition | |
VariableIsOfUnusablePolarity Name PolarityModality | |
UnequalLevel Comparison Level Level | |
UnequalTerms Comparison Term Term CompareAs | |
UnequalRelevance Comparison Term Term | The two function types have different relevance. |
UnequalQuantity Comparison Term Term | The two function types have different relevance. |
UnequalCohesion Comparison Term Term | The two function types have different cohesion. |
UnequalPolarity Comparison Term Term | The two function types have different polarity. |
UnequalFiniteness Comparison Term Term | One of the function types has a finite domain (i.e. is a |
UnequalHiding Term Term | The two function types have different hiding. |
UnequalSorts Sort Sort | |
NotLeqSort Sort Sort | |
MetaCannotDependOn MetaId Term Nat | The arguments are the meta variable, the proposed solution, and the parameter that it wants to depend on. |
MetaIrrelevantSolution MetaId Term | When solving |
MetaErasedSolution MetaId Term | When solving |
GenericError String | |
GenericDocError Doc | |
SortOfSplitVarError (Maybe Blocker) Doc | the meta is what we might be blocked on. |
WrongSharpArity QName | |
BuiltinMustBeConstructor BuiltinId Expr | |
BuiltinMustBeData BuiltinId Int | |
BuiltinMustBeDef BuiltinId | |
BuiltinMustBeFunction BuiltinId | |
BuiltinMustBePostulate BuiltinId | |
NoSuchBuiltinName String | |
InvalidBuiltin String | |
DuplicateBuiltinBinding BuiltinId Term Term | |
NoBindingForBuiltin BuiltinId | |
NoBindingForPrimitive PrimitiveId | |
NoSuchPrimitiveFunction String | |
DuplicatePrimitiveBinding PrimitiveId QName QName | |
WrongArgInfoForPrimitive PrimitiveId ArgInfo ArgInfo | |
ShadowedModule Name (List1 ModuleName) | |
BuiltinInParameterisedModule BuiltinId | |
IllegalDeclarationInDataDefinition (List1 Declaration) | The declaration list comes from a single |
IllegalLetInTelescope TypedBinding | |
IllegalPatternInTelescope Binder | |
AbsentRHSRequiresAbsurdPattern | |
TooManyFields QName [Name] (List1 Name) | Record type, fields not supplied by user, possibly non-fields but supplied. |
DuplicateFields (List1 Name) | |
DuplicateConstructors (List1 Name) | |
DuplicateOverlapPragma QName OverlapMode OverlapMode | |
WithOnFreeVariable Expr Term | |
UnexpectedWithPatterns (List1 Pattern) | |
WithClausePatternMismatch Pattern (NamedArg DeBruijnPattern) | |
IllTypedPatternAfterWithAbstraction Pattern | |
TooFewPatternsInWithClause | |
TooManyPatternsInWithClause | |
PathAbstractionFailed (Abs Type) | |
FieldOutsideRecord | |
ModuleArityMismatch ModuleName Telescope (Either (List1 (NamedArg Expr)) Args) | |
GeneralizeCyclicDependency | |
ReferencesFutureVariables Term (List1 Int) (Arg Term) Int | The first term references the given list of variables, which are in "the future" with respect to the given lock (and its leftmost variable) |
DoesNotMentionTicks Term Type (Arg Term) | Arguments: later term, its type, lock term. The lock term does not mention any @lock variables. |
MismatchedProjectionsError QName QName | |
AttributeKindNotEnabled String String String | |
InvalidProjectionParameter (NamedArg Expr) | |
TacticAttributeNotAllowed | |
CannotRewriteByNonEquation Type | |
MacroResultTypeMismatch Type | |
NamedWhereModuleInRefinedContext [Term] [String] | The lists should have the same length. TODO: enforce this by construction. |
ComatchingDisabledForRecord QName | |
IncorrectTypeForRewriteRelation Term IncorrectTypeForRewriteRelationReason | |
CannotGenerateHCompClause Type | Cannot generate |
CannotGenerateTransportClause QName (Closure (Abs Type)) | Cannot generate transport clause because type is not fibrant. |
CubicalNotErasure QName | Name was defined for |
CubicalPrimitiveNotFullyApplied QName | |
ExpectedIntervalLiteral Expr | Expected an interval literal (0 or 1) but found |
FaceConstraintDisjunction | |
FaceConstraintUnsatisfiable | |
PatternInPathLambda | Attempt to pattern match in an abstraction of interval type. |
PatternInSystem | Attempt to pattern or copattern match in a system. Data errors |
UnexpectedParameter LamBinding | |
NoParameterOfName ArgName | |
UnexpectedModalityAnnotationInParameter LamBinding | |
ExpectedBindingForParameter (Dom Type) (Abs Type) | |
UnexpectedTypeSignatureForParameter (List1 (NamedArg Binder)) | |
SortDoesNotAdmitDataDefinitions QName Sort | |
SortCannotDependOnItsIndex QName Type | |
UnusableAtModality WhyCheckModality Modality Term | |
SplitError SplitError | |
ImpossibleConstructor QName NegativeUnification | |
DatatypeIndexPolarity | An index of a data type has a polarity different from |
RecursiveRecordNeedsInductivity QName | A record type inferred as recursive needs a manual declaration whether it should be inductively or coinductively. |
CannotSolveSizeConstraints (List1 (ProblemConstraint, HypSizeConstraint)) Doc | The list of constraints is given redundantly as pairs of
|
ContradictorySizeConstraint (ProblemConstraint, HypSizeConstraint) | |
EmptyTypeOfSizes Term | This type, representing a type of sizes, might be empty. |
FunctionTypeInSizeUniv Term | This term, a function type constructor, lives in
|
PostulatedSizeInModule | |
LibraryError LibErrors | Collected errors when processing the |
LibTooFarDown TopLevelModuleName AgdaLibFile | The |
SolvedButOpenHoles | Some interaction points (holes) have not been filled by user.
There are not |
CyclicModuleDependency (List2 TopLevelModuleName) | The cycle starts and ends with the same module. |
FileNotFound TopLevelModuleName [AbsolutePath] | The list can be empty. |
OverlappingProjects AbsolutePath TopLevelModuleName TopLevelModuleName | |
AmbiguousTopLevelModuleName TopLevelModuleName (List2 AbsolutePath) | The given module has at least 2 file locations. |
ModuleNameUnexpected TopLevelModuleName TopLevelModuleName | Found module name, expected module name. |
ModuleNameDoesntMatchFileName TopLevelModuleName [AbsolutePath] | The list can be empty. |
ModuleDefinedInOtherFile TopLevelModuleName AbsolutePath AbsolutePath | Module name, file from which it was loaded, file which the include path says contains the module. |
InvalidFileName AbsolutePath InvalidFileNameReason | The file name does not correspond to a module name. |
ModuleNameHashCollision RawTopLevelModuleName (Maybe RawTopLevelModuleName) | |
BothWithAndRHS | |
AbstractConstructorNotInScope QName | |
CopatternHeadNotProjection QName | |
NotAllowedInDotPatterns NotAllowedInDotPatterns | |
NotInScope QName | |
NoSuchModule QName | |
AmbiguousName QName AmbiguousNameReason | |
AmbiguousModule QName (List1 ModuleName) | |
AmbiguousField Name (List2 ModuleName) | |
AmbiguousConstructor QName (List2 QName) | The list contains all interpretations of the name. |
ClashingDefinition QName QName (Maybe NiceDeclaration) | |
ClashingModule ModuleName ModuleName | |
DefinitionInDifferentModule QName | The given data/record definition rests in a different module than its signature. |
DuplicateImports QName (List1 ImportedName) | |
InvalidPattern Pattern | |
InvalidPun ConstructorOrPatternSynonym QName | Expected the identifier to be a variable, not a constructor or pattern synonym. |
RepeatedNamesInImportDirective (List1 (List2 ImportedName)) | Some names are bound several times by an |
RepeatedVariablesInPattern (List1 Name) | |
GeneralizeNotSupportedHere QName | |
GeneralizedVarInLetOpenedModule QName | |
MultipleFixityDecls (List1 (Name, Pair Fixity')) | |
MultiplePolarityPragmas (List1 Name) | |
ExplicitPolarityVsPragma QName | |
ConstructorNameOfNonRecord ResolvedName | |
CannotQuote CannotQuote | |
CannotQuoteTerm CannotQuoteTerm | |
DeclarationsAfterTopLevelModule | |
IllegalDeclarationBeforeTopLevelModule | |
MissingTypeSignature MissingTypeSignatureInfo | |
NotAnExpression Expr | |
NotAValidLetBinding (Maybe NotAValidLetBinding) | |
NotAValidLetExpression NotAValidLetExpression | |
NotValidBeforeField NiceDeclaration | |
OpenEverythingInRecordWhere | |
PrivateRecordField | |
QualifiedLocalModule | |
AsPatternInPatternSynonym | |
DotPatternInPatternSynonym | |
BadArgumentsToPatternSynonym AmbiguousQName | |
TooFewArgumentsToPatternSynonym AmbiguousQName | |
CannotResolveAmbiguousPatternSynonym (List1 (QName, PatternSynDefn)) | |
IllegalInstanceVariableInPatternSynonym Name | This variable is bound in the lhs of the pattern synonym in instance position, but not on the rhs. This is forbidden because expansion of pattern synonyms would not be faithful to availability of instances in instance search. |
PatternSynonymArgumentShadows ConstructorOrPatternSynonym Name (List1 AbstractName) | A variable to be bound in the pattern synonym resolved on the rhs as name of a constructor or a pattern synonym. The resolvents are given in the list. |
UnusedVariableInPatternSynonym Name | This variable is only bound on the lhs of the pattern synonym, not on the rhs. |
UnboundVariablesInPatternSynonym (List1 Name) | These variables are only bound on the rhs of the pattern synonym, not on the lhs. Operator errors |
NoParseForApplication (List2 Expr) | |
AmbiguousParseForApplication (List2 Expr) (List1 Expr) | |
NoParseForLHS LHSOrPatSyn [Pattern] Pattern | The list contains patterns that failed to be interpreted. If it is non-empty, the first entry could be printed as error hint. |
AmbiguousParseForLHS LHSOrPatSyn Pattern (List2 Pattern) | Pattern and its possible interpretations. |
AmbiguousProjection QName (List1 QName) | The list contains alternative interpretations of the name. |
AmbiguousOverloadedProjection (List1 QName) Doc | |
OperatorInformation [NotationSection] TypeError | The list of notations can be empty. |
InstanceNoCandidate Type [(Term, TCErr)] | The list can be empty. Reflection errors |
ExecError ExecError | |
UnquoteFailed UnquoteError | |
DeBruijnIndexOutOfScope Nat Telescope [Name] | The list can be empty. Language option errors |
NeedOptionAllowExec | |
NeedOptionCopatterns | |
NeedOptionCubical Cubical String | Flavor of cubical needed for the given reason. |
NeedOptionPatternMatching | |
NeedOptionProp | |
NeedOptionRewriting | |
NeedOptionSizedTypes String | Need |
NeedOptionTwoLevel | |
NeedOptionUniversePolymorphism | |
NonFatalErrors (Set1 TCWarning) | |
InstanceSearchDepthExhausted Term Type Int | |
TriedToCopyConstrainedPrim QName | |
InvalidInstanceHeadType Type WhyInvalidInstanceType | |
InteractionError InteractionError | |
BackendDoesNotSupportOnlyScopeChecking BackendName | The given backend does not support |
CubicalCompilationNotSupported Cubical | NYI: Compilation of files using the given flavor of |
CustomBackendError BackendName Doc | Used for backend-specific errors. The string is the backend name. |
GHCBackendError GHCBackendError | Errors raised by the GHC backend. |
JSBackendError JSBackendError | Errors raised by the JS backend. |
UnknownBackend BackendName (Set BackendName) | Unknown backend requested, known ones are in the |
Instances
PrettyTCM TypeError Source # | |||||
Defined in Agda.TypeChecking.Errors | |||||
NFData TypeError Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Generic TypeError Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
| |||||
Show TypeError Source # | |||||
type Rep TypeError Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep TypeError = D1 ('MetaData "TypeError" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (((((((C1 ('MetaCons "InternalError" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String)) :+: (C1 ('MetaCons "NotImplemented" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String)) :+: C1 ('MetaCons "NotSupported" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String)))) :+: ((C1 ('MetaCons "CompilationError" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String)) :+: C1 ('MetaCons "SyntaxError" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String))) :+: (C1 ('MetaCons "OptionError" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 OptionError)) :+: C1 ('MetaCons "NicifierError" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 DeclarationException'))))) :+: (((C1 ('MetaCons "DoNotationError" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String)) :+: C1 ('MetaCons "IdiomBracketError" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String))) :+: (C1 ('MetaCons "NoKnownRecordWithSuchFields" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Name])) :+: C1 ('MetaCons "ShouldEndInApplicationOfTheDatatype" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)))) :+: ((C1 ('MetaCons "ConstructorPatternInWrongDatatype" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "CantResolveOverloadedConstructorsTargetingSameDatatype" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 QName)))) :+: (C1 ('MetaCons "ConstructorDoesNotTargetGivenType" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)) :+: C1 ('MetaCons "InvalidDottedExpression" 'PrefixI 'False) (U1 :: Type -> Type))))) :+: (((C1 ('MetaCons "LiteralTooBig" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "NegativeLiteralInPattern" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "WrongHidingInLHS" 'PrefixI 'False) (U1 :: Type -> Type))) :+: ((C1 ('MetaCons "WrongHidingInLambda" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)) :+: C1 ('MetaCons "WrongHidingInApplication" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type))) :+: (C1 ('MetaCons "WrongHidingInProjection" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "IllegalHidingInPostfixProjection" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (NamedArg Expr)))))) :+: (((C1 ('MetaCons "WrongNamedArgument" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (NamedArg Expr)) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 NamedName))) :+: C1 ('MetaCons "WrongAnnotationInLambda" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "WrongIrrelevanceInLambda" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "WrongQuantityInLambda" 'PrefixI 'False) (U1 :: Type -> Type))) :+: ((C1 ('MetaCons "WrongCohesionInLambda" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "WrongPolarityInLambda" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "QuantityMismatch" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Quantity) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Quantity)) :+: C1 ('MetaCons "HidingMismatch" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Hiding) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Hiding))))))) :+: ((((C1 ('MetaCons "RelevanceMismatch" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Relevance) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Relevance)) :+: (C1 ('MetaCons "ForcedConstructorNotInstantiated" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Pattern)) :+: C1 ('MetaCons "IllformedProjectionPatternAbstract" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Pattern)))) :+: ((C1 ('MetaCons "IllformedProjectionPatternConcrete" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Pattern)) :+: C1 ('MetaCons "CannotEliminateWithPattern" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe Blocker)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (NamedArg Pattern)) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)))) :+: (C1 ('MetaCons "CannotEliminateWithProjection" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Arg Type)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName))) :+: C1 ('MetaCons "WrongNumberOfConstructorArguments" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Nat) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Nat)))))) :+: (((C1 ('MetaCons "ShouldBeEmpty" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [DeBruijnPattern])) :+: C1 ('MetaCons "ShouldBeASort" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type))) :+: (C1 ('MetaCons "ShouldBePi" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)) :+: C1 ('MetaCons "ShouldBePath" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)))) :+: ((C1 ('MetaCons "ShouldBeRecordType" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)) :+: C1 ('MetaCons "ShouldBeRecordPattern" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 DeBruijnPattern))) :+: (C1 ('MetaCons "CannotApply" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)) :+: C1 ('MetaCons "InvalidTypeSort" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Sort)))))) :+: (((C1 ('MetaCons "SplitOnCoinductive" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "SplitOnIrrelevant" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Dom Type))) :+: C1 ('MetaCons "SplitOnUnusableCohesion" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Dom Type))))) :+: ((C1 ('MetaCons "SplitOnUnusablePolarity" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Dom Type))) :+: C1 ('MetaCons "SplitOnNonVariable" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type))) :+: (C1 ('MetaCons "SplitOnNonEtaRecord" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "SplitOnAbstract" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName))))) :+: (((C1 ('MetaCons "SplitOnUnchecked" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "SplitOnPartial" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Dom Type)))) :+: (C1 ('MetaCons "SplitInProp" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 DataOrRecordE)) :+: C1 ('MetaCons "DefinitionIsIrrelevant" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)))) :+: ((C1 ('MetaCons "DefinitionIsErased" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "ProjectionIsIrrelevant" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName))) :+: (C1 ('MetaCons "VariableIsIrrelevant" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Name)) :+: C1 ('MetaCons "VariableIsErased" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Name)))))))) :+: (((((C1 ('MetaCons "VariableIsOfUnusableCohesion" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Name) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Cohesion)) :+: (C1 ('MetaCons "LambdaIsErased" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "RecordIsErased" 'PrefixI 'False) (U1 :: Type -> Type))) :+: ((C1 ('MetaCons "InvalidModalTelescopeUse" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Modality)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Modality) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Definition))) :+: C1 ('MetaCons "VariableIsOfUnusablePolarity" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Name) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 PolarityModality))) :+: (C1 ('MetaCons "UnequalLevel" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Level) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Level))) :+: C1 ('MetaCons "UnequalTerms" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 CompareAs)))))) :+: (((C1 ('MetaCons "UnequalRelevance" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term))) :+: C1 ('MetaCons "UnequalQuantity" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)))) :+: (C1 ('MetaCons "UnequalCohesion" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term))) :+: C1 ('MetaCons "UnequalPolarity" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term))))) :+: ((C1 ('MetaCons "UnequalFiniteness" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term))) :+: C1 ('MetaCons "UnequalHiding" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term))) :+: (C1 ('MetaCons "UnequalSorts" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Sort) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Sort)) :+: C1 ('MetaCons "NotLeqSort" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Sort) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Sort)))))) :+: (((C1 ('MetaCons "MetaCannotDependOn" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 MetaId) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Nat))) :+: (C1 ('MetaCons "MetaIrrelevantSolution" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 MetaId) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)) :+: C1 ('MetaCons "MetaErasedSolution" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 MetaId) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)))) :+: ((C1 ('MetaCons "GenericError" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String)) :+: C1 ('MetaCons "GenericDocError" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Doc))) :+: (C1 ('MetaCons "SortOfSplitVarError" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe Blocker)) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Doc)) :+: C1 ('MetaCons "WrongSharpArity" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName))))) :+: (((C1 ('MetaCons "BuiltinMustBeConstructor" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 BuiltinId) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr)) :+: C1 ('MetaCons "BuiltinMustBeData" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 BuiltinId) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Int))) :+: (C1 ('MetaCons "BuiltinMustBeDef" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 BuiltinId)) :+: C1 ('MetaCons "BuiltinMustBeFunction" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 BuiltinId)))) :+: ((C1 ('MetaCons "BuiltinMustBePostulate" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 BuiltinId)) :+: C1 ('MetaCons "NoSuchBuiltinName" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String))) :+: (C1 ('MetaCons "InvalidBuiltin" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String)) :+: C1 ('MetaCons "DuplicateBuiltinBinding" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 BuiltinId) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)))))))) :+: ((((C1 ('MetaCons "NoBindingForBuiltin" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 BuiltinId)) :+: (C1 ('MetaCons "NoBindingForPrimitive" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 PrimitiveId)) :+: C1 ('MetaCons "NoSuchPrimitiveFunction" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String)))) :+: ((C1 ('MetaCons "DuplicatePrimitiveBinding" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 PrimitiveId) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName))) :+: C1 ('MetaCons "WrongArgInfoForPrimitive" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 PrimitiveId) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ArgInfo) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ArgInfo)))) :+: (C1 ('MetaCons "ShadowedModule" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Name) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 ModuleName))) :+: C1 ('MetaCons "BuiltinInParameterisedModule" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 BuiltinId))))) :+: (((C1 ('MetaCons "IllegalDeclarationInDataDefinition" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 Declaration))) :+: C1 ('MetaCons "IllegalLetInTelescope" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 TypedBinding))) :+: (C1 ('MetaCons "IllegalPatternInTelescope" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Binder)) :+: C1 ('MetaCons "AbsentRHSRequiresAbsurdPattern" 'PrefixI 'False) (U1 :: Type -> Type))) :+: ((C1 ('MetaCons "TooManyFields" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Name]) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 Name)))) :+: C1 ('MetaCons "DuplicateFields" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 Name)))) :+: (C1 ('MetaCons "DuplicateConstructors" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 Name))) :+: C1 ('MetaCons "DuplicateOverlapPragma" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 OverlapMode) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 OverlapMode))))))) :+: (((C1 ('MetaCons "WithOnFreeVariable" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)) :+: (C1 ('MetaCons "UnexpectedWithPatterns" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 Pattern))) :+: C1 ('MetaCons "WithClausePatternMismatch" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Pattern) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (NamedArg DeBruijnPattern))))) :+: ((C1 ('MetaCons "IllTypedPatternAfterWithAbstraction" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Pattern)) :+: C1 ('MetaCons "TooFewPatternsInWithClause" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "TooManyPatternsInWithClause" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "PathAbstractionFailed" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Abs Type)))))) :+: (((C1 ('MetaCons "FieldOutsideRecord" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "ModuleArityMismatch" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ModuleName) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Telescope) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Either (List1 (NamedArg Expr)) Args))))) :+: (C1 ('MetaCons "GeneralizeCyclicDependency" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "ReferencesFutureVariables" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 Int))) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Arg Term)) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Int))))) :+: ((C1 ('MetaCons "DoesNotMentionTicks" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Arg Term)))) :+: C1 ('MetaCons "MismatchedProjectionsError" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName))) :+: (C1 ('MetaCons "AttributeKindNotEnabled" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String))) :+: C1 ('MetaCons "InvalidProjectionParameter" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (NamedArg Expr)))))))))) :+: ((((((C1 ('MetaCons "TacticAttributeNotAllowed" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "CannotRewriteByNonEquation" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)) :+: C1 ('MetaCons "MacroResultTypeMismatch" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)))) :+: ((C1 ('MetaCons "NamedWhereModuleInRefinedContext" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Term]) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [String])) :+: C1 ('MetaCons "ComatchingDisabledForRecord" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName))) :+: (C1 ('MetaCons "IncorrectTypeForRewriteRelation" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 IncorrectTypeForRewriteRelationReason)) :+: C1 ('MetaCons "CannotGenerateHCompClause" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type))))) :+: (((C1 ('MetaCons "CannotGenerateTransportClause" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Closure (Abs Type)))) :+: C1 ('MetaCons "CubicalNotErasure" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName))) :+: (C1 ('MetaCons "CubicalPrimitiveNotFullyApplied" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "ExpectedIntervalLiteral" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr)))) :+: ((C1 ('MetaCons "FaceConstraintDisjunction" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "FaceConstraintUnsatisfiable" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "PatternInPathLambda" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "PatternInSystem" 'PrefixI 'False) (U1 :: Type -> Type))))) :+: (((C1 ('MetaCons "UnexpectedParameter" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 LamBinding)) :+: (C1 ('MetaCons "NoParameterOfName" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ArgName)) :+: C1 ('MetaCons "UnexpectedModalityAnnotationInParameter" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 LamBinding)))) :+: ((C1 ('MetaCons "ExpectedBindingForParameter" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Dom Type)) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Abs Type))) :+: C1 ('MetaCons "UnexpectedTypeSignatureForParameter" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 (NamedArg Binder))))) :+: (C1 ('MetaCons "SortDoesNotAdmitDataDefinitions" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Sort)) :+: C1 ('MetaCons "SortCannotDependOnItsIndex" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type))))) :+: (((C1 ('MetaCons "UnusableAtModality" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 WhyCheckModality) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Modality) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term))) :+: C1 ('MetaCons "SplitError" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 SplitError))) :+: (C1 ('MetaCons "ImpossibleConstructor" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 NegativeUnification)) :+: C1 ('MetaCons "DatatypeIndexPolarity" 'PrefixI 'False) (U1 :: Type -> Type))) :+: ((C1 ('MetaCons "RecursiveRecordNeedsInductivity" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "CannotSolveSizeConstraints" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 (ProblemConstraint, HypSizeConstraint))) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Doc))) :+: (C1 ('MetaCons "ContradictorySizeConstraint" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (ProblemConstraint, HypSizeConstraint))) :+: C1 ('MetaCons "EmptyTypeOfSizes" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term))))))) :+: ((((C1 ('MetaCons "FunctionTypeInSizeUniv" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)) :+: (C1 ('MetaCons "PostulatedSizeInModule" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "LibraryError" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 LibErrors)))) :+: ((C1 ('MetaCons "LibTooFarDown" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 TopLevelModuleName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 AgdaLibFile)) :+: C1 ('MetaCons "SolvedButOpenHoles" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "CyclicModuleDependency" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List2 TopLevelModuleName))) :+: C1 ('MetaCons "FileNotFound" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 TopLevelModuleName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [AbsolutePath]))))) :+: (((C1 ('MetaCons "OverlappingProjects" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 AbsolutePath) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 TopLevelModuleName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 TopLevelModuleName))) :+: C1 ('MetaCons "AmbiguousTopLevelModuleName" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 TopLevelModuleName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List2 AbsolutePath)))) :+: (C1 ('MetaCons "ModuleNameUnexpected" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 TopLevelModuleName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 TopLevelModuleName)) :+: C1 ('MetaCons "ModuleNameDoesntMatchFileName" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 TopLevelModuleName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [AbsolutePath])))) :+: ((C1 ('MetaCons "ModuleDefinedInOtherFile" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 TopLevelModuleName) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 AbsolutePath) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 AbsolutePath))) :+: C1 ('MetaCons "InvalidFileName" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 AbsolutePath) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 InvalidFileNameReason))) :+: (C1 ('MetaCons "ModuleNameHashCollision" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 RawTopLevelModuleName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe RawTopLevelModuleName))) :+: C1 ('MetaCons "BothWithAndRHS" 'PrefixI 'False) (U1 :: Type -> Type))))) :+: (((C1 ('MetaCons "AbstractConstructorNotInScope" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: (C1 ('MetaCons "CopatternHeadNotProjection" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "NotAllowedInDotPatterns" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 NotAllowedInDotPatterns)))) :+: ((C1 ('MetaCons "NotInScope" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "NoSuchModule" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName))) :+: (C1 ('MetaCons "AmbiguousName" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 AmbiguousNameReason)) :+: C1 ('MetaCons "AmbiguousModule" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 ModuleName)))))) :+: (((C1 ('MetaCons "AmbiguousField" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Name) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List2 ModuleName))) :+: C1 ('MetaCons "AmbiguousConstructor" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List2 QName)))) :+: (C1 ('MetaCons "ClashingDefinition" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe NiceDeclaration)))) :+: C1 ('MetaCons "ClashingModule" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ModuleName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ModuleName)))) :+: ((C1 ('MetaCons "DefinitionInDifferentModule" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "DuplicateImports" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 ImportedName)))) :+: (C1 ('MetaCons "InvalidPattern" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Pattern)) :+: C1 ('MetaCons "InvalidPun" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ConstructorOrPatternSynonym) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)))))))) :+: (((((C1 ('MetaCons "RepeatedNamesInImportDirective" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 (List2 ImportedName)))) :+: (C1 ('MetaCons "RepeatedVariablesInPattern" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 Name))) :+: C1 ('MetaCons "GeneralizeNotSupportedHere" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)))) :+: ((C1 ('MetaCons "GeneralizedVarInLetOpenedModule" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "MultipleFixityDecls" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 (Name, Pair Fixity'))))) :+: (C1 ('MetaCons "MultiplePolarityPragmas" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 Name))) :+: C1 ('MetaCons "ExplicitPolarityVsPragma" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName))))) :+: (((C1 ('MetaCons "ConstructorNameOfNonRecord" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ResolvedName)) :+: C1 ('MetaCons "CannotQuote" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 CannotQuote))) :+: (C1 ('MetaCons "CannotQuoteTerm" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 CannotQuoteTerm)) :+: C1 ('MetaCons "DeclarationsAfterTopLevelModule" 'PrefixI 'False) (U1 :: Type -> Type))) :+: ((C1 ('MetaCons "IllegalDeclarationBeforeTopLevelModule" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "MissingTypeSignature" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 MissingTypeSignatureInfo))) :+: (C1 ('MetaCons "NotAnExpression" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr)) :+: C1 ('MetaCons "NotAValidLetBinding" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe NotAValidLetBinding))))))) :+: (((C1 ('MetaCons "NotAValidLetExpression" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 NotAValidLetExpression)) :+: (C1 ('MetaCons "NotValidBeforeField" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 NiceDeclaration)) :+: C1 ('MetaCons "OpenEverythingInRecordWhere" 'PrefixI 'False) (U1 :: Type -> Type))) :+: ((C1 ('MetaCons "PrivateRecordField" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "QualifiedLocalModule" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "AsPatternInPatternSynonym" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "DotPatternInPatternSynonym" 'PrefixI 'False) (U1 :: Type -> Type)))) :+: (((C1 ('MetaCons "BadArgumentsToPatternSynonym" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 AmbiguousQName)) :+: C1 ('MetaCons "TooFewArgumentsToPatternSynonym" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 AmbiguousQName))) :+: (C1 ('MetaCons "CannotResolveAmbiguousPatternSynonym" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 (QName, PatternSynDefn)))) :+: C1 ('MetaCons "IllegalInstanceVariableInPatternSynonym" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Name)))) :+: ((C1 ('MetaCons "PatternSynonymArgumentShadows" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ConstructorOrPatternSynonym) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Name) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 AbstractName)))) :+: C1 ('MetaCons "UnusedVariableInPatternSynonym" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Name))) :+: (C1 ('MetaCons "UnboundVariablesInPatternSynonym" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 Name))) :+: C1 ('MetaCons "NoParseForApplication" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List2 Expr)))))))) :+: ((((C1 ('MetaCons "AmbiguousParseForApplication" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List2 Expr)) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 Expr))) :+: (C1 ('MetaCons "NoParseForLHS" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 LHSOrPatSyn) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Pattern]) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Pattern))) :+: C1 ('MetaCons "AmbiguousParseForLHS" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 LHSOrPatSyn) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Pattern) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List2 Pattern)))))) :+: ((C1 ('MetaCons "AmbiguousProjection" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 QName))) :+: C1 ('MetaCons "AmbiguousOverloadedProjection" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 QName)) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Doc))) :+: (C1 ('MetaCons "OperatorInformation" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [NotationSection]) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 TypeError)) :+: C1 ('MetaCons "InstanceNoCandidate" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [(Term, TCErr)]))))) :+: (((C1 ('MetaCons "ExecError" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ExecError)) :+: C1 ('MetaCons "UnquoteFailed" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 UnquoteError))) :+: (C1 ('MetaCons "DeBruijnIndexOutOfScope" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Nat) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Telescope) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Name]))) :+: C1 ('MetaCons "NeedOptionAllowExec" 'PrefixI 'False) (U1 :: Type -> Type))) :+: ((C1 ('MetaCons "NeedOptionCopatterns" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "NeedOptionCubical" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Cubical) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String))) :+: (C1 ('MetaCons "NeedOptionPatternMatching" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "NeedOptionProp" 'PrefixI 'False) (U1 :: Type -> Type))))) :+: (((C1 ('MetaCons "NeedOptionRewriting" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "NeedOptionSizedTypes" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String)) :+: C1 ('MetaCons "NeedOptionTwoLevel" 'PrefixI 'False) (U1 :: Type -> Type))) :+: ((C1 ('MetaCons "NeedOptionUniversePolymorphism" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "NonFatalErrors" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Set1 TCWarning)))) :+: (C1 ('MetaCons "InstanceSearchDepthExhausted" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Int))) :+: C1 ('MetaCons "TriedToCopyConstrainedPrim" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName))))) :+: (((C1 ('MetaCons "InvalidInstanceHeadType" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 WhyInvalidInstanceType)) :+: C1 ('MetaCons "InteractionError" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 InteractionError))) :+: (C1 ('MetaCons "BackendDoesNotSupportOnlyScopeChecking" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 BackendName)) :+: C1 ('MetaCons "CubicalCompilationNotSupported" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Cubical)))) :+: ((C1 ('MetaCons "CustomBackendError" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 BackendName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Doc)) :+: C1 ('MetaCons "GHCBackendError" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 GHCBackendError))) :+: (C1 ('MetaCons "JSBackendError" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 JSBackendError)) :+: C1 ('MetaCons "UnknownBackend" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 BackendName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Set BackendName))))))))))) |
data UnificationFailure Source #
Constructors
UnifyIndicesNotVars Telescope Type Term Term Args | Failed to apply injectivity to constructor of indexed datatype |
UnifyRecursiveEq Telescope Type Int Term | Can't solve equation because variable occurs in (type of) lhs |
UnifyReflexiveEq Telescope Type Term | Can't solve reflexive equation because --without-K is enabled |
UnifyUnusableModality Telescope Type Int Term Modality | Can't solve equation because solution modality is less "usable" |
Instances
PrettyTCM UnificationFailure Source # | |||||
Defined in Agda.TypeChecking.Errors Methods prettyTCM :: MonadPretty m => UnificationFailure -> m Doc Source # | |||||
NFData UnificationFailure Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: UnificationFailure -> () # | |||||
Generic UnificationFailure Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
Methods from :: UnificationFailure -> Rep UnificationFailure x # to :: Rep UnificationFailure x -> UnificationFailure # | |||||
Show UnificationFailure Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods showsPrec :: Int -> UnificationFailure -> ShowS # show :: UnificationFailure -> String # showList :: [UnificationFailure] -> ShowS # | |||||
type Rep UnificationFailure Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep UnificationFailure = D1 ('MetaData "UnificationFailure" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) ((C1 ('MetaCons "UnifyIndicesNotVars" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Telescope) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Args)))) :+: C1 ('MetaCons "UnifyRecursiveEq" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Telescope) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Int) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)))) :+: (C1 ('MetaCons "UnifyReflexiveEq" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Telescope) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term))) :+: C1 ('MetaCons "UnifyUnusableModality" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Telescope) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Int) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Modality)))))) |
data UnquoteError Source #
Constructors
BlockedOnMeta TCState Blocker | |
CannotDeclareHiddenFunction QName | |
CommitAfterDef | |
ConInsteadOfDef QName String String | |
DefineDataNotData QName | |
DefInsteadOfCon QName String String | |
MissingDeclaration QName | |
MissingDefinition QName | |
NakedUnquote | |
NonCanonical String Term | |
PatLamWithoutClauses Term | |
StaleMeta TopLevelModuleName MetaId | Attempt to unquote a serialized meta. |
TooManyParameters Nat Expr | Attempt to shave of |
UnboundName QName |
Instances
PrettyTCM UnquoteError Source # | |||||
Defined in Agda.TypeChecking.Errors Methods prettyTCM :: MonadPretty m => UnquoteError -> m Doc Source # | |||||
NFData UnquoteError Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: UnquoteError -> () # | |||||
Generic UnquoteError Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
| |||||
Show UnquoteError Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods showsPrec :: Int -> UnquoteError -> ShowS # show :: UnquoteError -> String # showList :: [UnquoteError] -> ShowS # | |||||
type Rep UnquoteError Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep UnquoteError = D1 ('MetaData "UnquoteError" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (((C1 ('MetaCons "BlockedOnMeta" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 TCState) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Blocker)) :+: (C1 ('MetaCons "CannotDeclareHiddenFunction" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "CommitAfterDef" 'PrefixI 'False) (U1 :: Type -> Type))) :+: ((C1 ('MetaCons "ConInsteadOfDef" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String))) :+: C1 ('MetaCons "DefineDataNotData" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName))) :+: (C1 ('MetaCons "DefInsteadOfCon" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String))) :+: C1 ('MetaCons "MissingDeclaration" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName))))) :+: ((C1 ('MetaCons "MissingDefinition" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: (C1 ('MetaCons "NakedUnquote" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "NonCanonical" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)))) :+: ((C1 ('MetaCons "PatLamWithoutClauses" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)) :+: C1 ('MetaCons "StaleMeta" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 TopLevelModuleName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 MetaId))) :+: (C1 ('MetaCons "TooManyParameters" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Nat) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr)) :+: C1 ('MetaCons "UnboundName" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)))))) |
data UnquoteFlags Source #
Constructors
UnquoteFlags | |
Fields |
Instances
NFData UnquoteFlags Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: UnquoteFlags -> () # | |||||
Generic UnquoteFlags Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
| |||||
type Rep UnquoteFlags Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep UnquoteFlags = D1 ('MetaData "UnquoteFlags" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "UnquoteFlags" 'PrefixI 'True) (S1 ('MetaSel ('Just "_unquoteNormalise") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool))) |
A non-fatal error is an error which does not prevent us from checking the document further and interacting with the user.
Constructors
NicifierIssue DeclarationWarning | |
TerminationIssue (List1 TerminationError) | |
UnreachableClauses QName (List1 Range) |
|
CoverageIssue QName (List1 (Telescope, [NamedArg DeBruijnPattern])) | `CoverageIssue f pss` means that |
CoverageNoExactSplit QName (List1 Clause) | |
InlineNoExactSplit QName Clause |
|
NotStrictlyPositive QName (Seq OccursWhere) | |
ConstructorDoesNotFitInData DataOrRecord_ QName Sort Sort TCErr | Checking whether constructor |
CoinductiveEtaRecord QName | A record type declared as both |
UnsolvedMetaVariables (Set1 Range) | Do not use directly with |
UnsolvedInteractionMetas (Set1 Range) | Do not use directly with |
UnsolvedConstraints (List1 ProblemConstraint) | Do not use directly with |
InteractionMetaBoundaries (Set1 Range) | Do not use directly with |
CantGeneralizeOverSorts (Set1 MetaId) | |
AbsurdPatternRequiresAbsentRHS | |
OldBuiltin BuiltinId BuiltinId | In `OldBuiltin old new`, the BUILTIN old has been replaced by new. |
BuiltinDeclaresIdentifier BuiltinId | The builtin declares a new identifier, so it should not be in scope. |
DuplicateRecordDirective RecordDirective | The given record directive is conflicting with a prior one in the same record declaration. |
EmptyRewritePragma | If the user wrote just |
EmptyWhere | An empty |
FixingRelevance String Relevance Relevance | Auto-correcting relevance pertaining to |
FixingCohesion String Cohesion Cohesion | Auto-correcting cohesion pertaining to |
FixingPolarity String PolarityModality PolarityModality | Auto-correcting polarity pertaining to |
IllformedAsClause String | If the user wrote something other than an unqualified name
in the |
InvalidCharacterLiteral Char | A character literal Agda does not support, e.g. surrogate code points. |
ClashesViaRenaming NameOrModule (Set1 Name) | If a |
UselessPatternDeclarationForRecord String | The 'pattern' declaration is useless in the presence
of either |
UselessPragma Range Doc | Warning when pragma is useless and thus ignored.
|
UselessPublic UselessPublicReason | If the user opens a module public before the module header. (See issue #2377.) |
UselessHiding (List1 ImportedName) | Names in |
UselessInline QName | |
UselessTactic | A tactic attribute applied to a non-hidden (visible or instance) argument. |
WrongInstanceDeclaration | |
InstanceWithExplicitArg QName | An instance was declared with an implicit argument, which means it will never actually be considered by instance search. |
InstanceNoOutputTypeName Doc | The type of an instance argument doesn't end in a named or variable type, so it will never be considered by instance search. |
InstanceArgWithExplicitArg Doc | As InstanceWithExplicitArg, but for local bindings rather than top-level instances. |
InversionDepthReached QName | The --inversion-max-depth was reached. |
SafeFlagPostulate QName | |
SafeFlagPragma (Set String) | Unsafe OPTIONS. |
SafeFlagWithoutKFlagPrimEraseEquality | |
WithoutKFlagPrimEraseEquality | |
ConflictingPragmaOptions String String | `ConflictingPragmaOptions a b`: Inconsistent options `--a` and `--no-b`, since `--a` implies `--b`. Ignoring `--no-b`. |
OptionWarning OptionWarning | |
ParseWarning ParseWarning | |
LibraryWarning LibWarning | |
DeprecationWarning String String String | `DeprecationWarning old new version`:
|
UserWarning Text | User-defined warning (e.g. to mention that a name is deprecated) |
DuplicateUsing (List1 ImportedName) | Duplicate mentions of the same name in |
FixityInRenamingModule (List1 Range) | Fixity of modules cannot be changed via renaming (since modules have no fixity). |
ModuleDoesntExport QName [Name] [Name] (List1 ImportedName) | Some imported names are not actually exported by the source module. The second argument is the names that could be exported. The third argument is the module names that could be exported. |
InfectiveImport Doc | Importing a file using an infective option into one which doesn't |
CoInfectiveImport Doc | Importing a file not using a coinfective option from one which does |
ConfluenceCheckingIncompleteBecauseOfMeta QName | Confluence checking incomplete because the definition of the |
ConfluenceForCubicalNotSupported | Confluence checking with |
NotARewriteRule QName IsAmbiguous |
|
IllegalRewriteRule QName IllegalRewriteRuleReason | |
RewriteNonConfluent Term Term Term Doc | Confluence checker found critical pair and equality checking resulted in a type error |
RewriteMaybeNonConfluent Term Term [Doc] | Confluence checker got stuck on computing overlap between two rewrite rules |
RewriteAmbiguousRules Term Term Term | The global confluence checker found a term |
RewriteMissingRule Term Term Term | The global confluence checker found a term |
PragmaCompileErased BackendName QName | COMPILE directive for an erased symbol. |
PragmaCompileList |
|
PragmaCompileMaybe |
|
PragmaCompileUnparsable String |
|
PragmaCompileWrong QName String | Wrong |
PragmaCompileWrongName QName IsAmbiguous |
|
PragmaExpectsDefinedSymbol String QName | |
PragmaExpectsUnambiguousConstructorOrFunction String QName IsAmbiguous | Pragma |
PragmaExpectsUnambiguousProjectionOrFunction String QName IsAmbiguous | Pragma |
NoMain TopLevelModuleName | Compiler run on module that does not have a |
NotInScopeW QName | Out of scope error we can recover from. |
UnsupportedIndexedMatch Doc | Was not able to compute a full equivalence when splitting. |
AsPatternShadowsConstructorOrPatternSynonym ConstructorOrPatternSynonym | The as-name in an as-pattern may not shadow a constructor or pattern synonym name, because this can be confusing to read. |
PatternShadowsConstructor Name QName | A pattern variable has the name of a constructor (data constructor or matchable record constructor). |
PlentyInHardCompileTimeMode QωOrigin | Explicit use of @ |
RecordFieldWarning RecordFieldWarning | |
MissingTypeSignatureForOpaque QName IsOpaque | An |
NotAffectedByOpaque | |
UnfoldingWrongName QName | Name in |
UnfoldTransparentName QName | |
UselessOpaque | |
HiddenNotInArgumentPosition Expr | |
InstanceNotInArgumentPosition Expr | |
MacroInLetBindings | |
AbstractInLetBindings | |
InvalidDisplayForm QName String | |
UnusedVariablesInDisplayForm (List1 Name) | The given names are bound in the lhs of the display form but not used on the rhs. This can indicate a user misunderstanding of display forms. |
TooManyArgumentsToSort QName (List1 (NamedArg Expr)) | Extra arguments to sort (will be ignored). |
WithClauseProjectionFixityMismatch | The with-clause uses projection in a different fixity style than the parent clause. |
TooManyPolarities QName PragmaPolarities | Too many polarities given in POLARITY pragma for |
TopLevelPolarity QName PolarityModality | Definition with non-default polarity annotation. |
FaceConstraintCannotBeHidden ArgInfo | Face constraint patterns |
FaceConstraintCannotBeNamed NamedName | Face constraint patterns |
CustomBackendWarning String Doc | Used for backend-specific warnings. The string is the backend name. |
Instances
PrettyTCM Warning Source # | |||||
Defined in Agda.TypeChecking.Pretty.Warning | |||||
EmbPrj Warning Source # | |||||
NFData Warning Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Generic Warning Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
| |||||
Show Warning Source # | |||||
type Rep Warning Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep Warning = D1 ('MetaData "Warning" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) ((((((C1 ('MetaCons "NicifierIssue" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 DeclarationWarning)) :+: C1 ('MetaCons "TerminationIssue" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 TerminationError)))) :+: (C1 ('MetaCons "UnreachableClauses" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 Range))) :+: (C1 ('MetaCons "CoverageIssue" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 (Telescope, [NamedArg DeBruijnPattern])))) :+: C1 ('MetaCons "CoverageNoExactSplit" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 Clause)))))) :+: ((C1 ('MetaCons "InlineNoExactSplit" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Clause)) :+: (C1 ('MetaCons "NotStrictlyPositive" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Seq OccursWhere))) :+: C1 ('MetaCons "ConstructorDoesNotFitInData" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 DataOrRecord_) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Sort) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Sort) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 TCErr)))))) :+: (C1 ('MetaCons "CoinductiveEtaRecord" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: (C1 ('MetaCons "UnsolvedMetaVariables" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Set1 Range))) :+: C1 ('MetaCons "UnsolvedInteractionMetas" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Set1 Range))))))) :+: (((C1 ('MetaCons "UnsolvedConstraints" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 ProblemConstraint))) :+: (C1 ('MetaCons "InteractionMetaBoundaries" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Set1 Range))) :+: C1 ('MetaCons "CantGeneralizeOverSorts" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Set1 MetaId))))) :+: (C1 ('MetaCons "AbsurdPatternRequiresAbsentRHS" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "OldBuiltin" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 BuiltinId) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 BuiltinId)) :+: C1 ('MetaCons "BuiltinDeclaresIdentifier" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 BuiltinId))))) :+: ((C1 ('MetaCons "DuplicateRecordDirective" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 RecordDirective)) :+: (C1 ('MetaCons "EmptyRewritePragma" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "EmptyWhere" 'PrefixI 'False) (U1 :: Type -> Type))) :+: (C1 ('MetaCons "FixingRelevance" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Relevance) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Relevance))) :+: (C1 ('MetaCons "FixingCohesion" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Cohesion) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Cohesion))) :+: C1 ('MetaCons "FixingPolarity" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 PolarityModality) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 PolarityModality)))))))) :+: ((((C1 ('MetaCons "IllformedAsClause" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String)) :+: (C1 ('MetaCons "InvalidCharacterLiteral" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Char)) :+: C1 ('MetaCons "ClashesViaRenaming" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 NameOrModule) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Set1 Name))))) :+: (C1 ('MetaCons "UselessPatternDeclarationForRecord" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String)) :+: (C1 ('MetaCons "UselessPragma" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Doc)) :+: C1 ('MetaCons "UselessPublic" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 UselessPublicReason))))) :+: ((C1 ('MetaCons "UselessHiding" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 ImportedName))) :+: (C1 ('MetaCons "UselessInline" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "UselessTactic" 'PrefixI 'False) (U1 :: Type -> Type))) :+: (C1 ('MetaCons "WrongInstanceDeclaration" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "InstanceWithExplicitArg" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "InstanceNoOutputTypeName" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Doc)))))) :+: (((C1 ('MetaCons "InstanceArgWithExplicitArg" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Doc)) :+: (C1 ('MetaCons "InversionDepthReached" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "SafeFlagPostulate" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)))) :+: (C1 ('MetaCons "SafeFlagPragma" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Set String))) :+: (C1 ('MetaCons "SafeFlagWithoutKFlagPrimEraseEquality" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "WithoutKFlagPrimEraseEquality" 'PrefixI 'False) (U1 :: Type -> Type)))) :+: ((C1 ('MetaCons "ConflictingPragmaOptions" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String)) :+: (C1 ('MetaCons "OptionWarning" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 OptionWarning)) :+: C1 ('MetaCons "ParseWarning" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ParseWarning)))) :+: (C1 ('MetaCons "LibraryWarning" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 LibWarning)) :+: (C1 ('MetaCons "DeprecationWarning" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String))) :+: C1 ('MetaCons "UserWarning" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Text)))))))) :+: (((((C1 ('MetaCons "DuplicateUsing" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 ImportedName))) :+: C1 ('MetaCons "FixityInRenamingModule" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 Range)))) :+: (C1 ('MetaCons "ModuleDoesntExport" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Name])) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Name]) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 ImportedName)))) :+: (C1 ('MetaCons "InfectiveImport" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Doc)) :+: C1 ('MetaCons "CoInfectiveImport" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Doc))))) :+: ((C1 ('MetaCons "ConfluenceCheckingIncompleteBecauseOfMeta" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: (C1 ('MetaCons "ConfluenceForCubicalNotSupported" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "NotARewriteRule" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 IsAmbiguous)))) :+: (C1 ('MetaCons "IllegalRewriteRule" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 IllegalRewriteRuleReason)) :+: (C1 ('MetaCons "RewriteNonConfluent" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Doc))) :+: C1 ('MetaCons "RewriteMaybeNonConfluent" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Doc]))))))) :+: (((C1 ('MetaCons "RewriteAmbiguousRules" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term))) :+: (C1 ('MetaCons "RewriteMissingRule" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term))) :+: C1 ('MetaCons "PragmaCompileErased" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 BackendName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)))) :+: (C1 ('MetaCons "PragmaCompileList" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "PragmaCompileMaybe" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "PragmaCompileUnparsable" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String))))) :+: ((C1 ('MetaCons "PragmaCompileWrong" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String)) :+: (C1 ('MetaCons "PragmaCompileWrongName" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 IsAmbiguous)) :+: C1 ('MetaCons "PragmaExpectsDefinedSymbol" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)))) :+: (C1 ('MetaCons "PragmaExpectsUnambiguousConstructorOrFunction" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 IsAmbiguous))) :+: (C1 ('MetaCons "PragmaExpectsUnambiguousProjectionOrFunction" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 IsAmbiguous))) :+: C1 ('MetaCons "NoMain" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 TopLevelModuleName))))))) :+: ((((C1 ('MetaCons "NotInScopeW" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: (C1 ('MetaCons "UnsupportedIndexedMatch" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Doc)) :+: C1 ('MetaCons "AsPatternShadowsConstructorOrPatternSynonym" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ConstructorOrPatternSynonym)))) :+: (C1 ('MetaCons "PatternShadowsConstructor" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Name) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: (C1 ('MetaCons "PlentyInHardCompileTimeMode" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QωOrigin)) :+: C1 ('MetaCons "RecordFieldWarning" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 RecordFieldWarning))))) :+: ((C1 ('MetaCons "MissingTypeSignatureForOpaque" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 IsOpaque)) :+: (C1 ('MetaCons "NotAffectedByOpaque" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "UnfoldingWrongName" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)))) :+: (C1 ('MetaCons "UnfoldTransparentName" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: (C1 ('MetaCons "UselessOpaque" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "HiddenNotInArgumentPosition" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr)))))) :+: (((C1 ('MetaCons "InstanceNotInArgumentPosition" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr)) :+: (C1 ('MetaCons "MacroInLetBindings" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "AbstractInLetBindings" 'PrefixI 'False) (U1 :: Type -> Type))) :+: (C1 ('MetaCons "InvalidDisplayForm" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String)) :+: (C1 ('MetaCons "UnusedVariablesInDisplayForm" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 Name))) :+: C1 ('MetaCons "TooManyArgumentsToSort" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 (NamedArg Expr))))))) :+: ((C1 ('MetaCons "WithClauseProjectionFixityMismatch" 'PrefixI 'True) ((S1 ('MetaSel ('Just "withClausePattern") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (NamedArg Pattern)) :*: S1 ('MetaSel ('Just "withClauseProjectionOrigin") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ProjOrigin)) :*: (S1 ('MetaSel ('Just "parentPattern") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (NamedArg DeBruijnPattern)) :*: S1 ('MetaSel ('Just "parentProjectionOrigin") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ProjOrigin))) :+: (C1 ('MetaCons "TooManyPolarities" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 PragmaPolarities)) :+: C1 ('MetaCons "TopLevelPolarity" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 PolarityModality)))) :+: (C1 ('MetaCons "FaceConstraintCannotBeHidden" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ArgInfo)) :+: (C1 ('MetaCons "FaceConstraintCannotBeNamed" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 NamedName)) :+: C1 ('MetaCons "CustomBackendWarning" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Doc))))))))) |
data WarningsAndNonFatalErrors Source #
Assorted warnings and errors to be displayed to the user
Constructors
WarningsAndNonFatalErrors | |
Fields |
Instances
EncodeTCM DisplayInfo Source # | |
Defined in Agda.Interaction.JSONTop | |
EncodeTCM Response Source # | |
Null WarningsAndNonFatalErrors Source # | |
Defined in Agda.TypeChecking.Monad.Base Methods empty :: WarningsAndNonFatalErrors Source # null :: WarningsAndNonFatalErrors -> Bool Source # |
data WhyCheckModality Source #
Why are we performing a modality check?
Constructors
ConstructorType | Because --without-K is enabled, so the types of data constructors must be usable at the context's modality. |
IndexedClause | Because --without-K is enabled, so the result type of clauses must be usable at the context's modality. |
IndexedClauseArg Name Name | Because --without-K is enabled, so any argument (second name) which mentions a dotted argument (first name) must have a type which is usable at the context's modality. |
GeneratedClause | Because we double-check the --cubical-compatible clauses. This is an internal error! |
Instances
NFData WhyCheckModality Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: WhyCheckModality -> () # | |||||
Generic WhyCheckModality Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
Methods from :: WhyCheckModality -> Rep WhyCheckModality x # to :: Rep WhyCheckModality x -> WhyCheckModality # | |||||
Show WhyCheckModality Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods showsPrec :: Int -> WhyCheckModality -> ShowS # show :: WhyCheckModality -> String # showList :: [WhyCheckModality] -> ShowS # | |||||
type Rep WhyCheckModality Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep WhyCheckModality = D1 ('MetaData "WhyCheckModality" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) ((C1 ('MetaCons "ConstructorType" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "IndexedClause" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "IndexedClauseArg" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Name) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Name)) :+: C1 ('MetaCons "GeneratedClause" 'PrefixI 'False) (U1 :: Type -> Type))) |
data WhyInvalidInstanceType Source #
Reason for why the instance type is invalid.
Constructors
ImproperInstHead | The type isn't headed by a local, a definition, or a postulate (e.g. it's a universe) |
ImproperInstTele | The type we're looking for has a visible argument |
Instances
NFData WhyInvalidInstanceType Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: WhyInvalidInstanceType -> () # | |||||
Generic WhyInvalidInstanceType Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
Methods from :: WhyInvalidInstanceType -> Rep WhyInvalidInstanceType x # to :: Rep WhyInvalidInstanceType x -> WhyInvalidInstanceType # | |||||
Show WhyInvalidInstanceType Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods showsPrec :: Int -> WhyInvalidInstanceType -> ShowS # show :: WhyInvalidInstanceType -> String # showList :: [WhyInvalidInstanceType] -> ShowS # | |||||
type Rep WhyInvalidInstanceType Source # | |||||
Defined in Agda.TypeChecking.Monad.Base |
data WhyNotAHaskellType Source #
Extra information for NotAHaskellType
error.
Constructors
NoPragmaFor QName | |
WrongPragmaFor Range QName | |
BadLambda Term | |
BadMeta Term | |
BadDontCare Term | |
NotCompiled QName |
Instances
NFData WhyNotAHaskellType Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: WhyNotAHaskellType -> () # | |||||
Generic WhyNotAHaskellType Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Associated Types
Methods from :: WhyNotAHaskellType -> Rep WhyNotAHaskellType x # to :: Rep WhyNotAHaskellType x -> WhyNotAHaskellType # | |||||
Show WhyNotAHaskellType Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods showsPrec :: Int -> WhyNotAHaskellType -> ShowS # show :: WhyNotAHaskellType -> String # showList :: [WhyNotAHaskellType] -> ShowS # | |||||
type Rep WhyNotAHaskellType Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep WhyNotAHaskellType = D1 ('MetaData "WhyNotAHaskellType" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) ((C1 ('MetaCons "NoPragmaFor" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: (C1 ('MetaCons "WrongPragmaFor" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "BadLambda" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)))) :+: (C1 ('MetaCons "BadMeta" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)) :+: (C1 ('MetaCons "BadDontCare" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)) :+: C1 ('MetaCons "NotCompiled" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName))))) |
data CannotQuoteTerm Source #
Extra information for error CannotQuoteTerm
.
Constructors
CannotQuoteTermHidden | |
CannotQuoteTermNothing |
Instances
NFData CannotQuoteTerm Source # | |||||
Defined in Agda.Interaction.Options.Errors Methods rnf :: CannotQuoteTerm -> () # | |||||
Bounded CannotQuoteTerm Source # | |||||
Defined in Agda.Interaction.Options.Errors | |||||
Enum CannotQuoteTerm Source # | |||||
Defined in Agda.Interaction.Options.Errors Methods succ :: CannotQuoteTerm -> CannotQuoteTerm # pred :: CannotQuoteTerm -> CannotQuoteTerm # toEnum :: Int -> CannotQuoteTerm # fromEnum :: CannotQuoteTerm -> Int # enumFrom :: CannotQuoteTerm -> [CannotQuoteTerm] # enumFromThen :: CannotQuoteTerm -> CannotQuoteTerm -> [CannotQuoteTerm] # enumFromTo :: CannotQuoteTerm -> CannotQuoteTerm -> [CannotQuoteTerm] # enumFromThenTo :: CannotQuoteTerm -> CannotQuoteTerm -> CannotQuoteTerm -> [CannotQuoteTerm] # | |||||
Generic CannotQuoteTerm Source # | |||||
Defined in Agda.Interaction.Options.Errors Associated Types
Methods from :: CannotQuoteTerm -> Rep CannotQuoteTerm x # to :: Rep CannotQuoteTerm x -> CannotQuoteTerm # | |||||
Show CannotQuoteTerm Source # | |||||
Defined in Agda.Interaction.Options.Errors Methods showsPrec :: Int -> CannotQuoteTerm -> ShowS # show :: CannotQuoteTerm -> String # showList :: [CannotQuoteTerm] -> ShowS # | |||||
type Rep CannotQuoteTerm Source # | |||||
Defined in Agda.Interaction.Options.Errors |
data ErasedDatatypeReason Source #
The reason for an ErasedDatatype
error.
Constructors
SeveralConstructors | There are several constructors. |
NoErasedMatches | The flag |
NoK | The K rule is not activated. |
Instances
NFData ErasedDatatypeReason Source # | |||||
Defined in Agda.Interaction.Options.Errors Methods rnf :: ErasedDatatypeReason -> () # | |||||
Bounded ErasedDatatypeReason Source # | |||||
Defined in Agda.Interaction.Options.Errors | |||||
Enum ErasedDatatypeReason Source # | |||||
Defined in Agda.Interaction.Options.Errors Methods succ :: ErasedDatatypeReason -> ErasedDatatypeReason # pred :: ErasedDatatypeReason -> ErasedDatatypeReason # toEnum :: Int -> ErasedDatatypeReason # fromEnum :: ErasedDatatypeReason -> Int # enumFrom :: ErasedDatatypeReason -> [ErasedDatatypeReason] # enumFromThen :: ErasedDatatypeReason -> ErasedDatatypeReason -> [ErasedDatatypeReason] # enumFromTo :: ErasedDatatypeReason -> ErasedDatatypeReason -> [ErasedDatatypeReason] # enumFromThenTo :: ErasedDatatypeReason -> ErasedDatatypeReason -> ErasedDatatypeReason -> [ErasedDatatypeReason] # | |||||
Generic ErasedDatatypeReason Source # | |||||
Defined in Agda.Interaction.Options.Errors Associated Types
Methods from :: ErasedDatatypeReason -> Rep ErasedDatatypeReason x # to :: Rep ErasedDatatypeReason x -> ErasedDatatypeReason # | |||||
Show ErasedDatatypeReason Source # | |||||
Defined in Agda.Interaction.Options.Errors Methods showsPrec :: Int -> ErasedDatatypeReason -> ShowS # show :: ErasedDatatypeReason -> String # showList :: [ErasedDatatypeReason] -> ShowS # | |||||
type Rep ErasedDatatypeReason Source # | |||||
Defined in Agda.Interaction.Options.Errors type Rep ErasedDatatypeReason = D1 ('MetaData "ErasedDatatypeReason" "Agda.Interaction.Options.Errors" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "SeveralConstructors" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "NoErasedMatches" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "NoK" 'PrefixI 'False) (U1 :: Type -> Type))) |
data NotAValidLetBinding Source #
Reasons for error NotAValidLetBinding
.
Constructors
MissingRHS | |
NotAValidLetPattern | |
WhereClausesNotAllowed |
Instances
NFData NotAValidLetBinding Source # | |||||
Defined in Agda.Interaction.Options.Errors Methods rnf :: NotAValidLetBinding -> () # | |||||
Bounded NotAValidLetBinding Source # | |||||
Defined in Agda.Interaction.Options.Errors | |||||
Enum NotAValidLetBinding Source # | |||||
Defined in Agda.Interaction.Options.Errors Methods succ :: NotAValidLetBinding -> NotAValidLetBinding # pred :: NotAValidLetBinding -> NotAValidLetBinding # toEnum :: Int -> NotAValidLetBinding # fromEnum :: NotAValidLetBinding -> Int # enumFrom :: NotAValidLetBinding -> [NotAValidLetBinding] # enumFromThen :: NotAValidLetBinding -> NotAValidLetBinding -> [NotAValidLetBinding] # enumFromTo :: NotAValidLetBinding -> NotAValidLetBinding -> [NotAValidLetBinding] # enumFromThenTo :: NotAValidLetBinding -> NotAValidLetBinding -> NotAValidLetBinding -> [NotAValidLetBinding] # | |||||
Generic NotAValidLetBinding Source # | |||||
Defined in Agda.Interaction.Options.Errors Associated Types
Methods from :: NotAValidLetBinding -> Rep NotAValidLetBinding x # to :: Rep NotAValidLetBinding x -> NotAValidLetBinding # | |||||
Show NotAValidLetBinding Source # | |||||
Defined in Agda.Interaction.Options.Errors Methods showsPrec :: Int -> NotAValidLetBinding -> ShowS # show :: NotAValidLetBinding -> String # showList :: [NotAValidLetBinding] -> ShowS # | |||||
type Rep NotAValidLetBinding Source # | |||||
Defined in Agda.Interaction.Options.Errors type Rep NotAValidLetBinding = D1 ('MetaData "NotAValidLetBinding" "Agda.Interaction.Options.Errors" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "MissingRHS" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "NotAValidLetPattern" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "WhereClausesNotAllowed" 'PrefixI 'False) (U1 :: Type -> Type))) |
data NotAValidLetExpression Source #
Reasons for error NotAValidLetExpression
.
Constructors
MissingBody |
Instances
NFData NotAValidLetExpression Source # | |||||
Defined in Agda.Interaction.Options.Errors Methods rnf :: NotAValidLetExpression -> () # | |||||
Bounded NotAValidLetExpression Source # | |||||
Defined in Agda.Interaction.Options.Errors | |||||
Enum NotAValidLetExpression Source # | |||||
Defined in Agda.Interaction.Options.Errors Methods succ :: NotAValidLetExpression -> NotAValidLetExpression # pred :: NotAValidLetExpression -> NotAValidLetExpression # toEnum :: Int -> NotAValidLetExpression # fromEnum :: NotAValidLetExpression -> Int # enumFrom :: NotAValidLetExpression -> [NotAValidLetExpression] # enumFromThen :: NotAValidLetExpression -> NotAValidLetExpression -> [NotAValidLetExpression] # enumFromTo :: NotAValidLetExpression -> NotAValidLetExpression -> [NotAValidLetExpression] # enumFromThenTo :: NotAValidLetExpression -> NotAValidLetExpression -> NotAValidLetExpression -> [NotAValidLetExpression] # | |||||
Generic NotAValidLetExpression Source # | |||||
Defined in Agda.Interaction.Options.Errors Associated Types
Methods from :: NotAValidLetExpression -> Rep NotAValidLetExpression x # to :: Rep NotAValidLetExpression x -> NotAValidLetExpression # | |||||
Show NotAValidLetExpression Source # | |||||
Defined in Agda.Interaction.Options.Errors Methods showsPrec :: Int -> NotAValidLetExpression -> ShowS # show :: NotAValidLetExpression -> String # showList :: [NotAValidLetExpression] -> ShowS # | |||||
type Rep NotAValidLetExpression Source # | |||||
data NotAllowedInDotPatterns Source #
Things not allowed in dot patterns.
Constructors
LetExpressions | |
PatternLambdas |
Instances
NFData NotAllowedInDotPatterns Source # | |||||
Defined in Agda.Interaction.Options.Errors Methods rnf :: NotAllowedInDotPatterns -> () # | |||||
Bounded NotAllowedInDotPatterns Source # | |||||
Defined in Agda.Interaction.Options.Errors | |||||
Enum NotAllowedInDotPatterns Source # | |||||
Defined in Agda.Interaction.Options.Errors Methods succ :: NotAllowedInDotPatterns -> NotAllowedInDotPatterns # pred :: NotAllowedInDotPatterns -> NotAllowedInDotPatterns # toEnum :: Int -> NotAllowedInDotPatterns # fromEnum :: NotAllowedInDotPatterns -> Int # enumFrom :: NotAllowedInDotPatterns -> [NotAllowedInDotPatterns] # enumFromThen :: NotAllowedInDotPatterns -> NotAllowedInDotPatterns -> [NotAllowedInDotPatterns] # enumFromTo :: NotAllowedInDotPatterns -> NotAllowedInDotPatterns -> [NotAllowedInDotPatterns] # enumFromThenTo :: NotAllowedInDotPatterns -> NotAllowedInDotPatterns -> NotAllowedInDotPatterns -> [NotAllowedInDotPatterns] # | |||||
Generic NotAllowedInDotPatterns Source # | |||||
Defined in Agda.Interaction.Options.Errors Associated Types
Methods from :: NotAllowedInDotPatterns -> Rep NotAllowedInDotPatterns x # to :: Rep NotAllowedInDotPatterns x -> NotAllowedInDotPatterns # | |||||
Show NotAllowedInDotPatterns Source # | |||||
Defined in Agda.Interaction.Options.Errors Methods showsPrec :: Int -> NotAllowedInDotPatterns -> ShowS # show :: NotAllowedInDotPatterns -> String # showList :: [NotAllowedInDotPatterns] -> ShowS # | |||||
type Rep NotAllowedInDotPatterns Source # | |||||
Defined in Agda.Interaction.Options.Errors |
Unique identifier of a file.
Instances
GetFileId FileToId Source # | |||||
GetIdFile IdToFile Source # | |||||
NFData FileId Source # | |||||
Defined in Agda.Utils.FileId | |||||
Enum FileId Source # | |||||
Defined in Agda.Utils.FileId | |||||
Generic FileId Source # | |||||
Defined in Agda.Utils.FileId Associated Types
| |||||
Num FileId Source # | |||||
Show FileId Source # | |||||
Eq FileId Source # | |||||
Ord FileId Source # | |||||
type Rep FileId Source # | |||||
Defined in Agda.Utils.FileId |
class Monad m => MonadFileId (m :: Type -> Type) where Source #
Minimal complete definition
Nothing
Methods
fileFromId :: FileId -> m File Source #
default fileFromId :: forall (t :: (Type -> Type) -> Type -> Type) (n :: Type -> Type). (MonadTrans t, MonadFileId n, m ~ t n) => FileId -> m File Source #
idFromFile :: File -> m FileId Source #
default idFromFile :: forall (t :: (Type -> Type) -> Type -> Type) (n :: Type -> Type). (MonadTrans t, MonadFileId n, m ~ t n) => File -> m FileId Source #
Instances
MonadFileId IM Source # | |
Defined in Agda.Interaction.Monad | |
MonadIO m => MonadFileId (TCMT m) Source # | |
Defined in Agda.TypeChecking.Monad.State | |
MonadFileId m => MonadFileId (ExceptT e m) Source # | |
Defined in Agda.Utils.FileId | |
MonadFileId m => MonadFileId (ReaderT r m) Source # | |
Defined in Agda.Utils.FileId | |
MonadFileId m => MonadFileId (StateT s m) Source # | |
Defined in Agda.Utils.FileId |
class (Functor m, Applicative m, Monad m) => HasOptions (m :: Type -> Type) where Source #
Minimal complete definition
Nothing
Methods
pragmaOptions :: m PragmaOptions Source #
Returns the pragma options which are currently in effect.
default pragmaOptions :: forall (n :: Type -> Type) (t :: (Type -> Type) -> Type -> Type). (HasOptions n, MonadTrans t, m ~ t n) => m PragmaOptions Source #
commandLineOptions :: m CommandLineOptions Source #
Returns the command line options which are currently in effect.
default commandLineOptions :: forall (n :: Type -> Type) (t :: (Type -> Type) -> Type -> Type). (HasOptions n, MonadTrans t, m ~ t n) => m CommandLineOptions Source #
Instances
HasOptions IM Source # | |
Defined in Agda.Interaction.Monad Methods | |
HasOptions TerM Source # | |
Defined in Agda.Termination.Monad Methods | |
HasOptions ReduceM Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
HasOptions NLM Source # | |
Defined in Agda.TypeChecking.Rewriting.NonLinMatch Methods | |
HasOptions m => HasOptions (PureConversionT m) Source # | |
Defined in Agda.TypeChecking.Conversion.Pure | |
HasOptions m => HasOptions (BlockT m) Source # | |
Defined in Agda.TypeChecking.Monad.Base Methods | |
MonadIO m => HasOptions (TCMT m) Source # | |
Defined in Agda.TypeChecking.Monad.Base Methods | |
HasOptions m => HasOptions (NamesT m) Source # | |
Defined in Agda.TypeChecking.Names Methods | |
HasOptions m => HasOptions (ListT m) Source # | |
Defined in Agda.Interaction.Options.HasOptions Methods | |
HasOptions m => HasOptions (ChangeT m) Source # | |
Defined in Agda.Interaction.Options.HasOptions Methods | |
HasOptions m => HasOptions (MaybeT m) Source # | |
Defined in Agda.Interaction.Options.HasOptions Methods | |
HasOptions m => HasOptions (ExceptT e m) Source # | |
Defined in Agda.Interaction.Options.HasOptions Methods pragmaOptions :: ExceptT e m PragmaOptions Source # commandLineOptions :: ExceptT e m CommandLineOptions Source # | |
HasOptions m => HasOptions (IdentityT m) Source # | |
Defined in Agda.Interaction.Options.HasOptions Methods pragmaOptions :: IdentityT m PragmaOptions Source # commandLineOptions :: IdentityT m CommandLineOptions Source # | |
HasOptions m => HasOptions (ReaderT r m) Source # | |
Defined in Agda.Interaction.Options.HasOptions Methods pragmaOptions :: ReaderT r m PragmaOptions Source # commandLineOptions :: ReaderT r m CommandLineOptions Source # | |
HasOptions m => HasOptions (StateT s m) Source # | |
Defined in Agda.Interaction.Options.HasOptions Methods pragmaOptions :: StateT s m PragmaOptions Source # commandLineOptions :: StateT s m CommandLineOptions Source # | |
(HasOptions m, Monoid w) => HasOptions (WriterT w m) Source # | |
Defined in Agda.Interaction.Options.HasOptions Methods pragmaOptions :: WriterT w m PragmaOptions Source # commandLineOptions :: WriterT w m CommandLineOptions Source # |
data RecordFieldWarning Source #
Instances
EmbPrj RecordFieldWarning Source # | |||||
Defined in Agda.TypeChecking.Serialise.Instances.Errors | |||||
NFData RecordFieldWarning Source # | |||||
Defined in Agda.TypeChecking.Monad.Base Methods rnf :: RecordFieldWarning -> () # | |||||
Generic RecordFieldWarning Source # | |||||
Defined in Agda.TypeChecking.Monad.Base.Warning Associated Types
Methods from :: RecordFieldWarning -> Rep RecordFieldWarning x # to :: Rep RecordFieldWarning x -> RecordFieldWarning # | |||||
Show RecordFieldWarning Source # | |||||
Defined in Agda.TypeChecking.Monad.Base.Warning Methods showsPrec :: Int -> RecordFieldWarning -> ShowS # show :: RecordFieldWarning -> String # showList :: [RecordFieldWarning] -> ShowS # | |||||
type Rep RecordFieldWarning Source # | |||||
Defined in Agda.TypeChecking.Monad.Base.Warning type Rep RecordFieldWarning = D1 ('MetaData "RecordFieldWarning" "Agda.TypeChecking.Monad.Base.Warning" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "DuplicateFields" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 (Name, Range)))) :+: C1 ('MetaCons "TooManyFields" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Name]) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 (Name, Range)))))) |
data UselessPublicReason Source #
Instances
EmbPrj UselessPublicReason Source # | |||||
Defined in Agda.TypeChecking.Serialise.Instances.Errors | |||||
NFData UselessPublicReason Source # | |||||
Defined in Agda.TypeChecking.Monad.Base.Warning Methods rnf :: UselessPublicReason -> () # | |||||
Bounded UselessPublicReason Source # | |||||
Defined in Agda.TypeChecking.Monad.Base.Warning | |||||
Enum UselessPublicReason Source # | |||||
Defined in Agda.TypeChecking.Monad.Base.Warning Methods succ :: UselessPublicReason -> UselessPublicReason # pred :: UselessPublicReason -> UselessPublicReason # toEnum :: Int -> UselessPublicReason # fromEnum :: UselessPublicReason -> Int # enumFrom :: UselessPublicReason -> [UselessPublicReason] # enumFromThen :: UselessPublicReason -> UselessPublicReason -> [UselessPublicReason] # enumFromTo :: UselessPublicReason -> UselessPublicReason -> [UselessPublicReason] # enumFromThenTo :: UselessPublicReason -> UselessPublicReason -> UselessPublicReason -> [UselessPublicReason] # | |||||
Generic UselessPublicReason Source # | |||||
Defined in Agda.TypeChecking.Monad.Base.Warning Associated Types
Methods from :: UselessPublicReason -> Rep UselessPublicReason x # to :: Rep UselessPublicReason x -> UselessPublicReason # | |||||
Show UselessPublicReason Source # | |||||
Defined in Agda.TypeChecking.Monad.Base.Warning Methods showsPrec :: Int -> UselessPublicReason -> ShowS # show :: UselessPublicReason -> String # showList :: [UselessPublicReason] -> ShowS # | |||||
type Rep UselessPublicReason Source # | |||||
Defined in Agda.TypeChecking.Monad.Base.Warning type Rep UselessPublicReason = D1 ('MetaData "UselessPublicReason" "Agda.TypeChecking.Monad.Base.Warning" "Agda-2.8.0-inplace" 'False) ((C1 ('MetaCons "UselessPublicPreamble" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "UselessPublicLet" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "UselessPublicNoOpen" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "UselessPublicAnonymousModule" 'PrefixI 'False) (U1 :: Type -> Type))) |
Orphan instances
Pretty Comparison Source # | |
Methods pretty :: Comparison -> Doc Source # prettyPrec :: Int -> Comparison -> Doc Source # prettyList :: [Comparison] -> Doc Source # | |
Pretty ModuleToSource Source # | |
Methods pretty :: ModuleToSource -> Doc Source # prettyPrec :: Int -> ModuleToSource -> Doc Source # prettyList :: [ModuleToSource] -> Doc Source # | |
Pretty NamedMeta Source # | |
Pretty Polarity Source # | |
KillRange Polarity Source # | |
Methods | |
GetFileId FileDictWithBuiltins Source # | |
GetIdFile FileDictWithBuiltins Source # | |
NFData Comparison Source # | |
Methods rnf :: Comparison -> () # | |
NFData HighlightingLevel Source # | |
Methods rnf :: HighlightingLevel -> () # | |
NFData HighlightingMethod Source # | |
Methods rnf :: HighlightingMethod -> () # | |
NFData Polarity Source # | |
NFData RecordFieldWarning Source # | |
Methods rnf :: RecordFieldWarning -> () # | |
NFData (BiMap RawTopLevelModuleName ModuleNameHash) Source # | |
Methods rnf :: BiMap RawTopLevelModuleName ModuleNameHash -> () # |