Safe Haskell | None |
---|---|
Language | Haskell2010 |
- Type checking state
- st-prefixed lenses
- Fresh things
- Open things
- Judgements
- Highlighting levels
- Type checking environment
- e-prefixed lenses
- Type checking warnings (aka non-fatal errors)
- Type checking errors
- Accessing options
- The reduce monad
- Monad with read-only
TCEnv
- Monad with mutable
TCState
- Type checking monad transformer
- Names for generated definitions
- KillRange instances
Synopsis
- data TCState = TCSt {}
- class Monad m => ReadTCState m where
- getTCState :: m TCState
- locallyTCState :: Lens' a TCState -> (a -> a) -> m b -> m b
- withTCState :: (TCState -> TCState) -> m a -> m a
- data PreScopeState = PreScopeState {
- stPreTokens :: !CompressedFile
- stPreImports :: !Signature
- stPreImportedModules :: !(Set ModuleName)
- stPreModuleToSource :: !ModuleToSource
- stPreVisitedModules :: !VisitedModules
- stPreScope :: !ScopeInfo
- stPrePatternSyns :: !PatternSynDefns
- stPrePatternSynImports :: !PatternSynDefns
- stPreGeneralizedVars :: !(Maybe (Set QName))
- stPrePragmaOptions :: !PragmaOptions
- stPreImportedBuiltins :: !(BuiltinThings PrimFun)
- stPreImportedDisplayForms :: !DisplayForms
- stPreImportedInstanceDefs :: !InstanceTable
- stPreForeignCode :: !(Map BackendName [ForeignCode])
- stPreFreshInteractionId :: !InteractionId
- stPreImportedUserWarnings :: !(Map QName String)
- stPreLocalUserWarnings :: !(Map QName String)
- stPreWarningOnImport :: !(Maybe String)
- stPreImportedPartialDefs :: !(Set QName)
- type DisambiguatedNames = IntMap QName
- type ConcreteNames = Map Name [Name]
- data PostScopeState = PostScopeState {
- stPostSyntaxInfo :: !CompressedFile
- stPostDisambiguatedNames :: !DisambiguatedNames
- stPostMetaStore :: !MetaStore
- stPostInteractionPoints :: !InteractionPoints
- stPostAwakeConstraints :: !Constraints
- stPostSleepingConstraints :: !Constraints
- stPostDirty :: !Bool
- stPostOccursCheckDefs :: !(Set QName)
- stPostSignature :: !Signature
- stPostModuleCheckpoints :: !(Map ModuleName CheckpointId)
- stPostImportsDisplayForms :: !DisplayForms
- stPostCurrentModule :: !(Maybe ModuleName)
- stPostInstanceDefs :: !TempInstanceTable
- stPostConcreteNames :: !ConcreteNames
- stPostUsedNames :: !(Map RawName [RawName])
- stPostShadowingNames :: !(Map Name [RawName])
- stPostStatistics :: !Statistics
- stPostTCWarnings :: ![TCWarning]
- stPostMutualBlocks :: !(Map MutualId MutualBlock)
- stPostLocalBuiltins :: !(BuiltinThings PrimFun)
- stPostFreshMetaId :: !MetaId
- stPostFreshMutualId :: !MutualId
- stPostFreshProblemId :: !ProblemId
- stPostFreshCheckpointId :: !CheckpointId
- stPostFreshInt :: !Int
- stPostFreshNameId :: !NameId
- stPostAreWeCaching :: !Bool
- stPostPostponeInstanceSearch :: !Bool
- stPostConsideringInstance :: !Bool
- stPostInstantiateBlocking :: !Bool
- stPostLocalPartialDefs :: !(Set QName)
- data MutualBlock = MutualBlock {}
- data PersistentTCState = PersistentTCSt {}
- data LoadedFileCache = LoadedFileCache {}
- type CachedTypeCheckLog = [(TypeCheckAction, PostScopeState)]
- type CurrentTypeCheckLog = [(TypeCheckAction, PostScopeState)]
- data TypeCheckAction
- initPersistentState :: PersistentTCState
- initPreScopeState :: PreScopeState
- initPostScopeState :: PostScopeState
- initState :: TCState
- stTokens :: Lens' CompressedFile TCState
- stImports :: Lens' Signature TCState
- stImportedModules :: Lens' (Set ModuleName) TCState
- stModuleToSource :: Lens' ModuleToSource TCState
- stVisitedModules :: Lens' VisitedModules TCState
- stScope :: Lens' ScopeInfo TCState
- stPatternSyns :: Lens' PatternSynDefns TCState
- stPatternSynImports :: Lens' PatternSynDefns TCState
- stGeneralizedVars :: Lens' (Maybe (Set QName)) TCState
- stPragmaOptions :: Lens' PragmaOptions TCState
- stImportedBuiltins :: Lens' (BuiltinThings PrimFun) TCState
- stForeignCode :: Lens' (Map BackendName [ForeignCode]) TCState
- stFreshInteractionId :: Lens' InteractionId TCState
- stImportedUserWarnings :: Lens' (Map QName String) TCState
- stLocalUserWarnings :: Lens' (Map QName String) TCState
- getUserWarnings :: ReadTCState m => m (Map QName String)
- stWarningOnImport :: Lens' (Maybe String) TCState
- stImportedPartialDefs :: Lens' (Set QName) TCState
- stLocalPartialDefs :: Lens' (Set QName) TCState
- getPartialDefs :: ReadTCState m => m (Set QName)
- stLoadedFileCache :: Lens' (Maybe LoadedFileCache) TCState
- stBackends :: Lens' [Backend] TCState
- stFreshNameId :: Lens' NameId TCState
- stSyntaxInfo :: Lens' CompressedFile TCState
- stDisambiguatedNames :: Lens' DisambiguatedNames TCState
- stMetaStore :: Lens' MetaStore TCState
- stInteractionPoints :: Lens' InteractionPoints TCState
- stAwakeConstraints :: Lens' Constraints TCState
- stSleepingConstraints :: Lens' Constraints TCState
- stDirty :: Lens' Bool TCState
- stOccursCheckDefs :: Lens' (Set QName) TCState
- stSignature :: Lens' Signature TCState
- stModuleCheckpoints :: Lens' (Map ModuleName CheckpointId) TCState
- stImportsDisplayForms :: Lens' DisplayForms TCState
- stImportedDisplayForms :: Lens' DisplayForms TCState
- stCurrentModule :: Lens' (Maybe ModuleName) TCState
- stImportedInstanceDefs :: Lens' InstanceTable TCState
- stInstanceDefs :: Lens' TempInstanceTable TCState
- stConcreteNames :: Lens' ConcreteNames TCState
- stUsedNames :: Lens' (Map RawName [RawName]) TCState
- stShadowingNames :: Lens' (Map Name [RawName]) TCState
- stStatistics :: Lens' Statistics TCState
- stTCWarnings :: Lens' [TCWarning] TCState
- stMutualBlocks :: Lens' (Map MutualId MutualBlock) TCState
- stLocalBuiltins :: Lens' (BuiltinThings PrimFun) TCState
- stFreshMetaId :: Lens' MetaId TCState
- stFreshMutualId :: Lens' MutualId TCState
- stFreshProblemId :: Lens' ProblemId TCState
- stFreshCheckpointId :: Lens' CheckpointId TCState
- stFreshInt :: Lens' Int TCState
- stAreWeCaching :: Lens' Bool TCState
- stPostponeInstanceSearch :: Lens' Bool TCState
- stConsideringInstance :: Lens' Bool TCState
- stInstantiateBlocking :: Lens' Bool TCState
- stBuiltinThings :: TCState -> BuiltinThings PrimFun
- class Enum i => HasFresh i where
- freshLens :: Lens' i TCState
- nextFresh' :: i -> i
- nextFresh :: HasFresh i => TCState -> (i, TCState)
- class Monad m => MonadFresh i m where
- fresh :: m i
- newtype ProblemId = ProblemId Nat
- newtype CheckpointId = CheckpointId Int
- 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
- class FreshName a where
- freshName_ :: MonadFresh NameId m => a -> m Name
- type ModuleToSource = Map TopLevelModuleName AbsolutePath
- type SourceToModule = Map AbsolutePath TopLevelModuleName
- sourceToModule :: TCM SourceToModule
- lookupModuleFromSource :: ReadTCState m => AbsolutePath -> m (Maybe TopLevelModuleName)
- class Monad m => MonadStConcreteNames m where
- runStConcreteNames :: StateT ConcreteNames m a -> m a
- useConcreteNames :: m ConcreteNames
- modifyConcreteNames :: (ConcreteNames -> ConcreteNames) -> m ()
- data ModuleInfo = ModuleInfo {}
- type VisitedModules = Map TopLevelModuleName ModuleInfo
- type DecodedModules = Map TopLevelModuleName Interface
- data ForeignCode = ForeignCode Range String
- data Interface = Interface {
- iSourceHash :: Hash
- iSource :: Text
- iFileType :: FileType
- iImportedModules :: [(ModuleName, Hash)]
- iModuleName :: ModuleName
- iScope :: Map ModuleName Scope
- iInsideScope :: ScopeInfo
- iSignature :: Signature
- iDisplayForms :: DisplayForms
- iUserWarnings :: Map QName String
- iImportWarning :: Maybe String
- iBuiltin :: BuiltinThings (String, QName)
- iForeignCode :: Map BackendName [ForeignCode]
- iHighlighting :: HighlightingInfo
- iPragmaOptions :: [OptionsPragma]
- iOptionsUsed :: PragmaOptions
- iPatternSyns :: PatternSynDefns
- iWarnings :: [TCWarning]
- iPartialDefs :: Set QName
- iFullHash :: Interface -> Hash
- data Closure a = Closure {}
- class LensClosure a b | b -> a where
- lensClosure :: Lens' (Closure a) b
- buildClosure :: (MonadTCEnv m, ReadTCState m) => a -> m (Closure a)
- type Constraints = [ProblemConstraint]
- data ProblemConstraint = PConstr {}
- data Constraint
- = ValueCmp Comparison CompareAs Term Term
- | ValueCmpOnFace Comparison Term Type Term Term
- | ElimCmp [Polarity] [IsForced] Type Term [Elim] [Elim]
- | TelCmp Type Type Comparison Telescope Telescope
- | SortCmp Comparison Sort Sort
- | LevelCmp Comparison Level Level
- | HasBiggerSort Sort
- | HasPTSRule (Dom Type) (Abs Sort)
- | CheckMetaInst MetaId
- | UnBlock MetaId
- | Guarded Constraint ProblemId
- | IsEmpty Range Type
- | CheckSizeLtSat Term
- | FindInstance MetaId (Maybe MetaId) (Maybe [Candidate])
- | CheckFunDef Delayed DefInfo QName [Clause]
- | UnquoteTactic (Maybe MetaId) Term Term Type
- data Comparison
- data CompareDirection
- fromCmp :: Comparison -> CompareDirection
- flipCmp :: CompareDirection -> CompareDirection
- dirToCmp :: (Comparison -> a -> a -> c) -> CompareDirection -> a -> a -> c
- data CompareAs
- data Open a = OpenThing {}
- data Judgement a
- = HasType {
- jMetaId :: a
- jComparison :: Comparison
- jMetaType :: Type
- | IsSort { }
- = HasType {
- data DoGeneralize
- data GeneralizedValue = GeneralizedValue {}
- data MetaVariable = MetaVar {}
- data Listener
- data Frozen
- data MetaInstantiation
- data CheckedTarget
- data TypeCheckingProblem
- = CheckExpr Comparison Expr Type
- | CheckArgs ExpandHidden Range [NamedArg Expr] Type Type ([Maybe Range] -> Elims -> Type -> CheckedTarget -> TCM Term)
- | CheckProjAppToKnownPrincipalArg Comparison Expr ProjOrigin (NonEmpty QName) Args Type Int Term Type
- | CheckLambda Comparison (Arg ([WithHiding Name], Maybe Type)) Expr Type
- | DoQuoteTerm Comparison Term Type
- newtype MetaPriority = MetaPriority Int
- data RunMetaOccursCheck
- data MetaInfo = MetaInfo {}
- type MetaNameSuggestion = String
- data NamedMeta = NamedMeta {}
- type MetaStore = IntMap MetaVariable
- normalMetaPriority :: MetaPriority
- lowMetaPriority :: MetaPriority
- highMetaPriority :: MetaPriority
- getMetaInfo :: MetaVariable -> Closure Range
- getMetaScope :: MetaVariable -> ScopeInfo
- getMetaEnv :: MetaVariable -> TCEnv
- getMetaSig :: MetaVariable -> Signature
- getMetaRelevance :: MetaVariable -> Relevance
- getMetaModality :: MetaVariable -> Modality
- metaFrozen :: Lens' Frozen MetaVariable
- _mvInfo :: Lens' MetaInfo MetaVariable
- data InteractionPoint = InteractionPoint {}
- type InteractionPoints = Map InteractionId InteractionPoint
- data Overapplied
- data IPBoundary' t = IPBoundary {
- ipbEquations :: [(t, t)]
- ipbValue :: t
- ipbMetaApp :: t
- ipbOverapplied :: Overapplied
- type IPBoundary = IPBoundary' Term
- data IPClause
- = IPClause {
- ipcQName :: QName
- ipcClauseNo :: Int
- ipcType :: Type
- ipcWithSub :: Maybe Substitution
- ipcClause :: SpineClause
- ipcClosure :: Closure ()
- ipcBoundary :: [Closure IPBoundary]
- | IPNoClause
- = IPClause {
- data Signature = Sig {}
- sigSections :: Lens' Sections Signature
- sigDefinitions :: Lens' Definitions Signature
- sigRewriteRules :: Lens' RewriteRuleMap Signature
- type Sections = Map ModuleName Section
- type Definitions = HashMap QName Definition
- type RewriteRuleMap = HashMap QName RewriteRules
- type DisplayForms = HashMap QName [LocalDisplayForm]
- newtype Section = Section {}
- secTelescope :: Lens' Telescope Section
- emptySignature :: Signature
- data DisplayForm = Display {
- dfFreeVars :: Nat
- dfPats :: Elims
- dfRHS :: DisplayTerm
- type LocalDisplayForm = Open DisplayForm
- data DisplayTerm
- = DWithApp DisplayTerm [DisplayTerm] Elims
- | DCon ConHead ConInfo [Arg DisplayTerm]
- | DDef QName [Elim' DisplayTerm]
- | DDot Term
- | DTerm Term
- defaultDisplayForm :: QName -> [LocalDisplayForm]
- data NLPat
- type PElims = [Elim' NLPat]
- data NLPType = NLPType {}
- data NLPSort
- type RewriteRules = [RewriteRule]
- data RewriteRule = RewriteRule {}
- data Definition = Defn {
- defArgInfo :: ArgInfo
- defName :: QName
- defType :: Type
- defPolarity :: [Polarity]
- defArgOccurrences :: [Occurrence]
- defArgGeneralizable :: NumGeneralizableArgs
- defGeneralizedParams :: [Maybe Name]
- defDisplay :: [LocalDisplayForm]
- defMutual :: MutualId
- defCompiledRep :: CompiledRepresentation
- defInstance :: Maybe QName
- defCopy :: Bool
- defMatchable :: Set QName
- defNoCompilation :: Bool
- defInjective :: Bool
- defCopatternLHS :: Bool
- defBlocked :: Blocked_
- theDef :: Defn
- data NumGeneralizableArgs
- theDefLens :: Lens' Defn Definition
- defaultDefn :: ArgInfo -> QName -> Type -> Defn -> Definition
- data Polarity
- data IsForced
- data CompilerPragma = CompilerPragma Range String
- type BackendName = String
- jsBackendName :: BackendName
- ghcBackendName :: BackendName
- type CompiledRepresentation = Map BackendName [CompilerPragma]
- noCompiledRep :: CompiledRepresentation
- type Face = [(Term, Bool)]
- data System = System {
- systemTel :: Telescope
- systemClauses :: [(Face, Term)]
- data ExtLamInfo = ExtLamInfo {
- extLamModule :: ModuleName
- extLamSys :: !(Maybe System)
- modifySystem :: (System -> System) -> ExtLamInfo -> ExtLamInfo
- data Projection = Projection {}
- newtype ProjLams = ProjLams {
- getProjLams :: [Arg ArgName]
- projDropPars :: Projection -> ProjOrigin -> Term
- projArgInfo :: Projection -> ArgInfo
- data EtaEquality
- = Specified {
- theEtaEquality :: !HasEta
- | Inferred {
- theEtaEquality :: !HasEta
- = Specified {
- setEtaEquality :: EtaEquality -> HasEta -> EtaEquality
- data FunctionFlag
- data CompKit = CompKit {}
- emptyCompKit :: CompKit
- data Defn
- = Axiom
- | DataOrRecSig {
- datarecPars :: Int
- | GeneralizableVar
- | AbstractDefn Defn
- | Function {
- funClauses :: [Clause]
- funCompiled :: Maybe CompiledClauses
- funSplitTree :: Maybe SplitTree
- funTreeless :: Maybe Compiled
- funCovering :: [Clause]
- funInv :: FunctionInverse
- funMutual :: Maybe [QName]
- funAbstr :: IsAbstract
- funDelayed :: Delayed
- funProjection :: Maybe Projection
- funFlags :: Set FunctionFlag
- funTerminates :: Maybe Bool
- funExtLam :: Maybe ExtLamInfo
- funWith :: Maybe QName
- | Datatype {
- dataPars :: Nat
- dataIxs :: Nat
- dataClause :: Maybe Clause
- dataCons :: [QName]
- dataSort :: Sort
- dataMutual :: Maybe [QName]
- dataAbstr :: IsAbstract
- dataPathCons :: [QName]
- | Record {
- recPars :: Nat
- recClause :: Maybe Clause
- recConHead :: ConHead
- recNamedCon :: Bool
- recFields :: [Dom QName]
- recTel :: Telescope
- recMutual :: Maybe [QName]
- recEtaEquality' :: EtaEquality
- recInduction :: Maybe Induction
- recAbstr :: IsAbstract
- recComp :: CompKit
- | Constructor { }
- | Primitive { }
- recRecursive :: Defn -> Bool
- recEtaEquality :: Defn -> HasEta
- emptyFunction :: Defn
- funFlag :: FunctionFlag -> Lens' Bool Defn
- funStatic :: Lens' Bool Defn
- funInline :: Lens' Bool Defn
- funMacro :: Lens' Bool Defn
- isMacro :: Defn -> Bool
- isEmptyFunction :: Defn -> Bool
- isCopatternLHS :: [Clause] -> Bool
- recCon :: Defn -> QName
- defIsRecord :: Defn -> Bool
- defIsDataOrRecord :: Defn -> Bool
- defConstructors :: Defn -> [QName]
- newtype Fields = Fields [(Name, Type)]
- data Simplification
- data Reduced no yes
- = NoReduction no
- | YesReduction Simplification yes
- redReturn :: a -> ReduceM (Reduced a' a)
- redBind :: ReduceM (Reduced a a') -> (a -> b) -> (a' -> ReduceM (Reduced b b')) -> ReduceM (Reduced b b')
- data IsReduced
- = NotReduced
- | Reduced (Blocked ())
- data MaybeReduced a = MaybeRed {
- isReduced :: IsReduced
- ignoreReduced :: a
- type MaybeReducedArgs = [MaybeReduced (Arg Term)]
- type MaybeReducedElims = [MaybeReduced Elim]
- notReduced :: a -> MaybeReduced a
- reduced :: Blocked (Arg Term) -> MaybeReduced (Arg Term)
- data AllowedReduction
- type AllowedReductions = SmallSet AllowedReduction
- allReductions :: AllowedReductions
- reallyAllReductions :: AllowedReductions
- data PrimitiveImpl = PrimImpl Type PrimFun
- data PrimFun = PrimFun {
- primFunName :: QName
- primFunArity :: Arity
- primFunImplementation :: [Arg Term] -> Int -> ReduceM (Reduced MaybeReducedArgs Term)
- primFun :: QName -> Arity -> ([Arg Term] -> ReduceM (Reduced MaybeReducedArgs Term)) -> PrimFun
- defClauses :: Definition -> [Clause]
- defCompiled :: Definition -> Maybe CompiledClauses
- defParameters :: Definition -> Maybe Nat
- defInverse :: Definition -> FunctionInverse
- defCompilerPragmas :: BackendName -> Definition -> [CompilerPragma]
- defDelayed :: Definition -> Delayed
- defNonterminating :: Definition -> Bool
- defTerminationUnconfirmed :: Definition -> Bool
- defAbstract :: Definition -> IsAbstract
- defForced :: Definition -> [IsForced]
- type FunctionInverse = FunctionInverse' Clause
- type InversionMap c = Map TermHead [c]
- data FunctionInverse' c
- = NotInjective
- | Inverse (InversionMap c)
- data TermHead
- newtype MutualId = MutId Int32
- type Statistics = Map String Integer
- data Call
- = CheckClause Type SpineClause
- | CheckLHS SpineLHS
- | CheckPattern Pattern Telescope Type
- | 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 Range [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
- | CheckConstructorFitsIn QName Type Sort
- | CheckFunDefCall Range QName [Clause]
- | CheckPragma Range Pragma
- | CheckPrimitive Range QName Expr
- | CheckIsEmpty Range Type
- | CheckConfluence QName QName
- | CheckWithFunctionType Type
- | CheckSectionApplication Range ModuleName ModuleApplication
- | CheckNamedWhere ModuleName
- | ScopeCheckExpr Expr
- | ScopeCheckDeclaration NiceDeclaration
- | ScopeCheckLHS QName Pattern
- | NoHighlighting
- | ModuleContents
- | SetRange Range
- type InstanceTable = Map QName (Set QName)
- type TempInstanceTable = (InstanceTable, Set QName)
- data BuiltinDescriptor
- = BuiltinData (TCM Type) [String]
- | BuiltinDataCons (TCM Type)
- | BuiltinPrim String (Term -> TCM ())
- | BuiltinPostulate Relevance (TCM Type)
- | BuiltinUnknown (Maybe (TCM Type)) (Term -> Type -> TCM ())
- data BuiltinInfo = BuiltinInfo {}
- type BuiltinThings pf = Map String (Builtin pf)
- data Builtin pf
- data HighlightingLevel
- data HighlightingMethod
- ifTopLevelAndHighlightingLevelIsOr :: MonadTCM tcm => HighlightingLevel -> Bool -> tcm () -> tcm ()
- ifTopLevelAndHighlightingLevelIs :: MonadTCM tcm => HighlightingLevel -> tcm () -> tcm ()
- data TCEnv = TCEnv {
- envContext :: Context
- envLetBindings :: LetBindings
- envCurrentModule :: ModuleName
- envCurrentPath :: Maybe AbsolutePath
- envAnonymousModules :: [(ModuleName, Nat)]
- envImportPath :: [TopLevelModuleName]
- envMutualBlock :: Maybe MutualId
- envTerminationCheck :: TerminationCheck ()
- envCoverageCheck :: CoverageCheck
- envMakeCase :: Bool
- envSolvingConstraints :: Bool
- envCheckingWhere :: Bool
- envWorkingOnTypes :: Bool
- envAssignMetas :: Bool
- envActiveProblems :: Set ProblemId
- envAbstractMode :: AbstractMode
- envModality :: Modality
- envDisplayFormsEnabled :: Bool
- envRange :: Range
- envHighlightingRange :: Range
- envClause :: IPClause
- envCall :: Maybe (Closure Call)
- envHighlightingLevel :: HighlightingLevel
- envHighlightingMethod :: HighlightingMethod
- envModuleNestingLevel :: !Int
- envExpandLast :: ExpandHidden
- envAppDef :: Maybe QName
- envSimplification :: Simplification
- envAllowedReductions :: AllowedReductions
- 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
- envCheckOptionConsistency :: Bool
- envActiveBackendName :: Maybe BackendName
- initEnv :: TCEnv
- class LensTCEnv a where
- data UnquoteFlags = UnquoteFlags {}
- defaultUnquoteFlags :: UnquoteFlags
- unquoteNormalise :: Lens' Bool UnquoteFlags
- eUnquoteNormalise :: Lens' Bool TCEnv
- eContext :: Lens' Context TCEnv
- eLetBindings :: Lens' LetBindings TCEnv
- eCurrentModule :: Lens' ModuleName TCEnv
- eCurrentPath :: Lens' (Maybe AbsolutePath) TCEnv
- eAnonymousModules :: Lens' [(ModuleName, Nat)] TCEnv
- eImportPath :: Lens' [TopLevelModuleName] TCEnv
- eMutualBlock :: Lens' (Maybe MutualId) TCEnv
- eTerminationCheck :: Lens' (TerminationCheck ()) TCEnv
- eCoverageCheck :: Lens' CoverageCheck TCEnv
- eMakeCase :: Lens' Bool TCEnv
- eSolvingConstraints :: Lens' Bool TCEnv
- eCheckingWhere :: Lens' Bool TCEnv
- eWorkingOnTypes :: Lens' Bool TCEnv
- eAssignMetas :: Lens' Bool TCEnv
- eActiveProblems :: Lens' (Set ProblemId) TCEnv
- eAbstractMode :: Lens' AbstractMode TCEnv
- eModality :: Lens' Modality TCEnv
- eRelevance :: Lens' Relevance TCEnv
- eQuantity :: Lens' Quantity TCEnv
- eDisplayFormsEnabled :: Lens' Bool TCEnv
- eRange :: Lens' Range TCEnv
- eHighlightingRange :: Lens' Range TCEnv
- eCall :: Lens' (Maybe (Closure Call)) TCEnv
- eHighlightingLevel :: Lens' HighlightingLevel TCEnv
- eHighlightingMethod :: Lens' HighlightingMethod TCEnv
- eModuleNestingLevel :: Lens' Int TCEnv
- eExpandLast :: Lens' ExpandHidden TCEnv
- eAppDef :: Lens' (Maybe QName) TCEnv
- eSimplification :: Lens' Simplification TCEnv
- eAllowedReductions :: Lens' AllowedReductions TCEnv
- eInjectivityDepth :: Lens' Int TCEnv
- eCompareBlocked :: Lens' Bool TCEnv
- ePrintDomainFreePi :: Lens' Bool TCEnv
- eInsideDotPattern :: Lens' Bool TCEnv
- eUnquoteFlags :: Lens' UnquoteFlags TCEnv
- eInstanceDepth :: Lens' Int TCEnv
- eIsDebugPrinting :: Lens' Bool TCEnv
- ePrintingPatternLambdas :: Lens' [QName] TCEnv
- eCallByNeed :: Lens' Bool TCEnv
- eCurrentCheckpoint :: Lens' CheckpointId TCEnv
- eCheckpoints :: Lens' (Map CheckpointId Substitution) TCEnv
- eGeneralizeMetas :: Lens' DoGeneralize TCEnv
- eGeneralizedVars :: Lens' (Map QName GeneralizedValue) TCEnv
- eActiveBackendName :: Lens' (Maybe BackendName) TCEnv
- type Context = [ContextEntry]
- type ContextEntry = Dom (Name, Type)
- type LetBindings = Map Name (Open (Term, Dom Type))
- data AbstractMode
- aDefToMode :: IsAbstract -> AbstractMode
- aModeToDef :: AbstractMode -> Maybe IsAbstract
- data ExpandHidden
- isDontExpandLast :: ExpandHidden -> Bool
- data Candidate = Candidate {}
- data Warning
- = NicifierIssue DeclarationWarning
- | TerminationIssue [TerminationError]
- | UnreachableClauses QName [Range]
- | CoverageIssue QName [(Telescope, [NamedArg DeBruijnPattern])]
- | CoverageNoExactSplit QName [Clause]
- | NotStrictlyPositive QName (Seq OccursWhere)
- | UnsolvedMetaVariables [Range]
- | UnsolvedInteractionMetas [Range]
- | UnsolvedConstraints Constraints
- | CantGeneralizeOverSorts [MetaId]
- | AbsurdPatternRequiresNoRHS [NamedArg DeBruijnPattern]
- | OldBuiltin String String
- | EmptyRewritePragma
- | IllformedAsClause String
- | ClashesViaRenaming NameOrModule [Name]
- | UselessPublic
- | UselessInline QName
- | WrongInstanceDeclaration
- | InstanceWithExplicitArg QName
- | InstanceNoOutputTypeName Doc
- | InstanceArgWithExplicitArg Doc
- | InversionDepthReached QName
- | GenericWarning Doc
- | GenericNonFatalError Doc
- | SafeFlagPostulate Name
- | SafeFlagPragma [String]
- | SafeFlagNonTerminating
- | SafeFlagTerminating
- | SafeFlagWithoutKFlagPrimEraseEquality
- | WithoutKFlagPrimEraseEquality
- | SafeFlagNoPositivityCheck
- | SafeFlagPolarity
- | SafeFlagNoUniverseCheck
- | SafeFlagNoCoverageCheck
- | SafeFlagInjective
- | SafeFlagEta
- | ParseWarning ParseWarning
- | LibraryWarning LibWarning
- | DeprecationWarning String String String
- | UserWarning String
- | FixityInRenamingModule (NonEmpty Range)
- | ModuleDoesntExport QName [ImportedName]
- | InfectiveImport String ModuleName
- | CoInfectiveImport String ModuleName
- | RewriteNonConfluent Term Term Term Doc
- | RewriteMaybeNonConfluent Term Term [Doc]
- | PragmaCompileErased BackendName QName
- | NotInScopeW [QName]
- warningName :: Warning -> WarningName
- data TCWarning = TCWarning {}
- tcWarningOrigin :: TCWarning -> SrcFile
- data CallInfo = CallInfo {}
- data TerminationError = TerminationError {
- termErrFunctions :: [QName]
- termErrCalls :: [CallInfo]
- data SplitError
- = NotADatatype (Closure Type)
- | IrrelevantDatatype (Closure Type)
- | ErasedDatatype Bool (Closure Type)
- | CoinductiveDatatype (Closure Type)
- | UnificationStuck { }
- | CosplitCatchall
- | CosplitNoTarget
- | CosplitNoRecordType (Closure Type)
- | CannotCreateMissingClause QName (Telescope, [NamedArg DeBruijnPattern]) Doc (Closure (Abs Type))
- | GenericSplitError String
- data NegativeUnification
- data UnificationFailure
- data UnquoteError
- data TypeError
- = InternalError String
- | NotImplemented String
- | NotSupported String
- | CompilationError String
- | PropMustBeSingleton
- | DataMustEndInSort Term
- | ShouldEndInApplicationOfTheDatatype Type
- | ShouldBeAppliedToTheDatatypeParameters Term Term
- | ShouldBeApplicationOf Type QName
- | ConstructorPatternInWrongDatatype QName QName
- | CantResolveOverloadedConstructorsTargetingSameDatatype QName [QName]
- | DoesNotConstructAnElementOf QName Type
- | WrongHidingInLHS
- | WrongHidingInLambda Type
- | WrongHidingInApplication Type
- | WrongNamedArgument (NamedArg Expr) [NamedName]
- | WrongIrrelevanceInLambda
- | WrongQuantityInLambda
- | WrongCohesionInLambda
- | QuantityMismatch Quantity Quantity
- | HidingMismatch Hiding Hiding
- | RelevanceMismatch Relevance Relevance
- | UninstantiatedDotPattern Expr
- | ForcedConstructorNotInstantiated Pattern
- | IlltypedPattern Pattern Type
- | IllformedProjectionPattern Pattern
- | CannotEliminateWithPattern (NamedArg Pattern) Type
- | WrongNumberOfConstructorArguments QName Nat Nat
- | ShouldBeEmpty Type [DeBruijnPattern]
- | ShouldBeASort Type
- | ShouldBePi Type
- | ShouldBePath Type
- | ShouldBeRecordType Type
- | ShouldBeRecordPattern DeBruijnPattern
- | NotAProjectionPattern (NamedArg Pattern)
- | NotAProperTerm
- | InvalidTypeSort Sort
- | InvalidType Term
- | FunctionTypeInSizeUniv Term
- | SplitOnIrrelevant (Dom Type)
- | SplitOnUnusableCohesion (Dom Type)
- | SplitOnNonVariable Term Type
- | DefinitionIsIrrelevant QName
- | DefinitionIsErased QName
- | VariableIsIrrelevant Name
- | VariableIsErased Name
- | VariableIsOfUnusableCohesion Name Cohesion
- | UnequalLevel Comparison Level Level
- | UnequalTerms Comparison Term Term CompareAs
- | UnequalTypes Comparison Type Type
- | UnequalRelevance Comparison Term Term
- | UnequalQuantity Comparison Term Term
- | UnequalCohesion Comparison Term Term
- | UnequalHiding Term Term
- | UnequalSorts Sort Sort
- | UnequalBecauseOfUniverseConflict Comparison Term Term
- | NotLeqSort Sort Sort
- | MetaCannotDependOn MetaId Nat
- | MetaOccursInItself MetaId
- | MetaIrrelevantSolution MetaId Term
- | MetaErasedSolution MetaId Term
- | GenericError String
- | GenericDocError Doc
- | BuiltinMustBeConstructor String Expr
- | NoSuchBuiltinName String
- | DuplicateBuiltinBinding String Term Term
- | NoBindingForBuiltin String
- | NoSuchPrimitiveFunction String
- | DuplicatePrimitiveBinding String QName QName
- | ShadowedModule Name [ModuleName]
- | BuiltinInParameterisedModule String
- | IllegalLetInTelescope TypedBinding
- | IllegalPatternInTelescope Binder
- | NoRHSRequiresAbsurdPattern [NamedArg Pattern]
- | TooManyFields QName [Name] [Name]
- | DuplicateFields [Name]
- | DuplicateConstructors [Name]
- | WithOnFreeVariable Expr Term
- | UnexpectedWithPatterns [Pattern]
- | WithClausePatternMismatch Pattern (NamedArg DeBruijnPattern)
- | FieldOutsideRecord
- | ModuleArityMismatch ModuleName Telescope [NamedArg Expr]
- | GeneralizeCyclicDependency
- | GeneralizeUnsolvedMeta
- | SplitError SplitError
- | ImpossibleConstructor QName NegativeUnification
- | TooManyPolarities QName Int
- | LocalVsImportedModuleClash ModuleName
- | SolvedButOpenHoles
- | CyclicModuleDependency [TopLevelModuleName]
- | FileNotFound TopLevelModuleName [AbsolutePath]
- | OverlappingProjects AbsolutePath TopLevelModuleName TopLevelModuleName
- | AmbiguousTopLevelModuleName TopLevelModuleName [AbsolutePath]
- | ModuleNameUnexpected TopLevelModuleName TopLevelModuleName
- | ModuleNameDoesntMatchFileName TopLevelModuleName [AbsolutePath]
- | ClashingFileNamesFor ModuleName [AbsolutePath]
- | ModuleDefinedInOtherFile TopLevelModuleName AbsolutePath AbsolutePath
- | BothWithAndRHS
- | AbstractConstructorNotInScope QName
- | NotInScope [QName]
- | NoSuchModule QName
- | AmbiguousName QName (NonEmpty QName)
- | AmbiguousModule QName (NonEmpty ModuleName)
- | ClashingDefinition QName QName
- | ClashingModule ModuleName ModuleName
- | ClashingImport Name QName
- | ClashingModuleImport Name ModuleName
- | PatternShadowsConstructor Name QName
- | DuplicateImports QName [ImportedName]
- | InvalidPattern Pattern
- | RepeatedVariablesInPattern [Name]
- | GeneralizeNotSupportedHere QName
- | MultipleFixityDecls [(Name, [Fixity'])]
- | MultiplePolarityPragmas [Name]
- | NotAModuleExpr Expr
- | NotAnExpression Expr
- | NotAValidLetBinding NiceDeclaration
- | NotValidBeforeField NiceDeclaration
- | NothingAppliedToHiddenArg Expr
- | NothingAppliedToInstanceArg Expr
- | BadArgumentsToPatternSynonym AmbiguousQName
- | TooFewArgumentsToPatternSynonym AmbiguousQName
- | CannotResolveAmbiguousPatternSynonym (NonEmpty (QName, PatternSynDefn))
- | UnusedVariableInPatternSynonym
- | NoParseForApplication [Expr]
- | AmbiguousParseForApplication [Expr] [Expr]
- | NoParseForLHS LHSOrPatSyn Pattern
- | AmbiguousParseForLHS LHSOrPatSyn Pattern [Pattern]
- | OperatorInformation [NotationSection] TypeError
- | InstanceNoCandidate Type [(Term, TCErr)]
- | UnquoteFailed UnquoteError
- | DeBruijnIndexOutOfScope Nat Telescope [Name]
- | NeedOptionCopatterns
- | NeedOptionRewriting
- | NeedOptionProp
- | NonFatalErrors [TCWarning]
- | InstanceSearchDepthExhausted Term Type Int
- | TriedToCopyConstrainedPrim QName
- data LHSOrPatSyn
- data TCErr
- class (Functor m, Applicative m, Monad m) => HasOptions m where
- sizedTypesOption :: HasOptions m => m Bool
- guardednessOption :: HasOptions m => m Bool
- withoutKOption :: HasOptions m => m Bool
- getIncludeDirs :: HasOptions m => m [AbsolutePath]
- enableCaching :: HasOptions m => m Bool
- data ReduceEnv = ReduceEnv {}
- mapRedEnv :: (TCEnv -> TCEnv) -> ReduceEnv -> ReduceEnv
- mapRedSt :: (TCState -> TCState) -> ReduceEnv -> ReduceEnv
- mapRedEnvSt :: (TCEnv -> TCEnv) -> (TCState -> TCState) -> ReduceEnv -> ReduceEnv
- reduceEnv :: Lens' TCEnv ReduceEnv
- reduceSt :: Lens' TCState ReduceEnv
- newtype ReduceM a = ReduceM {}
- onReduceEnv :: (ReduceEnv -> ReduceEnv) -> ReduceM a -> ReduceM a
- fmapReduce :: (a -> b) -> ReduceM a -> ReduceM b
- apReduce :: ReduceM (a -> b) -> ReduceM a -> ReduceM b
- bindReduce :: ReduceM a -> (a -> ReduceM b) -> ReduceM b
- runReduceM :: ReduceM a -> TCM a
- runReduceF :: (a -> ReduceM b) -> TCM (a -> b)
- useR :: ReadTCState m => Lens' a TCState -> m a
- askR :: ReduceM ReduceEnv
- localR :: (ReduceEnv -> ReduceEnv) -> ReduceM a -> ReduceM a
- class (Applicative m, MonadTCEnv m, ReadTCState m, HasOptions m) => MonadReduce m where
- liftReduce :: ReduceM a -> m a
- class Monad m => MonadTCEnv m where
- asksTC :: MonadTCEnv m => (TCEnv -> a) -> m a
- viewTC :: MonadTCEnv m => Lens' a TCEnv -> m a
- locallyTC :: MonadTCEnv m => Lens' a TCEnv -> (a -> a) -> m b -> m b
- class Monad m => MonadTCState m where
- getsTC :: ReadTCState m => (TCState -> a) -> m a
- modifyTC' :: MonadTCState m => (TCState -> TCState) -> m ()
- useTC :: ReadTCState m => Lens' a TCState -> m a
- setTCLens :: MonadTCState m => Lens' a TCState -> a -> m ()
- modifyTCLens :: MonadTCState m => Lens' a TCState -> (a -> a) -> m ()
- modifyTCLensM :: MonadTCState m => Lens' a TCState -> (a -> m a) -> m ()
- stateTCLens :: MonadTCState m => Lens' a TCState -> (a -> (r, a)) -> m r
- stateTCLensM :: MonadTCState m => Lens' a TCState -> (a -> m (r, a)) -> m r
- newtype TCMT m a = TCM {}
- type TCM = TCMT IO
- mapTCMT :: (forall a. m a -> n a) -> TCMT m a -> TCMT n a
- pureTCM :: MonadIO m => (TCState -> TCEnv -> a) -> TCMT m a
- returnTCMT :: MonadIO m => a -> TCMT m a
- bindTCMT :: MonadIO m => TCMT m a -> (a -> TCMT m b) -> TCMT m b
- thenTCMT :: MonadIO m => TCMT m a -> TCMT m b -> TCMT m b
- fmapTCMT :: MonadIO m => (a -> b) -> TCMT m a -> TCMT m b
- apTCMT :: MonadIO m => TCMT m (a -> b) -> TCMT m a -> TCMT m b
- type IM = TCMT (InputT IO)
- runIM :: IM a -> TCM a
- catchError_ :: TCM a -> (TCErr -> TCM a) -> TCM a
- finally_ :: TCM a -> TCM b -> TCM a
- class (Applicative tcm, MonadIO tcm, MonadTCEnv tcm, MonadTCState tcm, HasOptions tcm) => MonadTCM tcm where
- patternViolation :: MonadError TCErr m => m a
- internalError :: MonadTCM tcm => String -> tcm a
- genericError :: (MonadTCEnv m, ReadTCState m, MonadError TCErr m) => String -> m a
- genericDocError :: (MonadTCEnv m, ReadTCState m, MonadError TCErr m) => Doc -> m a
- typeError :: (MonadTCEnv m, ReadTCState m, MonadError TCErr m) => TypeError -> m a
- typeError_ :: (MonadTCEnv m, ReadTCState m) => TypeError -> m TCErr
- 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
- runSafeTCM :: TCM a -> TCState -> IO (a, TCState)
- forkTCM :: TCM a -> TCM ()
- patternInTeleName :: String
- extendedLambdaName :: String
- isExtendedLambdaName :: QName -> Bool
- absurdLambdaName :: String
- isAbsurdLambdaName :: QName -> Bool
- generalizedFieldName :: String
- getGeneralizedFieldName :: QName -> Maybe String
Type checking state
TCSt | |
|
Instances
class Monad m => ReadTCState m where Source #
getTCState :: m TCState Source #
locallyTCState :: Lens' a TCState -> (a -> a) -> m b -> m b Source #
withTCState :: (TCState -> TCState) -> m a -> m a Source #
Instances
data PreScopeState Source #
PreScopeState | |
|
type DisambiguatedNames = IntMap QName Source #
data PostScopeState Source #
PostScopeState | |
|
data MutualBlock Source #
A mutual block of names in the signature.
MutualBlock | |
|
Instances
Eq MutualBlock Source # | |
Defined in Agda.TypeChecking.Monad.Base (==) :: MutualBlock -> MutualBlock -> Bool # (/=) :: MutualBlock -> MutualBlock -> Bool # | |
Show MutualBlock Source # | |
Defined in Agda.TypeChecking.Monad.Base showsPrec :: Int -> MutualBlock -> ShowS # show :: MutualBlock -> String # showList :: [MutualBlock] -> ShowS # | |
Null MutualBlock Source # | |
Defined in Agda.TypeChecking.Monad.Base empty :: MutualBlock Source # null :: MutualBlock -> Bool Source # |
data PersistentTCState Source #
A part of the state which is not reverted when an error is thrown or the state is reset.
PersistentTCSt | |
|
Instances
data LoadedFileCache 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.
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 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.
EnterSection !ModuleInfo !ModuleName !Telescope | |
LeaveSection !ModuleName | |
Decl !Declaration | Never a Section or ScopeDecl |
Pragmas !PragmaOptions |
initPersistentState :: PersistentTCState Source #
Empty persistent state.
initPreScopeState :: PreScopeState Source #
Empty state of type checker.
st-prefixed lenses
getUserWarnings :: ReadTCState m => m (Map QName String) Source #
getPartialDefs :: ReadTCState m => m (Set QName) Source #
Fresh things
class Monad m => MonadFresh i m where Source #
Instances
HasFresh i => MonadFresh i TCM Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
Monad m => MonadFresh Int (PureConversionT m) Source # | |
Defined in Agda.TypeChecking.Conversion.Pure fresh :: PureConversionT m Int Source # | |
Monad m => MonadFresh NameId (PureConversionT m) Source # | |
Defined in Agda.TypeChecking.Conversion.Pure fresh :: PureConversionT m NameId Source # | |
Monad m => MonadFresh ProblemId (PureConversionT m) Source # | |
Defined in Agda.TypeChecking.Conversion.Pure fresh :: PureConversionT m ProblemId Source # | |
MonadFresh i m => MonadFresh i (StateT s m) Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
MonadFresh i m => MonadFresh i (ReaderT r m) Source # | |
Defined in Agda.TypeChecking.Monad.Base |
Instances
newtype CheckpointId Source #
Instances
freshNoName :: MonadFresh NameId m => Range -> m Name Source #
freshNoName_ :: MonadFresh NameId m => m Name Source #
freshRecordName :: MonadFresh NameId m => m Name Source #
class FreshName a where Source #
Create a fresh name from a
.
freshName_ :: MonadFresh NameId m => a -> m Name Source #
Instances
FreshName () Source # | |
Defined in Agda.TypeChecking.Monad.Base freshName_ :: MonadFresh NameId m => () -> m Name Source # | |
FreshName String Source # | |
Defined in Agda.TypeChecking.Monad.Base freshName_ :: MonadFresh NameId m => String -> m Name Source # | |
FreshName Range Source # | |
Defined in Agda.TypeChecking.Monad.Base freshName_ :: MonadFresh NameId m => Range -> m Name Source # | |
FreshName (Range, String) Source # | |
Defined in Agda.TypeChecking.Monad.Base freshName_ :: MonadFresh NameId m => (Range, String) -> m Name Source # |
Managing file names
type ModuleToSource = Map TopLevelModuleName AbsolutePath Source #
Maps top-level module names to the corresponding source file names.
type SourceToModule = Map AbsolutePath TopLevelModuleName Source #
Maps source file names to the corresponding top-level module names.
sourceToModule :: TCM SourceToModule Source #
Creates a SourceToModule
map based on stModuleToSource
.
O(n log n).
For a single reverse lookup in stModuleToSource
,
rather use lookupModuleFromSourse
.
lookupModuleFromSource :: ReadTCState m => AbsolutePath -> m (Maybe TopLevelModuleName) Source #
Lookup an AbsolutePath
in sourceToModule
.
O(n).
Associating concrete names to an abstract name
class Monad m => MonadStConcreteNames m 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).
runStConcreteNames :: StateT ConcreteNames m a -> m a Source #
useConcreteNames :: m ConcreteNames Source #
modifyConcreteNames :: (ConcreteNames -> ConcreteNames) -> m () Source #
Instances
Interface
data ModuleInfo Source #
ModuleInfo | |
|
data ForeignCode Source #
Instances
Show ForeignCode Source # | |
Defined in Agda.TypeChecking.Monad.Base showsPrec :: Int -> ForeignCode -> ShowS # show :: ForeignCode -> String # showList :: [ForeignCode] -> ShowS # | |
EmbPrj ForeignCode Source # | |
Interface | |
|
iFullHash :: Interface -> Hash Source #
Combines the source hash and the (full) hashes of the imported modules.
Closure
Closure | |
|
Instances
class LensClosure a b | b -> a where Source #
lensClosure :: Lens' (Closure a) b Source #
Instances
LensClosure Range MetaInfo Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
LensClosure Range MetaVariable Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
LensClosure a (Closure a) Source # | |
Defined in Agda.TypeChecking.Monad.Base |
buildClosure :: (MonadTCEnv m, ReadTCState m) => a -> m (Closure a) Source #
Constraints
type Constraints = [ProblemConstraint] Source #
data ProblemConstraint Source #
Instances
data Constraint Source #
ValueCmp Comparison CompareAs Term Term | |
ValueCmpOnFace Comparison Term Type Term Term | |
ElimCmp [Polarity] [IsForced] Type Term [Elim] [Elim] | |
TelCmp Type Type Comparison Telescope Telescope | the two types are for the error message only |
SortCmp Comparison Sort Sort | |
LevelCmp Comparison Level Level | |
HasBiggerSort Sort | |
HasPTSRule (Dom Type) (Abs Sort) | |
CheckMetaInst MetaId | |
UnBlock MetaId | |
Guarded Constraint ProblemId | |
IsEmpty Range Type | The range is the one of the absurd pattern. |
CheckSizeLtSat Term | Check that the |
FindInstance MetaId (Maybe MetaId) (Maybe [Candidate]) | the first argument is the instance argument, the second one is the meta on which the constraint may be blocked on and the third one is the list of candidates (or Nothing if we haven’t determined the list of candidates yet) |
CheckFunDef Delayed DefInfo QName [Clause] | |
UnquoteTactic (Maybe MetaId) Term Term Type | First argument is computation and the others are hole and goal type |
Instances
data Comparison Source #
Instances
data CompareDirection Source #
An extension of Comparison
to >=
.
Instances
Eq CompareDirection Source # | |
Defined in Agda.TypeChecking.Monad.Base (==) :: CompareDirection -> CompareDirection -> Bool # (/=) :: CompareDirection -> CompareDirection -> Bool # | |
Show CompareDirection Source # | |
Defined in Agda.TypeChecking.Monad.Base showsPrec :: Int -> CompareDirection -> ShowS # show :: CompareDirection -> String # showList :: [CompareDirection] -> ShowS # | |
Pretty CompareDirection Source # | |
Defined in Agda.TypeChecking.Monad.Base pretty :: CompareDirection -> Doc Source # prettyPrec :: Int -> CompareDirection -> Doc Source # prettyList :: [CompareDirection] -> Doc Source # |
fromCmp :: Comparison -> CompareDirection Source #
Embed Comparison
into CompareDirection
.
flipCmp :: CompareDirection -> CompareDirection Source #
Flip the direction of comparison.
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
We can either compare two terms at a given type, or compare two types without knowing (or caring about) their sorts.
AsTermsOf Type |
|
AsSizes | Replaces |
AsTypes |
Instances
Open things
A thing tagged with the context it came from.
Instances
Functor Open Source # | |
Foldable Open Source # | |
Defined in Agda.TypeChecking.Monad.Base 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 # | |
Decoration Open Source # | |
Data a => Data (Open a) Source # | |
Defined in Agda.TypeChecking.Monad.Base gfoldl :: (forall d b. Data d => c (d -> b) -> d -> c b) -> (forall g. g -> c g) -> Open a -> c (Open a) # gunfold :: (forall b r. Data b => c (b -> r) -> c r) -> (forall r. r -> c r) -> Constr -> c (Open a) # toConstr :: Open a -> Constr # dataTypeOf :: Open a -> DataType # dataCast1 :: Typeable t => (forall d. Data d => c (t d)) -> Maybe (c (Open a)) # dataCast2 :: Typeable t => (forall d e. (Data d, Data e) => c (t d e)) -> Maybe (c (Open a)) # gmapT :: (forall b. Data b => b -> b) -> Open a -> Open a # gmapQl :: (r -> r' -> r) -> r -> (forall d. Data d => d -> r') -> Open a -> r # gmapQr :: forall r r'. (r' -> r -> r) -> r -> (forall d. Data d => d -> r') -> Open a -> r # gmapQ :: (forall d. Data d => d -> u) -> Open a -> [u] # gmapQi :: Int -> (forall d. Data d => d -> u) -> Open a -> u # gmapM :: Monad m => (forall d. Data d => d -> m d) -> Open a -> m (Open a) # gmapMp :: MonadPlus m => (forall d. Data d => d -> m d) -> Open a -> m (Open a) # gmapMo :: MonadPlus m => (forall d. Data d => d -> m d) -> Open a -> m (Open a) # | |
Show a => Show (Open a) Source # | |
KillRange a => KillRange (Open a) Source # | |
Defined in Agda.TypeChecking.Monad.Base killRange :: KillRangeT (Open a) Source # | |
EmbPrj a => EmbPrj (Open a) Source # | |
NamesIn a => NamesIn (Open a) Source # | |
InstantiateFull t => InstantiateFull (Open t) Source # | |
Defined in Agda.TypeChecking.Reduce |
Judgements
Parametrized since it is used without MetaId when creating a new meta.
HasType | |
| |
IsSort | |
Generalizable variables
data DoGeneralize Source #
Instances
data GeneralizedValue Source #
The value of a generalizable variable. This is created to be a generalizable meta before checking the type to be generalized.
Instances
Data GeneralizedValue Source # | |
Defined in Agda.TypeChecking.Monad.Base gfoldl :: (forall d b. Data d => c (d -> b) -> d -> c b) -> (forall g. g -> c g) -> GeneralizedValue -> c GeneralizedValue # gunfold :: (forall b r. Data b => c (b -> r) -> c r) -> (forall r. r -> c r) -> Constr -> c GeneralizedValue # toConstr :: GeneralizedValue -> Constr # dataTypeOf :: GeneralizedValue -> DataType # dataCast1 :: Typeable t => (forall d. Data d => c (t d)) -> Maybe (c GeneralizedValue) # dataCast2 :: Typeable t => (forall d e. (Data d, Data e) => c (t d e)) -> Maybe (c GeneralizedValue) # gmapT :: (forall b. Data b => b -> b) -> GeneralizedValue -> GeneralizedValue # gmapQl :: (r -> r' -> r) -> r -> (forall d. Data d => d -> r') -> GeneralizedValue -> r # gmapQr :: forall r r'. (r' -> r -> r) -> r -> (forall d. Data d => d -> r') -> GeneralizedValue -> r # gmapQ :: (forall d. Data d => d -> u) -> GeneralizedValue -> [u] # gmapQi :: Int -> (forall d. Data d => d -> u) -> GeneralizedValue -> u # gmapM :: Monad m => (forall d. Data d => d -> m d) -> GeneralizedValue -> m GeneralizedValue # gmapMp :: MonadPlus m => (forall d. Data d => d -> m d) -> GeneralizedValue -> m GeneralizedValue # gmapMo :: MonadPlus m => (forall d. Data d => d -> m d) -> GeneralizedValue -> m GeneralizedValue # | |
Show GeneralizedValue Source # | |
Defined in Agda.TypeChecking.Monad.Base showsPrec :: Int -> GeneralizedValue -> ShowS # show :: GeneralizedValue -> String # showList :: [GeneralizedValue] -> ShowS # |
Meta variables
data MetaVariable Source #
MetaVar | |
|
Instances
SetRange MetaVariable Source # | |
Defined in Agda.TypeChecking.Monad.Base setRange :: Range -> MetaVariable -> MetaVariable Source # | |
HasRange MetaVariable Source # | |
Defined in Agda.TypeChecking.Monad.Base getRange :: MetaVariable -> Range Source # | |
LensClosure Range MetaVariable Source # | |
Defined in Agda.TypeChecking.Monad.Base |
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).
Frozen | Do not instantiate. |
Instantiable |
data MetaInstantiation Source #
InstV [Arg String] Term | solved by term (abstracted over some free variables) |
Open | unsolved |
OpenInstance | open, to be instantiated by instance search |
BlockedConst Term | solution blocked by unsolved constraints |
PostponedTypeCheckingProblem (Closure TypeCheckingProblem) (TCM Bool) |
Instances
Show MetaInstantiation Source # | |
Defined in Agda.TypeChecking.Monad.Base showsPrec :: Int -> MetaInstantiation -> ShowS # show :: MetaInstantiation -> String # showList :: [MetaInstantiation] -> 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.
data TypeCheckingProblem Source #
CheckExpr Comparison Expr Type | |
CheckArgs ExpandHidden Range [NamedArg Expr] Type Type ([Maybe Range] -> Elims -> Type -> CheckedTarget -> TCM Term) | |
CheckProjAppToKnownPrincipalArg Comparison Expr ProjOrigin (NonEmpty QName) Args Type Int Term Type | |
CheckLambda Comparison (Arg ([WithHiding Name], Maybe Type)) Expr Type |
|
DoQuoteTerm Comparison Term Type | Quote the given term and check type against |
Instances
PrettyTCM TypeCheckingProblem Source # | |
Defined in Agda.TypeChecking.Pretty prettyTCM :: MonadPretty m => TypeCheckingProblem -> m Doc Source # |
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.
Instances
Eq MetaPriority Source # | |
Defined in Agda.TypeChecking.Monad.Base (==) :: MetaPriority -> MetaPriority -> Bool # (/=) :: MetaPriority -> MetaPriority -> Bool # | |
Ord MetaPriority Source # | |
Defined in Agda.TypeChecking.Monad.Base compare :: MetaPriority -> MetaPriority -> Ordering # (<) :: MetaPriority -> MetaPriority -> Bool # (<=) :: MetaPriority -> MetaPriority -> Bool # (>) :: MetaPriority -> MetaPriority -> Bool # (>=) :: MetaPriority -> MetaPriority -> Bool # max :: MetaPriority -> MetaPriority -> MetaPriority # min :: MetaPriority -> MetaPriority -> MetaPriority # | |
Show MetaPriority Source # | |
Defined in Agda.TypeChecking.Monad.Base showsPrec :: Int -> MetaPriority -> ShowS # show :: MetaPriority -> String # showList :: [MetaPriority] -> ShowS # |
data RunMetaOccursCheck Source #
Instances
Eq RunMetaOccursCheck Source # | |
Defined in Agda.TypeChecking.Monad.Base (==) :: RunMetaOccursCheck -> RunMetaOccursCheck -> Bool # (/=) :: RunMetaOccursCheck -> RunMetaOccursCheck -> Bool # | |
Ord RunMetaOccursCheck Source # | |
Defined in Agda.TypeChecking.Monad.Base compare :: RunMetaOccursCheck -> RunMetaOccursCheck -> Ordering # (<) :: RunMetaOccursCheck -> RunMetaOccursCheck -> Bool # (<=) :: RunMetaOccursCheck -> RunMetaOccursCheck -> Bool # (>) :: RunMetaOccursCheck -> RunMetaOccursCheck -> Bool # (>=) :: RunMetaOccursCheck -> RunMetaOccursCheck -> Bool # max :: RunMetaOccursCheck -> RunMetaOccursCheck -> RunMetaOccursCheck # min :: RunMetaOccursCheck -> RunMetaOccursCheck -> RunMetaOccursCheck # | |
Show RunMetaOccursCheck Source # | |
Defined in Agda.TypeChecking.Monad.Base showsPrec :: Int -> RunMetaOccursCheck -> ShowS # show :: RunMetaOccursCheck -> String # showList :: [RunMetaOccursCheck] -> ShowS # |
MetaInfo
is cloned from one meta to the next during pruning.
MetaInfo | |
|
Instances
type MetaNameSuggestion = String Source #
Name suggestion for meta variable. Empty string means no suggestion.
For printing, we couple a meta with its name suggestion.
Instances
type MetaStore = IntMap MetaVariable Source #
getMetaInfo :: MetaVariable -> Closure Range Source #
getMetaScope :: MetaVariable -> ScopeInfo Source #
getMetaEnv :: MetaVariable -> TCEnv Source #
getMetaSig :: MetaVariable -> Signature Source #
Interaction meta variables
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.
Instances
Eq InteractionPoint Source # | |
Defined in Agda.TypeChecking.Monad.Base (==) :: InteractionPoint -> InteractionPoint -> Bool # (/=) :: InteractionPoint -> InteractionPoint -> Bool # |
type InteractionPoints = Map 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)
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).
Instances
Eq Overapplied Source # | |
Defined in Agda.TypeChecking.Monad.Base (==) :: Overapplied -> Overapplied -> Bool # (/=) :: Overapplied -> Overapplied -> Bool # | |
Data Overapplied Source # | |
Defined in Agda.TypeChecking.Monad.Base gfoldl :: (forall d b. Data d => c (d -> b) -> d -> c b) -> (forall g. g -> c g) -> Overapplied -> c Overapplied # gunfold :: (forall b r. Data b => c (b -> r) -> c r) -> (forall r. r -> c r) -> Constr -> c Overapplied # toConstr :: Overapplied -> Constr # dataTypeOf :: Overapplied -> DataType # dataCast1 :: Typeable t => (forall d. Data d => c (t d)) -> Maybe (c Overapplied) # dataCast2 :: Typeable t => (forall d e. (Data d, Data e) => c (t d e)) -> Maybe (c Overapplied) # gmapT :: (forall b. Data b => b -> b) -> Overapplied -> Overapplied # gmapQl :: (r -> r' -> r) -> r -> (forall d. Data d => d -> r') -> Overapplied -> r # gmapQr :: forall r r'. (r' -> r -> r) -> r -> (forall d. Data d => d -> r') -> Overapplied -> r # gmapQ :: (forall d. Data d => d -> u) -> Overapplied -> [u] # gmapQi :: Int -> (forall d. Data d => d -> u) -> Overapplied -> u # gmapM :: Monad m => (forall d. Data d => d -> m d) -> Overapplied -> m Overapplied # gmapMp :: MonadPlus m => (forall d. Data d => d -> m d) -> Overapplied -> m Overapplied # gmapMo :: MonadPlus m => (forall d. Data d => d -> m d) -> Overapplied -> m Overapplied # | |
Show Overapplied Source # | |
Defined in Agda.TypeChecking.Monad.Base showsPrec :: Int -> Overapplied -> ShowS # show :: Overapplied -> String # showList :: [Overapplied] -> ShowS # |
data IPBoundary' t Source #
Datatype representing a single boundary condition: x_0 = u_0, ... ,x_n = u_n ⊢ t = ?n es
IPBoundary | |
|
Instances
type IPBoundary = IPBoundary' Term Source #
Which clause is an interaction point located in?
IPClause | |
| |
IPNoClause | The interaction point is not in the rhs of a clause. |
Instances
Eq IPClause Source # | |
Data IPClause Source # | |
Defined in Agda.TypeChecking.Monad.Base gfoldl :: (forall d b. Data d => c (d -> b) -> d -> c b) -> (forall g. g -> c g) -> IPClause -> c IPClause # gunfold :: (forall b r. Data b => c (b -> r) -> c r) -> (forall r. r -> c r) -> Constr -> c IPClause # toConstr :: IPClause -> Constr # dataTypeOf :: IPClause -> DataType # dataCast1 :: Typeable t => (forall d. Data d => c (t d)) -> Maybe (c IPClause) # dataCast2 :: Typeable t => (forall d e. (Data d, Data e) => c (t d e)) -> Maybe (c IPClause) # gmapT :: (forall b. Data b => b -> b) -> IPClause -> IPClause # gmapQl :: (r -> r' -> r) -> r -> (forall d. Data d => d -> r') -> IPClause -> r # gmapQr :: forall r r'. (r' -> r -> r) -> r -> (forall d. Data d => d -> r') -> IPClause -> r # gmapQ :: (forall d. Data d => d -> u) -> IPClause -> [u] # gmapQi :: Int -> (forall d. Data d => d -> u) -> IPClause -> u # gmapM :: Monad m => (forall d. Data d => d -> m d) -> IPClause -> m IPClause # gmapMp :: MonadPlus m => (forall d. Data d => d -> m d) -> IPClause -> m IPClause # gmapMo :: MonadPlus m => (forall d. Data d => d -> m d) -> IPClause -> m IPClause # |
Signature
Sig | |
|
Instances
Data Signature Source # | |
Defined in Agda.TypeChecking.Monad.Base gfoldl :: (forall d b. Data d => c (d -> b) -> d -> c b) -> (forall g. g -> c g) -> Signature -> c Signature # gunfold :: (forall b r. Data b => c (b -> r) -> c r) -> (forall r. r -> c r) -> Constr -> c Signature # toConstr :: Signature -> Constr # dataTypeOf :: Signature -> DataType # dataCast1 :: Typeable t => (forall d. Data d => c (t d)) -> Maybe (c Signature) # dataCast2 :: Typeable t => (forall d e. (Data d, Data e) => c (t d e)) -> Maybe (c Signature) # gmapT :: (forall b. Data b => b -> b) -> Signature -> Signature # gmapQl :: (r -> r' -> r) -> r -> (forall d. Data d => d -> r') -> Signature -> r # gmapQr :: forall r r'. (r' -> r -> r) -> r -> (forall d. Data d => d -> r') -> Signature -> r # gmapQ :: (forall d. Data d => d -> u) -> Signature -> [u] # gmapQi :: Int -> (forall d. Data d => d -> u) -> Signature -> u # gmapM :: Monad m => (forall d. Data d => d -> m d) -> Signature -> m Signature # gmapMp :: MonadPlus m => (forall d. Data d => d -> m d) -> Signature -> m Signature # gmapMo :: MonadPlus m => (forall d. Data d => d -> m d) -> Signature -> m Signature # | |
Show Signature Source # | |
KillRange Signature Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
EmbPrj Signature Source # | |
InstantiateFull Signature Source # | |
Defined in Agda.TypeChecking.Reduce |
type Definitions = HashMap QName Definition Source #
type RewriteRuleMap = HashMap QName RewriteRules Source #
type DisplayForms = HashMap QName [LocalDisplayForm] Source #
Instances
Eq Section Source # | |
Data Section Source # | |
Defined in Agda.TypeChecking.Monad.Base gfoldl :: (forall d b. Data d => c (d -> b) -> d -> c b) -> (forall g. g -> c g) -> Section -> c Section # gunfold :: (forall b r. Data b => c (b -> r) -> c r) -> (forall r. r -> c r) -> Constr -> c Section # toConstr :: Section -> Constr # dataTypeOf :: Section -> DataType # dataCast1 :: Typeable t => (forall d. Data d => c (t d)) -> Maybe (c Section) # dataCast2 :: Typeable t => (forall d e. (Data d, Data e) => c (t d e)) -> Maybe (c Section) # gmapT :: (forall b. Data b => b -> b) -> Section -> Section # gmapQl :: (r -> r' -> r) -> r -> (forall d. Data d => d -> r') -> Section -> r # gmapQr :: forall r r'. (r' -> r -> r) -> r -> (forall d. Data d => d -> r') -> Section -> r # gmapQ :: (forall d. Data d => d -> u) -> Section -> [u] # gmapQi :: Int -> (forall d. Data d => d -> u) -> Section -> u # gmapM :: Monad m => (forall d. Data d => d -> m d) -> Section -> m Section # gmapMp :: MonadPlus m => (forall d. Data d => d -> m d) -> Section -> m Section # gmapMo :: MonadPlus m => (forall d. Data d => d -> m d) -> Section -> m Section # | |
Show Section Source # | |
Pretty Section Source # | |
KillRange Section Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
KillRange Sections Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
EmbPrj Section Source # | |
InstantiateFull Section Source # | |
Defined in Agda.TypeChecking.Reduce |
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 var 0
is interpreted
as a hole. Each occurrence of var 0
is a new hole (pattern var).
For each *occurrence* of var0
the rhs dt
has a free variable.
These are instantiated when matching a display form against a
term q vs
succeeds.
Display | |
|
Instances
type LocalDisplayForm = Open DisplayForm Source #
data DisplayTerm Source #
DWithApp DisplayTerm [DisplayTerm] Elims |
|
DCon ConHead ConInfo [Arg DisplayTerm] |
|
DDef QName [Elim' DisplayTerm] |
|
DDot Term |
|
DTerm Term |
|
Instances
defaultDisplayForm :: QName -> [LocalDisplayForm] Source #
By default, we have no display form.
Non-linear (non-constructor) first-order pattern.
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
Instances
Instances
type RewriteRules = [RewriteRule] Source #
data RewriteRule Source #
Rewrite rules can be added independently from function clauses.
Instances
data Definition Source #
Defn | |
|
Instances
data NumGeneralizableArgs Source #
NoGeneralizableArgs | |
SomeGeneralizableArgs Int | When lambda-lifting new args are generalizable if
|
Instances
defaultDefn :: ArgInfo -> QName -> Type -> Defn -> Definition Source #
Create a definition with sensible defaults.
Polarity for equality and subtype checking.
Covariant | monotone |
Contravariant | antitone |
Invariant | no information (mixed variance) |
Nonvariant | constant |
Instances
Eq Polarity Source # | |
Data Polarity Source # | |
Defined in Agda.TypeChecking.Monad.Base gfoldl :: (forall d b. Data d => c (d -> b) -> d -> c b) -> (forall g. g -> c g) -> Polarity -> c Polarity # gunfold :: (forall b r. Data b => c (b -> r) -> c r) -> (forall r. r -> c r) -> Constr -> c Polarity # toConstr :: Polarity -> Constr # dataTypeOf :: Polarity -> DataType # dataCast1 :: Typeable t => (forall d. Data d => c (t d)) -> Maybe (c Polarity) # dataCast2 :: Typeable t => (forall d e. (Data d, Data e) => c (t d e)) -> Maybe (c Polarity) # gmapT :: (forall b. Data b => b -> b) -> Polarity -> Polarity # gmapQl :: (r -> r' -> r) -> r -> (forall d. Data d => d -> r') -> Polarity -> r # gmapQr :: forall r r'. (r' -> r -> r) -> r -> (forall d. Data d => d -> r') -> Polarity -> r # gmapQ :: (forall d. Data d => d -> u) -> Polarity -> [u] # gmapQi :: Int -> (forall d. Data d => d -> u) -> Polarity -> u # gmapM :: Monad m => (forall d. Data d => d -> m d) -> Polarity -> m Polarity # gmapMp :: MonadPlus m => (forall d. Data d => d -> m d) -> Polarity -> m Polarity # gmapMo :: MonadPlus m => (forall d. Data d => d -> m d) -> Polarity -> m Polarity # | |
Show Polarity Source # | |
Pretty Polarity Source # | |
KillRange Polarity Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
EmbPrj Polarity Source # | |
PrettyTCM Polarity Source # | |
Defined in Agda.TypeChecking.Pretty | |
Abstract [Polarity] Source # | |
Apply [Polarity] Source # | |
Information about whether an argument is forced by the type of a function.
Instances
Eq IsForced Source # | |
Data IsForced Source # | |
Defined in Agda.TypeChecking.Monad.Base gfoldl :: (forall d b. Data d => c (d -> b) -> d -> c b) -> (forall g. g -> c g) -> IsForced -> c IsForced # gunfold :: (forall b r. Data b => c (b -> r) -> c r) -> (forall r. r -> c r) -> Constr -> c IsForced # toConstr :: IsForced -> Constr # dataTypeOf :: IsForced -> DataType # dataCast1 :: Typeable t => (forall d. Data d => c (t d)) -> Maybe (c IsForced) # dataCast2 :: Typeable t => (forall d e. (Data d, Data e) => c (t d e)) -> Maybe (c IsForced) # gmapT :: (forall b. Data b => b -> b) -> IsForced -> IsForced # gmapQl :: (r -> r' -> r) -> r -> (forall d. Data d => d -> r') -> IsForced -> r # gmapQr :: forall r r'. (r' -> r -> r) -> r -> (forall d. Data d => d -> r') -> IsForced -> r # gmapQ :: (forall d. Data d => d -> u) -> IsForced -> [u] # gmapQi :: Int -> (forall d. Data d => d -> u) -> IsForced -> u # gmapM :: Monad m => (forall d. Data d => d -> m d) -> IsForced -> m IsForced # gmapMp :: MonadPlus m => (forall d. Data d => d -> m d) -> IsForced -> m IsForced # gmapMo :: MonadPlus m => (forall d. Data d => d -> m d) -> IsForced -> m IsForced # | |
Show IsForced Source # | |
KillRange IsForced Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
EmbPrj IsForced Source # | |
PrettyTCM IsForced Source # | |
Defined in Agda.TypeChecking.Pretty | |
ChooseFlex IsForced Source # | |
Defined in Agda.TypeChecking.Rules.LHS.Problem chooseFlex :: IsForced -> IsForced -> FlexChoice Source # |
data CompilerPragma Source #
The backends are responsible for parsing their own pragmas.
Instances
type BackendName = String Source #
An alternative representation of partial elements in a telescope: Γ ⊢ λ Δ. [φ₁ u₁, ... , φₙ uₙ] : Δ → PartialP (∨_ᵢ φᵢ) T see cubicaltt paper (however we do not store the type T).
System | |
|
Instances
Data System Source # | |
Defined in Agda.TypeChecking.Monad.Base gfoldl :: (forall d b. Data d => c (d -> b) -> d -> c b) -> (forall g. g -> c g) -> System -> c System # gunfold :: (forall b r. Data b => c (b -> r) -> c r) -> (forall r. r -> c r) -> Constr -> c System # toConstr :: System -> Constr # dataTypeOf :: System -> DataType # dataCast1 :: Typeable t => (forall d. Data d => c (t d)) -> Maybe (c System) # dataCast2 :: Typeable t => (forall d e. (Data d, Data e) => c (t d e)) -> Maybe (c System) # gmapT :: (forall b. Data b => b -> b) -> System -> System # gmapQl :: (r -> r' -> r) -> r -> (forall d. Data d => d -> r') -> System -> r # gmapQr :: forall r r'. (r' -> r -> r) -> r -> (forall d. Data d => d -> r') -> System -> r # gmapQ :: (forall d. Data d => d -> u) -> System -> [u] # gmapQi :: Int -> (forall d. Data d => d -> u) -> System -> u # gmapM :: Monad m => (forall d. Data d => d -> m d) -> System -> m System # gmapMp :: MonadPlus m => (forall d. Data d => d -> m d) -> System -> m System # gmapMo :: MonadPlus m => (forall d. Data d => d -> m d) -> System -> m System # | |
Show System Source # | |
KillRange System Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
Abstract System Source # | |
Apply System Source # | |
EmbPrj System Source # | |
InstantiateFull System Source # | |
Defined in Agda.TypeChecking.Reduce | |
Reify (QNamed System) [Clause] Source # | |
data ExtLamInfo Source #
Additional information for extended lambdas.
ExtLamInfo | |
|
Instances
modifySystem :: (System -> System) -> ExtLamInfo -> ExtLamInfo Source #
data Projection Source #
Additional information for projection Function
s.
Projection | |
|
Instances
Abstractions to build projection function (dropping parameters).
ProjLams | |
|
Instances
Data ProjLams Source # | |
Defined in Agda.TypeChecking.Monad.Base gfoldl :: (forall d b. Data d => c (d -> b) -> d -> c b) -> (forall g. g -> c g) -> ProjLams -> c ProjLams # gunfold :: (forall b r. Data b => c (b -> r) -> c r) -> (forall r. r -> c r) -> Constr -> c ProjLams # toConstr :: ProjLams -> Constr # dataTypeOf :: ProjLams -> DataType # dataCast1 :: Typeable t => (forall d. Data d => c (t d)) -> Maybe (c ProjLams) # dataCast2 :: Typeable t => (forall d e. (Data d, Data e) => c (t d e)) -> Maybe (c ProjLams) # gmapT :: (forall b. Data b => b -> b) -> ProjLams -> ProjLams # gmapQl :: (r -> r' -> r) -> r -> (forall d. Data d => d -> r') -> ProjLams -> r # gmapQr :: forall r r'. (r' -> r -> r) -> r -> (forall d. Data d => d -> r') -> ProjLams -> r # gmapQ :: (forall d. Data d => d -> u) -> ProjLams -> [u] # gmapQi :: Int -> (forall d. Data d => d -> u) -> ProjLams -> u # gmapM :: Monad m => (forall d. Data d => d -> m d) -> ProjLams -> m ProjLams # gmapMp :: MonadPlus m => (forall d. Data d => d -> m d) -> ProjLams -> m ProjLams # gmapMo :: MonadPlus m => (forall d. Data d => d -> m d) -> ProjLams -> m ProjLams # | |
Show ProjLams Source # | |
Null ProjLams Source # | |
KillRange ProjLams Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
Abstract ProjLams Source # | |
Apply ProjLams Source # | |
EmbPrj ProjLams Source # | |
projDropPars :: Projection -> ProjOrigin -> Term Source #
Building the projection function (which drops the parameters).
projArgInfo :: Projection -> ArgInfo Source #
The info of the principal (record) argument.
data EtaEquality Source #
Should a record type admit eta-equality?
Specified | User specifed 'eta-equality' or 'no-eta-equality'. |
| |
Inferred | Positivity checker inferred whether eta is safe. |
|
Instances
setEtaEquality :: EtaEquality -> HasEta -> EtaEquality Source #
Make sure we do not overwrite a user specification.
data FunctionFlag Source #
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? |
Instances
Instances
Eq CompKit Source # | |
Data CompKit Source # | |
Defined in Agda.TypeChecking.Monad.Base gfoldl :: (forall d b. Data d => c (d -> b) -> d -> c b) -> (forall g. g -> c g) -> CompKit -> c CompKit # gunfold :: (forall b r. Data b => c (b -> r) -> c r) -> (forall r. r -> c r) -> Constr -> c CompKit # toConstr :: CompKit -> Constr # dataTypeOf :: CompKit -> DataType # dataCast1 :: Typeable t => (forall d. Data d => c (t d)) -> Maybe (c CompKit) # dataCast2 :: Typeable t => (forall d e. (Data d, Data e) => c (t d e)) -> Maybe (c CompKit) # gmapT :: (forall b. Data b => b -> b) -> CompKit -> CompKit # gmapQl :: (r -> r' -> r) -> r -> (forall d. Data d => d -> r') -> CompKit -> r # gmapQr :: forall r r'. (r' -> r -> r) -> r -> (forall d. Data d => d -> r') -> CompKit -> r # gmapQ :: (forall d. Data d => d -> u) -> CompKit -> [u] # gmapQi :: Int -> (forall d. Data d => d -> u) -> CompKit -> u # gmapM :: Monad m => (forall d. Data d => d -> m d) -> CompKit -> m CompKit # gmapMp :: MonadPlus m => (forall d. Data d => d -> m d) -> CompKit -> m CompKit # gmapMo :: MonadPlus m => (forall d. Data d => d -> m d) -> CompKit -> m CompKit # | |
Ord CompKit Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
Show CompKit Source # | |
KillRange CompKit Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
EmbPrj CompKit Source # | |
NamesIn CompKit Source # | |
Axiom | Postulate |
DataOrRecSig | Data or record type signature that doesn't yet have a definition |
| |
GeneralizableVar | Generalizable variable (introduced in |
AbstractDefn Defn | Returned by |
Function | |
| |
Datatype | |
| |
Record | |
| |
Constructor | |
| |
Primitive | Primitive or builtin functions. |
|
Instances
Data Defn Source # | |
Defined in Agda.TypeChecking.Monad.Base gfoldl :: (forall d b. Data d => c (d -> b) -> d -> c b) -> (forall g. g -> c g) -> Defn -> c Defn # gunfold :: (forall b r. Data b => c (b -> r) -> c r) -> (forall r. r -> c r) -> Constr -> c Defn # dataTypeOf :: Defn -> DataType # dataCast1 :: Typeable t => (forall d. Data d => c (t d)) -> Maybe (c Defn) # dataCast2 :: Typeable t => (forall d e. (Data d, Data e) => c (t d e)) -> Maybe (c Defn) # gmapT :: (forall b. Data b => b -> b) -> Defn -> Defn # gmapQl :: (r -> r' -> r) -> r -> (forall d. Data d => d -> r') -> Defn -> r # gmapQr :: forall r r'. (r' -> r -> r) -> r -> (forall d. Data d => d -> r') -> Defn -> r # gmapQ :: (forall d. Data d => d -> u) -> Defn -> [u] # gmapQi :: Int -> (forall d. Data d => d -> u) -> Defn -> u # gmapM :: Monad m => (forall d. Data d => d -> m d) -> Defn -> m Defn # gmapMp :: MonadPlus m => (forall d. Data d => d -> m d) -> Defn -> m Defn # gmapMo :: MonadPlus m => (forall d. Data d => d -> m d) -> Defn -> m Defn # | |
Show Defn Source # | |
Pretty Defn Source # | |
KillRange Defn Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
Abstract Defn Source # | |
Apply Defn Source # | |
EmbPrj Defn Source # | |
NamesIn Defn Source # | |
InstantiateFull Defn Source # | |
Defined in Agda.TypeChecking.Reduce | |
Occurs Defn Source # | |
recRecursive :: Defn -> Bool Source #
Is the record type recursive?
recEtaEquality :: Defn -> HasEta Source #
emptyFunction :: Defn Source #
A template for creating Function
definitions, with sensible defaults.
isEmptyFunction :: Defn -> Bool Source #
Checking whether we are dealing with a function yet to be defined.
isCopatternLHS :: [Clause] -> Bool Source #
defIsRecord :: Defn -> Bool Source #
defIsDataOrRecord :: Defn -> Bool Source #
defConstructors :: Defn -> [QName] Source #
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?
Instances
redBind :: ReduceM (Reduced a a') -> (a -> b) -> (a' -> ReduceM (Reduced b b')) -> ReduceM (Reduced b b') Source #
Conceptually: redBind m f k = either (return . Left . f) k =<< m
Three cases: 1. not reduced, 2. reduced, but blocked, 3. reduced, not blocked.
NotReduced | |
Reduced (Blocked ()) |
data MaybeReduced a Source #
MaybeRed | |
|
Instances
Functor MaybeReduced Source # | |
Defined in Agda.TypeChecking.Monad.Base fmap :: (a -> b) -> MaybeReduced a -> MaybeReduced b # (<$) :: a -> MaybeReduced b -> MaybeReduced a # | |
IsProjElim e => IsProjElim (MaybeReduced e) Source # | |
Defined in Agda.TypeChecking.Monad.Base isProjElim :: MaybeReduced e -> Maybe (ProjOrigin, QName) Source # | |
PrettyTCM a => PrettyTCM (MaybeReduced a) Source # | |
Defined in Agda.TypeChecking.Pretty prettyTCM :: MonadPretty m => MaybeReduced a -> m Doc Source # |
type MaybeReducedArgs = [MaybeReduced (Arg Term)] Source #
type MaybeReducedElims = [MaybeReduced Elim] Source #
notReduced :: a -> MaybeReduced a Source #
data AllowedReduction Source #
Controlling reduce
.
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
allReductions :: AllowedReductions Source #
Not quite all reductions (skip non-terminating reductions)
PrimFun | |
|
primFun :: QName -> Arity -> ([Arg Term] -> ReduceM (Reduced MaybeReducedArgs Term)) -> PrimFun Source #
defClauses :: Definition -> [Clause] Source #
defParameters :: Definition -> Maybe Nat Source #
defCompilerPragmas :: BackendName -> Definition -> [CompilerPragma] Source #
defDelayed :: Definition -> Delayed Source #
Are the clauses of this definition delayed?
defNonterminating :: Definition -> Bool Source #
Has the definition failed the termination checker?
defTerminationUnconfirmed :: Definition -> Bool Source #
Has the definition not termination checked or did the check fail?
defAbstract :: Definition -> IsAbstract Source #
defForced :: Definition -> [IsForced] Source #
Injectivity
type FunctionInverse = FunctionInverse' Clause Source #
type InversionMap c = Map TermHead [c] Source #
data FunctionInverse' c Source #
Instances
Instances
Eq TermHead Source # | |
Data TermHead Source # | |
Defined in Agda.TypeChecking.Monad.Base gfoldl :: (forall d b. Data d => c (d -> b) -> d -> c b) -> (forall g. g -> c g) -> TermHead -> c TermHead # gunfold :: (forall b r. Data b => c (b -> r) -> c r) -> (forall r. r -> c r) -> Constr -> c TermHead # toConstr :: TermHead -> Constr # dataTypeOf :: TermHead -> DataType # dataCast1 :: Typeable t => (forall d. Data d => c (t d)) -> Maybe (c TermHead) # dataCast2 :: Typeable t => (forall d e. (Data d, Data e) => c (t d e)) -> Maybe (c TermHead) # gmapT :: (forall b. Data b => b -> b) -> TermHead -> TermHead # gmapQl :: (r -> r' -> r) -> r -> (forall d. Data d => d -> r') -> TermHead -> r # gmapQr :: forall r r'. (r' -> r -> r) -> r -> (forall d. Data d => d -> r') -> TermHead -> r # gmapQ :: (forall d. Data d => d -> u) -> TermHead -> [u] # gmapQi :: Int -> (forall d. Data d => d -> u) -> TermHead -> u # gmapM :: Monad m => (forall d. Data d => d -> m d) -> TermHead -> m TermHead # gmapMp :: MonadPlus m => (forall d. Data d => d -> m d) -> TermHead -> m TermHead # gmapMo :: MonadPlus m => (forall d. Data d => d -> m d) -> TermHead -> m TermHead # | |
Ord TermHead Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
Show TermHead Source # | |
Pretty TermHead Source # | |
KillRange TermHead Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
EmbPrj TermHead Source # | |
Mutual blocks
Instances
Enum MutualId Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
Eq MutualId Source # | |
Data MutualId Source # | |
Defined in Agda.TypeChecking.Monad.Base gfoldl :: (forall d b. Data d => c (d -> b) -> d -> c b) -> (forall g. g -> c g) -> MutualId -> c MutualId # gunfold :: (forall b r. Data b => c (b -> r) -> c r) -> (forall r. r -> c r) -> Constr -> c MutualId # toConstr :: MutualId -> Constr # dataTypeOf :: MutualId -> DataType # dataCast1 :: Typeable t => (forall d. Data d => c (t d)) -> Maybe (c MutualId) # dataCast2 :: Typeable t => (forall d e. (Data d, Data e) => c (t d e)) -> Maybe (c MutualId) # gmapT :: (forall b. Data b => b -> b) -> MutualId -> MutualId # gmapQl :: (r -> r' -> r) -> r -> (forall d. Data d => d -> r') -> MutualId -> r # gmapQr :: forall r r'. (r' -> r -> r) -> r -> (forall d. Data d => d -> r') -> MutualId -> r # gmapQ :: (forall d. Data d => d -> u) -> MutualId -> [u] # gmapQi :: Int -> (forall d. Data d => d -> u) -> MutualId -> u # gmapM :: Monad m => (forall d. Data d => d -> m d) -> MutualId -> m MutualId # gmapMp :: MonadPlus m => (forall d. Data d => d -> m d) -> MutualId -> m MutualId # gmapMo :: MonadPlus m => (forall d. Data d => d -> m d) -> MutualId -> m MutualId # | |
Num MutualId Source # | |
Ord MutualId Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
Show MutualId Source # | |
KillRange MutualId Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
HasFresh MutualId Source # | |
EmbPrj MutualId Source # | |
Statistics
Trace
Instances
Data Call Source # | |
Defined in Agda.TypeChecking.Monad.Base gfoldl :: (forall d b. Data d => c (d -> b) -> d -> c b) -> (forall g. g -> c g) -> Call -> c Call # gunfold :: (forall b r. Data b => c (b -> r) -> c r) -> (forall r. r -> c r) -> Constr -> c Call # dataTypeOf :: Call -> DataType # dataCast1 :: Typeable t => (forall d. Data d => c (t d)) -> Maybe (c Call) # dataCast2 :: Typeable t => (forall d e. (Data d, Data e) => c (t d e)) -> Maybe (c Call) # gmapT :: (forall b. Data b => b -> b) -> Call -> Call # gmapQl :: (r -> r' -> r) -> r -> (forall d. Data d => d -> r') -> Call -> r # gmapQr :: forall r r'. (r' -> r -> r) -> r -> (forall d. Data d => d -> r') -> Call -> r # gmapQ :: (forall d. Data d => d -> u) -> Call -> [u] # gmapQi :: Int -> (forall d. Data d => d -> u) -> Call -> u # gmapM :: Monad m => (forall d. Data d => d -> m d) -> Call -> m Call # gmapMp :: MonadPlus m => (forall d. Data d => d -> m d) -> Call -> m Call # gmapMo :: MonadPlus m => (forall d. Data d => d -> m d) -> Call -> m Call # | |
Pretty Call Source # | |
HasRange Call Source # | |
PrettyTCM Call Source # | |
Defined in Agda.TypeChecking.Pretty.Call |
Instance table
type InstanceTable = Map QName (Set QName) Source #
The instance table is a Map
associating to every name of
recorddata typepostulate its list of instances
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.
Builtin things
data BuiltinDescriptor Source #
BuiltinData (TCM Type) [String] | |
BuiltinDataCons (TCM Type) | |
BuiltinPrim String (Term -> TCM ()) | |
BuiltinPostulate Relevance (TCM Type) | |
BuiltinUnknown (Maybe (TCM Type)) (Term -> Type -> TCM ()) | Builtin of any kind.
Type can be checked ( |
Instances
Functor Builtin Source # | |
Foldable Builtin Source # | |
Defined in Agda.TypeChecking.Monad.Base 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 # | |
Show pf => Show (Builtin pf) Source # | |
EmbPrj a => EmbPrj (Builtin a) Source # | |
InstantiateFull a => InstantiateFull (Builtin a) Source # | |
Defined in Agda.TypeChecking.Reduce |
Highlighting levels
data HighlightingLevel Source #
How much highlighting should be sent to the user interface?
None | |
NonInteractive | |
Interactive | This includes both non-interactive highlighting and interactive highlighting of the expression that is currently being type-checked. |
Instances
data HighlightingMethod Source #
How should highlighting be sent to the user interface?
Instances
ifTopLevelAndHighlightingLevelIsOr :: MonadTCM tcm => HighlightingLevel -> Bool -> tcm () -> tcm () Source #
ifTopLevelAndHighlightingLevelIs l b m
runs m
when we're
type-checking the top-level module and either the highlighting
level is at least l
or b
is True
.
ifTopLevelAndHighlightingLevelIs :: MonadTCM tcm => HighlightingLevel -> tcm () -> tcm () Source #
ifTopLevelAndHighlightingLevelIs l m
runs m
when we're
type-checking the top-level module and the highlighting level is
at least l
.
Type checking environment
TCEnv | |
|
Instances
Data TCEnv Source # | |
Defined in Agda.TypeChecking.Monad.Base gfoldl :: (forall d b. Data d => c (d -> b) -> d -> c b) -> (forall g. g -> c g) -> TCEnv -> c TCEnv # gunfold :: (forall b r. Data b => c (b -> r) -> c r) -> (forall r. r -> c r) -> Constr -> c TCEnv # dataTypeOf :: TCEnv -> DataType # dataCast1 :: Typeable t => (forall d. Data d => c (t d)) -> Maybe (c TCEnv) # dataCast2 :: Typeable t => (forall d e. (Data d, Data e) => c (t d e)) -> Maybe (c TCEnv) # gmapT :: (forall b. Data b => b -> b) -> TCEnv -> TCEnv # gmapQl :: (r -> r' -> r) -> r -> (forall d. Data d => d -> r') -> TCEnv -> r # gmapQr :: forall r r'. (r' -> r -> r) -> r -> (forall d. Data d => d -> r') -> TCEnv -> r # gmapQ :: (forall d. Data d => d -> u) -> TCEnv -> [u] # gmapQi :: Int -> (forall d. Data d => d -> u) -> TCEnv -> u # gmapM :: Monad m => (forall d. Data d => d -> m d) -> TCEnv -> m TCEnv # gmapMp :: MonadPlus m => (forall d. Data d => d -> m d) -> TCEnv -> m TCEnv # gmapMo :: MonadPlus m => (forall d. Data d => d -> m d) -> TCEnv -> m TCEnv # | |
LensIsAbstract TCEnv Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
LensRelevance TCEnv Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
LensQuantity TCEnv Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
LensModality TCEnv Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
LensTCEnv TCEnv Source # | |
data UnquoteFlags Source #
Instances
Data UnquoteFlags Source # | |
Defined in Agda.TypeChecking.Monad.Base gfoldl :: (forall d b. Data d => c (d -> b) -> d -> c b) -> (forall g. g -> c g) -> UnquoteFlags -> c UnquoteFlags # gunfold :: (forall b r. Data b => c (b -> r) -> c r) -> (forall r. r -> c r) -> Constr -> c UnquoteFlags # toConstr :: UnquoteFlags -> Constr # dataTypeOf :: UnquoteFlags -> DataType # dataCast1 :: Typeable t => (forall d. Data d => c (t d)) -> Maybe (c UnquoteFlags) # dataCast2 :: Typeable t => (forall d e. (Data d, Data e) => c (t d e)) -> Maybe (c UnquoteFlags) # gmapT :: (forall b. Data b => b -> b) -> UnquoteFlags -> UnquoteFlags # gmapQl :: (r -> r' -> r) -> r -> (forall d. Data d => d -> r') -> UnquoteFlags -> r # gmapQr :: forall r r'. (r' -> r -> r) -> r -> (forall d. Data d => d -> r') -> UnquoteFlags -> r # gmapQ :: (forall d. Data d => d -> u) -> UnquoteFlags -> [u] # gmapQi :: Int -> (forall d. Data d => d -> u) -> UnquoteFlags -> u # gmapM :: Monad m => (forall d. Data d => d -> m d) -> UnquoteFlags -> m UnquoteFlags # gmapMp :: MonadPlus m => (forall d. Data d => d -> m d) -> UnquoteFlags -> m UnquoteFlags # gmapMo :: MonadPlus m => (forall d. Data d => d -> m d) -> UnquoteFlags -> m UnquoteFlags # |
e-prefixed lenses
eAnonymousModules :: Lens' [(ModuleName, Nat)] TCEnv Source #
eTerminationCheck :: Lens' (TerminationCheck ()) TCEnv Source #
Context
type Context = [ContextEntry] Source #
The Context
is a stack of ContextEntry
s.
Let bindings
Abstract mode
data AbstractMode Source #
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
Eq AbstractMode Source # | |
Defined in Agda.TypeChecking.Monad.Base (==) :: AbstractMode -> AbstractMode -> Bool # (/=) :: AbstractMode -> AbstractMode -> Bool # | |
Data AbstractMode Source # | |
Defined in Agda.TypeChecking.Monad.Base gfoldl :: (forall d b. Data d => c (d -> b) -> d -> c b) -> (forall g. g -> c g) -> AbstractMode -> c AbstractMode # gunfold :: (forall b r. Data b => c (b -> r) -> c r) -> (forall r. r -> c r) -> Constr -> c AbstractMode # toConstr :: AbstractMode -> Constr # dataTypeOf :: AbstractMode -> DataType # dataCast1 :: Typeable t => (forall d. Data d => c (t d)) -> Maybe (c AbstractMode) # dataCast2 :: Typeable t => (forall d e. (Data d, Data e) => c (t d e)) -> Maybe (c AbstractMode) # gmapT :: (forall b. Data b => b -> b) -> AbstractMode -> AbstractMode # gmapQl :: (r -> r' -> r) -> r -> (forall d. Data d => d -> r') -> AbstractMode -> r # gmapQr :: forall r r'. (r' -> r -> r) -> r -> (forall d. Data d => d -> r') -> AbstractMode -> r # gmapQ :: (forall d. Data d => d -> u) -> AbstractMode -> [u] # gmapQi :: Int -> (forall d. Data d => d -> u) -> AbstractMode -> u # gmapM :: Monad m => (forall d. Data d => d -> m d) -> AbstractMode -> m AbstractMode # gmapMp :: MonadPlus m => (forall d. Data d => d -> m d) -> AbstractMode -> m AbstractMode # gmapMo :: MonadPlus m => (forall d. Data d => d -> m d) -> AbstractMode -> m AbstractMode # | |
Show AbstractMode Source # | |
Defined in Agda.TypeChecking.Monad.Base showsPrec :: Int -> AbstractMode -> ShowS # show :: AbstractMode -> String # showList :: [AbstractMode] -> ShowS # |
aDefToMode :: IsAbstract -> AbstractMode Source #
aModeToDef :: AbstractMode -> Maybe IsAbstract Source #
Insertion of implicit arguments
data ExpandHidden Source #
ExpandLast | Add implicit arguments in the end until type is no longer hidden |
DontExpandLast | Do not append implicit arguments. |
ReallyDontExpandLast | Makes |
Instances
Eq ExpandHidden Source # | |
Defined in Agda.TypeChecking.Monad.Base (==) :: ExpandHidden -> ExpandHidden -> Bool # (/=) :: ExpandHidden -> ExpandHidden -> Bool # | |
Data ExpandHidden Source # | |
Defined in Agda.TypeChecking.Monad.Base gfoldl :: (forall d b. Data d => c (d -> b) -> d -> c b) -> (forall g. g -> c g) -> ExpandHidden -> c ExpandHidden # gunfold :: (forall b r. Data b => c (b -> r) -> c r) -> (forall r. r -> c r) -> Constr -> c ExpandHidden # toConstr :: ExpandHidden -> Constr # dataTypeOf :: ExpandHidden -> DataType # dataCast1 :: Typeable t => (forall d. Data d => c (t d)) -> Maybe (c ExpandHidden) # dataCast2 :: Typeable t => (forall d e. (Data d, Data e) => c (t d e)) -> Maybe (c ExpandHidden) # gmapT :: (forall b. Data b => b -> b) -> ExpandHidden -> ExpandHidden # gmapQl :: (r -> r' -> r) -> r -> (forall d. Data d => d -> r') -> ExpandHidden -> r # gmapQr :: forall r r'. (r' -> r -> r) -> r -> (forall d. Data d => d -> r') -> ExpandHidden -> r # gmapQ :: (forall d. Data d => d -> u) -> ExpandHidden -> [u] # gmapQi :: Int -> (forall d. Data d => d -> u) -> ExpandHidden -> u # gmapM :: Monad m => (forall d. Data d => d -> m d) -> ExpandHidden -> m ExpandHidden # gmapMp :: MonadPlus m => (forall d. Data d => d -> m d) -> ExpandHidden -> m ExpandHidden # gmapMo :: MonadPlus m => (forall d. Data d => d -> m d) -> ExpandHidden -> m ExpandHidden # |
isDontExpandLast :: ExpandHidden -> Bool Source #
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.
Instances
Type checking warnings (aka non-fatal errors)
A non-fatal error is an error which does not prevent us from checking the document further and interacting with the user.
NicifierIssue DeclarationWarning | |
TerminationIssue [TerminationError] | |
UnreachableClauses QName [Range] | `UnreachableClauses f rs` means that the clauses in |
CoverageIssue QName [(Telescope, [NamedArg DeBruijnPattern])] | `CoverageIssue f pss` means that |
CoverageNoExactSplit QName [Clause] | |
NotStrictlyPositive QName (Seq OccursWhere) | |
UnsolvedMetaVariables [Range] | Do not use directly with |
UnsolvedInteractionMetas [Range] | Do not use directly with |
UnsolvedConstraints Constraints | Do not use directly with |
CantGeneralizeOverSorts [MetaId] | |
AbsurdPatternRequiresNoRHS [NamedArg DeBruijnPattern] | |
OldBuiltin String String | In `OldBuiltin old new`, the BUILTIN old has been replaced by new |
EmptyRewritePragma | If the user wrote just |
IllformedAsClause String | If the user wrote something other than an unqualified name
in the |
ClashesViaRenaming NameOrModule [Name] | If a |
UselessPublic | If the user opens a module public before the module header. (See issue #2377.) |
UselessInline QName | |
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. Generic warnings for one-off things |
GenericWarning Doc | Harmless generic warning (not an error) |
GenericNonFatalError Doc | Generic error which doesn't abort proceedings (not a warning) Safe flag errors |
SafeFlagPostulate Name | |
SafeFlagPragma [String] | Unsafe OPTIONS. |
SafeFlagNonTerminating | |
SafeFlagTerminating | |
SafeFlagWithoutKFlagPrimEraseEquality | |
WithoutKFlagPrimEraseEquality | |
SafeFlagNoPositivityCheck | |
SafeFlagPolarity | |
SafeFlagNoUniverseCheck | |
SafeFlagNoCoverageCheck | |
SafeFlagInjective | |
SafeFlagEta | ETA pragma is unsafe. |
ParseWarning ParseWarning | |
LibraryWarning LibWarning | |
DeprecationWarning String String String | `DeprecationWarning old new version`:
|
UserWarning String | User-defined warning (e.g. to mention that a name is deprecated) |
FixityInRenamingModule (NonEmpty Range) | Fixity of modules cannot be changed via renaming (since modules have no fixity). |
ModuleDoesntExport QName [ImportedName] | Some imported names are not actually exported by the source module |
InfectiveImport String ModuleName | Importing a file using an infective option into one which doesn't |
CoInfectiveImport String ModuleName | Importing a file not using a coinfective option from one which does |
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 |
PragmaCompileErased BackendName QName | COMPILE directive for an erased symbol |
NotInScopeW [QName] | Out of scope error we can recover from |
Instances
Data Warning Source # | |
Defined in Agda.TypeChecking.Monad.Base gfoldl :: (forall d b. Data d => c (d -> b) -> d -> c b) -> (forall g. g -> c g) -> Warning -> c Warning # gunfold :: (forall b r. Data b => c (b -> r) -> c r) -> (forall r. r -> c r) -> Constr -> c Warning # toConstr :: Warning -> Constr # dataTypeOf :: Warning -> DataType # dataCast1 :: Typeable t => (forall d. Data d => c (t d)) -> Maybe (c Warning) # dataCast2 :: Typeable t => (forall d e. (Data d, Data e) => c (t d e)) -> Maybe (c Warning) # gmapT :: (forall b. Data b => b -> b) -> Warning -> Warning # gmapQl :: (r -> r' -> r) -> r -> (forall d. Data d => d -> r') -> Warning -> r # gmapQr :: forall r r'. (r' -> r -> r) -> r -> (forall d. Data d => d -> r') -> Warning -> r # gmapQ :: (forall d. Data d => d -> u) -> Warning -> [u] # gmapQi :: Int -> (forall d. Data d => d -> u) -> Warning -> u # gmapM :: Monad m => (forall d. Data d => d -> m d) -> Warning -> m Warning # gmapMp :: MonadPlus m => (forall d. Data d => d -> m d) -> Warning -> m Warning # gmapMo :: MonadPlus m => (forall d. Data d => d -> m d) -> Warning -> m Warning # | |
Show Warning Source # | |
EmbPrj Warning Source # | |
PrettyTCM Warning Source # | |
Defined in Agda.TypeChecking.Pretty.Warning |
warningName :: Warning -> WarningName Source #
TCWarning | |
|
tcWarningOrigin :: TCWarning -> SrcFile Source #
Type checking errors
Information about a call.
CallInfo | |
|
Instances
Data CallInfo Source # | |
Defined in Agda.TypeChecking.Monad.Base gfoldl :: (forall d b. Data d => c (d -> b) -> d -> c b) -> (forall g. g -> c g) -> CallInfo -> c CallInfo # gunfold :: (forall b r. Data b => c (b -> r) -> c r) -> (forall r. r -> c r) -> Constr -> c CallInfo # toConstr :: CallInfo -> Constr # dataTypeOf :: CallInfo -> DataType # dataCast1 :: Typeable t => (forall d. Data d => c (t d)) -> Maybe (c CallInfo) # dataCast2 :: Typeable t => (forall d e. (Data d, Data e) => c (t d e)) -> Maybe (c CallInfo) # gmapT :: (forall b. Data b => b -> b) -> CallInfo -> CallInfo # gmapQl :: (r -> r' -> r) -> r -> (forall d. Data d => d -> r') -> CallInfo -> r # gmapQr :: forall r r'. (r' -> r -> r) -> r -> (forall d. Data d => d -> r') -> CallInfo -> r # gmapQ :: (forall d. Data d => d -> u) -> CallInfo -> [u] # gmapQi :: Int -> (forall d. Data d => d -> u) -> CallInfo -> u # gmapM :: Monad m => (forall d. Data d => d -> m d) -> CallInfo -> m CallInfo # gmapMp :: MonadPlus m => (forall d. Data d => d -> m d) -> CallInfo -> m CallInfo # gmapMo :: MonadPlus m => (forall d. Data d => d -> m d) -> CallInfo -> m CallInfo # | |
Show CallInfo Source # | |
Pretty CallInfo Source # | We only |
AllNames CallInfo Source # | |
PrettyTCM CallInfo Source # | |
Defined in Agda.TypeChecking.Pretty.Call |
data TerminationError Source #
Information about a mutual block which did not pass the termination checker.
TerminationError | |
|
Instances
Data TerminationError Source # | |
Defined in Agda.TypeChecking.Monad.Base gfoldl :: (forall d b. Data d => c (d -> b) -> d -> c b) -> (forall g. g -> c g) -> TerminationError -> c TerminationError # gunfold :: (forall b r. Data b => c (b -> r) -> c r) -> (forall r. r -> c r) -> Constr -> c TerminationError # toConstr :: TerminationError -> Constr # dataTypeOf :: TerminationError -> DataType # dataCast1 :: Typeable t => (forall d. Data d => c (t d)) -> Maybe (c TerminationError) # dataCast2 :: Typeable t => (forall d e. (Data d, Data e) => c (t d e)) -> Maybe (c TerminationError) # gmapT :: (forall b. Data b => b -> b) -> TerminationError -> TerminationError # gmapQl :: (r -> r' -> r) -> r -> (forall d. Data d => d -> r') -> TerminationError -> r # gmapQr :: forall r r'. (r' -> r -> r) -> r -> (forall d. Data d => d -> r') -> TerminationError -> r # gmapQ :: (forall d. Data d => d -> u) -> TerminationError -> [u] # gmapQi :: Int -> (forall d. Data d => d -> u) -> TerminationError -> u # gmapM :: Monad m => (forall d. Data d => d -> m d) -> TerminationError -> m TerminationError # gmapMp :: MonadPlus m => (forall d. Data d => d -> m d) -> TerminationError -> m TerminationError # gmapMo :: MonadPlus m => (forall d. Data d => d -> m d) -> TerminationError -> m TerminationError # | |
Show TerminationError Source # | |
Defined in Agda.TypeChecking.Monad.Base showsPrec :: Int -> TerminationError -> ShowS # show :: TerminationError -> String # showList :: [TerminationError] -> ShowS # |
data SplitError Source #
Error when splitting a pattern variable into possible constructor patterns.
NotADatatype (Closure Type) | Neither data type nor record. |
IrrelevantDatatype (Closure Type) | Data type, but in irrelevant position. |
ErasedDatatype Bool (Closure Type) | Data type, but in erased position.
If the boolean is |
CoinductiveDatatype (Closure Type) | Split on codata not allowed. UNUSED, but keep! -- | NoRecordConstructor Type -- ^ record type, but no constructor |
UnificationStuck | |
| |
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
Show SplitError Source # | |
Defined in Agda.TypeChecking.Monad.Base showsPrec :: Int -> SplitError -> ShowS # show :: SplitError -> String # showList :: [SplitError] -> ShowS # | |
PrettyTCM SplitError Source # | |
Defined in Agda.TypeChecking.Errors prettyTCM :: MonadPretty m => SplitError -> m Doc Source # |
data NegativeUnification Source #
Instances
Show NegativeUnification Source # | |
Defined in Agda.TypeChecking.Monad.Base showsPrec :: Int -> NegativeUnification -> ShowS # show :: NegativeUnification -> String # showList :: [NegativeUnification] -> ShowS # | |
PrettyTCM NegativeUnification Source # | |
Defined in Agda.TypeChecking.Errors prettyTCM :: MonadPretty m => NegativeUnification -> m Doc Source # |
data UnificationFailure Source #
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
Show UnificationFailure Source # | |
Defined in Agda.TypeChecking.Monad.Base showsPrec :: Int -> UnificationFailure -> ShowS # show :: UnificationFailure -> String # showList :: [UnificationFailure] -> ShowS # | |
PrettyTCM UnificationFailure Source # | |
Defined in Agda.TypeChecking.Errors prettyTCM :: MonadPretty m => UnificationFailure -> m Doc Source # |
data UnquoteError Source #
BadVisibility String (Arg Term) | |
ConInsteadOfDef QName String String | |
DefInsteadOfCon QName String String | |
NonCanonical String Term | |
BlockedOnMeta TCState MetaId | |
UnquotePanic String |
Instances
Show UnquoteError Source # | |
Defined in Agda.TypeChecking.Monad.Base showsPrec :: Int -> UnquoteError -> ShowS # show :: UnquoteError -> String # showList :: [UnquoteError] -> ShowS # |
Instances
data LHSOrPatSyn Source #
Distinguish error message when parsing lhs or pattern synonym, resp.
Instances
Eq LHSOrPatSyn Source # | |
Defined in Agda.TypeChecking.Monad.Base (==) :: LHSOrPatSyn -> LHSOrPatSyn -> Bool # (/=) :: LHSOrPatSyn -> LHSOrPatSyn -> Bool # | |
Show LHSOrPatSyn Source # | |
Defined in Agda.TypeChecking.Monad.Base showsPrec :: Int -> LHSOrPatSyn -> ShowS # show :: LHSOrPatSyn -> String # showList :: [LHSOrPatSyn] -> ShowS # |
Type-checking errors.
TypeError | |
| |
Exception Range Doc | |
IOException TCState Range IOException | The first argument is the state in which the error was raised. |
PatternErr | The exception which is usually caught.
Raised for pattern violations during unification ( |
Instances
Show TCErr Source # | |
Exception TCErr Source # | |
Defined in Agda.TypeChecking.Monad.Base toException :: TCErr -> SomeException # fromException :: SomeException -> Maybe TCErr # displayException :: TCErr -> String # | |
Error TCErr Source # | |
HasRange TCErr Source # | |
PrettyTCM TCErr Source # | |
Defined in Agda.TypeChecking.Errors | |
MonadError TCErr TCM Source # | |
Defined in Agda.TypeChecking.Monad.Base throwError :: TCErr -> TCM a # | |
MonadError TCErr IM Source # | |
Defined in Agda.TypeChecking.Monad.Base throwError :: TCErr -> IM a # | |
MonadError TCErr TerM Source # | |
Defined in Agda.Termination.Monad throwError :: TCErr -> TerM a # | |
Monad m => MonadError TCErr (PureConversionT m) Source # | |
Defined in Agda.TypeChecking.Conversion.Pure throwError :: TCErr -> PureConversionT m a # catchError :: PureConversionT m a -> (TCErr -> PureConversionT m a) -> PureConversionT m a # |
Accessing options
class (Functor m, Applicative m, Monad m) => HasOptions m where Source #
Nothing
pragmaOptions :: m PragmaOptions Source #
Returns the pragma options which are currently in effect.
default pragmaOptions :: (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 :: (HasOptions n, MonadTrans t, m ~ t n) => m CommandLineOptions Source #
Instances
sizedTypesOption :: HasOptions m => m Bool Source #
guardednessOption :: HasOptions m => m Bool Source #
withoutKOption :: HasOptions m => m Bool Source #
getIncludeDirs :: HasOptions m => m [AbsolutePath] Source #
Gets the include directories.
Precondition: optAbsoluteIncludePaths
must be nonempty (i.e.
setCommandLineOptions
must have run).
enableCaching :: HasOptions m => m Bool Source #
The reduce monad
Environment of the reduce monad.
Instances
fmapReduce :: (a -> b) -> ReduceM a -> ReduceM b Source #
runReduceM :: ReduceM a -> TCM a Source #
runReduceF :: (a -> ReduceM b) -> TCM (a -> b) Source #
class (Applicative m, MonadTCEnv m, ReadTCState m, HasOptions m) => MonadReduce m where Source #
liftReduce :: ReduceM a -> m a Source #
Instances
Monad with read-only TCEnv
class Monad m => MonadTCEnv m where Source #
MonadTCEnv
made into its own dedicated service class.
This allows us to use MonadReader
for ReaderT
extensions of TCM
.
Instances
MonadTCEnv ReduceM Source # | |
MonadTCEnv AbsToCon Source # | |
MonadTCEnv TerM Source # | |
MonadTCEnv m => MonadTCEnv (MaybeT m) Source # | |
MonadTCEnv m => MonadTCEnv (ListT m) Source # | |
MonadIO m => MonadTCEnv (TCMT m) Source # | |
MonadTCEnv m => MonadTCEnv (ChangeT m) Source # | |
MonadTCEnv m => MonadTCEnv (NamesT m) Source # | |
MonadTCEnv m => MonadTCEnv (PureConversionT m) Source # | |
Defined in Agda.TypeChecking.Conversion.Pure askTC :: PureConversionT m TCEnv Source # localTC :: (TCEnv -> TCEnv) -> PureConversionT m a -> PureConversionT m a Source # | |
MonadTCEnv m => MonadTCEnv (ExceptT err m) Source # | |
MonadTCEnv m => MonadTCEnv (ReaderT r m) Source # | |
MonadTCEnv m => MonadTCEnv (IdentityT m) Source # | |
MonadTCEnv m => MonadTCEnv (StateT s m) Source # | |
(Monoid w, MonadTCEnv m) => MonadTCEnv (WriterT w m) Source # | |
asksTC :: MonadTCEnv m => (TCEnv -> a) -> m a Source #
locallyTC :: MonadTCEnv m => Lens' a TCEnv -> (a -> a) -> m b -> m b Source #
Modify the lens-indicated part of the TCEnv
in a subcomputation.
Monad with mutable TCState
class Monad m => MonadTCState m where Source #
MonadTCState
made into its own dedicated service class.
This allows us to use MonadState
for StateT
extensions of TCM
.
Instances
MonadTCState TerM Source # | |
MonadTCState m => MonadTCState (MaybeT m) Source # | |
MonadTCState m => MonadTCState (ListT m) Source # | |
MonadIO m => MonadTCState (TCMT m) Source # | |
MonadTCState m => MonadTCState (ChangeT m) Source # | |
MonadTCState m => MonadTCState (NamesT m) Source # | |
MonadTCState m => MonadTCState (ExceptT err m) Source # | |
MonadTCState m => MonadTCState (ReaderT r m) Source # | |
MonadTCState m => MonadTCState (IdentityT m) Source # | |
MonadTCState m => MonadTCState (StateT s m) Source # | |
(Monoid w, MonadTCState m) => MonadTCState (WriterT w m) Source # | |
TCState
accessors (no lenses)
getsTC :: ReadTCState m => (TCState -> a) -> m a Source #
modifyTC' :: MonadTCState m => (TCState -> TCState) -> m () Source #
A variant of modifyTC
in which the computation is strict in the
new state.
TCState
accessors via lenses
setTCLens :: MonadTCState m => Lens' a TCState -> a -> m () infix 4 Source #
Overwrite the part of the TCState
focused on by the lens.
modifyTCLens :: MonadTCState m => Lens' a TCState -> (a -> a) -> m () Source #
Modify the part of the TCState
focused on by the lens.
modifyTCLensM :: MonadTCState m => Lens' a TCState -> (a -> m a) -> m () Source #
Modify a part of the state monadically.
stateTCLens :: MonadTCState m => Lens' a TCState -> (a -> (r, a)) -> m r Source #
Modify the part of the TCState
focused on by the lens, and return some result.
stateTCLensM :: MonadTCState m => Lens' a TCState -> (a -> m (r, a)) -> m r Source #
Modify a part of the state monadically, and return some result.
Type checking monad transformer
Instances
returnTCMT :: MonadIO m => a -> TCMT m a Source #
catchError_ :: TCM a -> (TCErr -> TCM a) -> TCM a Source #
Preserve the state of the failing computation.
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.
class (Applicative tcm, MonadIO tcm, MonadTCEnv tcm, MonadTCState tcm, HasOptions tcm) => MonadTCM tcm where Source #
Embedding a TCM computation.
Nothing
Instances
MonadTCM TerM Source # | |
MonadTCM tcm => MonadTCM (MaybeT tcm) Source # | |
MonadTCM tcm => MonadTCM (ListT tcm) Source # | |
MonadIO m => MonadTCM (TCMT m) Source # | |
MonadTCM tcm => MonadTCM (ChangeT tcm) Source # | |
MonadTCM m => MonadTCM (NamesT m) Source # | |
MonadTCM tcm => MonadTCM (ExceptT err tcm) Source # | |
MonadTCM tcm => MonadTCM (ReaderT r tcm) Source # | |
MonadTCM tcm => MonadTCM (IdentityT tcm) Source # | |
MonadTCM tcm => MonadTCM (StateT s tcm) Source # | |
(Monoid w, MonadTCM tcm) => MonadTCM (WriterT w tcm) Source # | |
patternViolation :: MonadError TCErr m => m a Source #
internalError :: MonadTCM tcm => String -> tcm a Source #
genericError :: (MonadTCEnv m, ReadTCState m, MonadError TCErr m) => String -> m a Source #
genericDocError :: (MonadTCEnv m, ReadTCState m, MonadError TCErr m) => Doc -> m a Source #
typeError :: (MonadTCEnv m, ReadTCState m, MonadError TCErr m) => TypeError -> m a Source #
typeError_ :: (MonadTCEnv m, ReadTCState m) => TypeError -> m TCErr Source #
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 #
runSafeTCM :: TCM a -> TCState -> IO (a, TCState) Source #
runSafeTCM
runs a safe TCM
action (a TCM
action which cannot fail)
in the initial environment.
forkTCM :: TCM a -> 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.
Names for generated definitions
patternInTeleName :: String Source #
Base name for patterns in telescopes
extendedLambdaName :: String Source #
Base name for extended lambda patterns
isExtendedLambdaName :: QName -> Bool Source #
Check whether we have an definition from an extended lambda.
absurdLambdaName :: String Source #
Name of absurdLambda definitions.
isAbsurdLambdaName :: QName -> Bool Source #
Check whether we have an definition from an absurd lambda.
generalizedFieldName :: String Source #
Base name for generalized variable projections
getGeneralizedFieldName :: QName -> Maybe String Source #
Check whether we have a generalized variable field