Safe Haskell | None |
---|---|
Language | Haskell2010 |
Interface for compiler backend writers.
Synopsis
- module Agda.Compiler.Backend.Base
- type Backend = Backend_boot TCM
- type Backend' opts env menv mod def = Backend'_boot TCM opts env menv mod def
- data Recompile menv mod
- data IsMain
- type Flag opts = opts -> OptM opts
- toTreeless :: EvaluationStrategy -> QName -> TCM (Maybe TTerm)
- module Agda.Syntax.Treeless
- data TCErr
- = TypeError { }
- | Exception Range Doc
- | IOException TCState Range IOException
- | PatternErr Blocker
- 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 InstanceInfo
- defCopy :: Bool
- defMatchable :: Set QName
- defNoCompilation :: Bool
- defInjective :: Bool
- defCopatternLHS :: Bool
- defBlocked :: Blocked_
- defLanguage :: !Language
- theDef :: Defn
- data Builtin pf
- = Builtin Term
- | Prim pf
- | BuiltinRewriteRelations (Set QName)
- data System = System {
- systemTel :: Telescope
- systemClauses :: [(Face, Term)]
- pattern Function :: [Clause] -> Maybe CompiledClauses -> Maybe SplitTree -> Maybe Compiled -> [Clause] -> FunctionInverse -> Maybe [QName] -> Either ProjectionLikenessMissing Projection -> SmallSet FunctionFlag -> Maybe Bool -> Maybe ExtLamInfo -> Maybe QName -> Maybe QName -> IsOpaque -> Defn
- type TCM = TCMT IO
- newtype TCMT (m :: Type -> Type) a = TCM {}
- class (Functor m, Applicative m, MonadFail m, HasOptions m, MonadDebug m, MonadTCEnv m) => HasConstInfo (m :: Type -> Type) where
- getConstInfo :: QName -> m Definition
- getConstInfo' :: QName -> m (Either SigError Definition)
- getRewriteRulesFor :: QName -> m RewriteRules
- getErasedConArgs :: HasConstInfo m => QName -> m [Bool]
- getTreeless :: HasConstInfo m => QName -> m (Maybe TTerm)
- data TypeCheckAction
- pattern Primitive :: IsAbstract -> PrimitiveId -> [Clause] -> FunctionInverse -> Maybe CompiledClauses -> IsOpaque -> Defn
- data MetaClass
- newtype ReduceM a = ReduceM {}
- runReduceM :: ReduceM a -> TCM a
- class (Functor m, Applicative m, Monad m) => MonadDebug (m :: Type -> Type) where
- formatDebugMessage :: VerboseKey -> VerboseLevel -> TCM Doc -> m String
- traceDebugMessage :: VerboseKey -> VerboseLevel -> String -> m a -> m a
- verboseBracket :: VerboseKey -> VerboseLevel -> String -> m a -> m a
- getVerbosity :: m Verbosity
- getProfileOptions :: m ProfileOptions
- isDebugPrinting :: m Bool
- nowDebugPrinting :: m a -> m a
- __IMPOSSIBLE_VERBOSE__ :: (HasCallStack, MonadDebug m) => String -> m a
- data HighlightingLevel
- data HighlightingMethod
- data Comparison
- data Polarity
- data Open a = OpenThing {}
- data MetaInfo = MetaInfo {}
- type Constraints = [ProblemConstraint]
- data Constraint
- = ValueCmp Comparison CompareAs Term Term
- | ValueCmpOnFace Comparison Term Type Term Term
- | ElimCmp [Polarity] [IsForced] Type Term [Elim] [Elim]
- | SortCmp Comparison Sort Sort
- | LevelCmp Comparison Level Level
- | HasBiggerSort Sort
- | HasPTSRule (Dom Type) (Abs Sort)
- | CheckDataSort QName Sort
- | CheckMetaInst MetaId
- | CheckType Type
- | UnBlock MetaId
- | IsEmpty Range Type
- | CheckSizeLtSat Term
- | FindInstance MetaId (Maybe [Candidate])
- | ResolveInstanceHead QName
- | CheckFunDef DefInfo QName [Clause] TCErr
- | UnquoteTactic Term Term Type
- | CheckLockedVars Term Type (Arg Term) Type
- | UsableAtModality WhyCheckModality (Maybe Sort) Modality Term
- data WarningsAndNonFatalErrors = WarningsAndNonFatalErrors {
- tcWarnings :: [TCWarning]
- nonFatalErrors :: [TCWarning]
- data Signature = Sig {}
- data Warning
- = NicifierIssue DeclarationWarning
- | TerminationIssue [TerminationError]
- | UnreachableClauses QName [Range]
- | CoverageIssue QName [(Telescope, [NamedArg DeBruijnPattern])]
- | CoverageNoExactSplit QName [Clause]
- | InlineNoExactSplit QName Clause
- | NotStrictlyPositive QName (Seq OccursWhere)
- | ConstructorDoesNotFitInData QName Sort Sort TCErr
- | CoinductiveEtaRecord QName
- | UnsolvedMetaVariables [Range]
- | UnsolvedInteractionMetas [Range]
- | UnsolvedConstraints Constraints
- | InteractionMetaBoundaries [Range]
- | CantGeneralizeOverSorts [MetaId]
- | AbsurdPatternRequiresNoRHS [NamedArg DeBruijnPattern]
- | OldBuiltin BuiltinId BuiltinId
- | BuiltinDeclaresIdentifier BuiltinId
- | DuplicateRecordDirective RecordDirective
- | EmptyRewritePragma
- | EmptyWhere
- | IllformedAsClause String
- | InvalidCharacterLiteral Char
- | ClashesViaRenaming NameOrModule [Name]
- | UselessPatternDeclarationForRecord String
- | UselessPragma Range Doc
- | UselessPublic
- | UselessHiding [ImportedName]
- | UselessInline QName
- | WrongInstanceDeclaration
- | InstanceWithExplicitArg QName
- | InstanceNoOutputTypeName Doc
- | InstanceArgWithExplicitArg Doc
- | InversionDepthReached QName
- | NoGuardednessFlag QName
- | SafeFlagPostulate Name
- | SafeFlagPragma [String]
- | SafeFlagWithoutKFlagPrimEraseEquality
- | WithoutKFlagPrimEraseEquality
- | ConflictingPragmaOptions String String
- | OptionWarning OptionWarning
- | ParseWarning ParseWarning
- | LibraryWarning LibWarning
- | DeprecationWarning String String String
- | UserWarning Text
- | DuplicateUsing (List1 ImportedName)
- | FixityInRenamingModule (List1 Range)
- | ModuleDoesntExport QName [Name] [Name] [ImportedName]
- | InfectiveImport Doc
- | CoInfectiveImport Doc
- | ConfluenceCheckingIncompleteBecauseOfMeta QName
- | ConfluenceForCubicalNotSupported
- | IllegalRewriteRule QName IllegalRewriteRuleReason
- | RewriteNonConfluent Term Term Term Doc
- | RewriteMaybeNonConfluent Term Term [Doc]
- | RewriteAmbiguousRules Term Term Term
- | RewriteMissingRule Term Term Term
- | PragmaCompileErased BackendName QName
- | PragmaCompileList
- | PragmaCompileMaybe
- | NoMain TopLevelModuleName
- | NotInScopeW [QName]
- | UnsupportedIndexedMatch Doc
- | AsPatternShadowsConstructorOrPatternSynonym LHSOrPatSyn
- | PatternShadowsConstructor Name QName
- | PlentyInHardCompileTimeMode QωOrigin
- | RecordFieldWarning RecordFieldWarning
- | MissingTypeSignatureForOpaque QName IsOpaque
- | NotAffectedByOpaque
- | UnfoldTransparentName QName
- | UselessOpaque
- | FaceConstraintCannotBeHidden ArgInfo
- | FaceConstraintCannotBeNamed NamedName
- | DuplicateInterfaceFiles AbsolutePath AbsolutePath
- | CustomBackendWarning String Doc
- inTopContext :: (MonadTCEnv tcm, ReadTCState tcm) => tcm a -> tcm a
- data TCWarning = TCWarning {}
- data NamedMeta = NamedMeta {}
- withInteractionId :: (MonadDebug m, MonadFail m, ReadTCState m, MonadError TCErr m, MonadTCEnv m, MonadTrace m) => InteractionId -> m a -> m a
- getInteractionRange :: (MonadInteractionPoints m, MonadFail m, MonadError TCErr m) => InteractionId -> m Range
- getMetaRange :: (HasCallStack, MonadDebug m, ReadTCState m) => MetaId -> m Range
- withMetaId :: (HasCallStack, MonadDebug m, MonadTCEnv m, MonadTrace m, ReadTCState m) => MetaId -> m a -> m a
- getIncludeDirs :: HasOptions m => m [AbsolutePath]
- libToTCM :: LibM a -> TCM a
- topLevelModuleName :: RawTopLevelModuleName -> TCM TopLevelModuleName
- type ModuleToSource = Map TopLevelModuleName AbsolutePath
- data ModuleInfo = ModuleInfo {
- miInterface :: Interface
- miWarnings :: [TCWarning]
- miPrimitive :: Bool
- miMode :: ModuleCheckMode
- pattern Constructor :: Int -> Int -> ConHead -> QName -> IsAbstract -> CompKit -> Maybe [QName] -> [IsForced] -> Maybe [Bool] -> Bool -> Bool -> Defn
- pattern Record :: Nat -> Maybe Clause -> ConHead -> Bool -> [Dom QName] -> Telescope -> Maybe [QName] -> EtaEquality -> PatternOrCopattern -> Maybe Induction -> Maybe Bool -> IsAbstract -> CompKit -> Defn
- pattern Datatype :: Nat -> Nat -> Maybe Clause -> [QName] -> Sort -> Maybe [QName] -> IsAbstract -> [QName] -> Maybe QName -> Maybe QName -> Defn
- class (MonadTCEnv m, ReadTCState m) => MonadTrace (m :: Type -> Type) where
- traceCall :: Call -> m a -> m a
- traceCallM :: m Call -> m a -> m a
- traceCallCPS :: Call -> ((a -> m b) -> m b) -> (a -> m b) -> m b
- traceClosureCall :: Closure Call -> m a -> m a
- printHighlightingInfo :: RemoveTokenBasedHighlighting -> HighlightingInfo -> m ()
- highlightAsTypeChecked :: MonadTrace m => Range -> Range -> m a -> m a
- type Definitions = HashMap QName Definition
- initState :: TCState
- setCommandLineOptions :: CommandLineOptions -> TCM ()
- withScope :: ReadTCState m => ScopeInfo -> m a -> m (a, ScopeInfo)
- type InteractionOutputCallback = Response_boot TCErr TCWarning WarningsAndNonFatalErrors -> TCM ()
- defaultInteractionOutputCallback :: InteractionOutputCallback
- data IPFace' t = IPFace' {}
- class (Functor m, Applicative m, Monad m) => HasOptions (m :: Type -> Type) where
- class Monad m => MonadTCEnv (m :: Type -> Type) where
- class (Applicative tcm, MonadIO tcm, MonadTCEnv tcm, MonadTCState tcm, HasOptions tcm) => MonadTCM (tcm :: Type -> Type) where
- class Monad m => MonadTCState (m :: Type -> Type) where
- class Monad m => ReadTCState (m :: Type -> Type) where
- getTCState :: m TCState
- locallyTCState :: Lens' TCState a -> (a -> a) -> m b -> m b
- withTCState :: (TCState -> TCState) -> m a -> m a
- mapTCMT :: (forall a1. m a1 -> n a1) -> TCMT m a -> TCMT n a
- data SomeBuiltin
- builtinById :: String -> Maybe BuiltinId
- isBuiltinNoDef :: BuiltinId -> Bool
- typeError :: (HasCallStack, MonadTCError m) => TypeError -> m a
- 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 (List1 QName)
- | DoesNotConstructAnElementOf QName Type
- | WrongHidingInLHS
- | WrongHidingInLambda Type
- | WrongHidingInApplication Type
- | WrongHidingInProjection QName
- | IllegalHidingInPostfixProjection (NamedArg Expr)
- | WrongNamedArgument (NamedArg Expr) [NamedName]
- | WrongIrrelevanceInLambda
- | WrongQuantityInLambda
- | WrongCohesionInLambda
- | QuantityMismatch Quantity Quantity
- | HidingMismatch Hiding Hiding
- | RelevanceMismatch Relevance Relevance
- | UninstantiatedDotPattern Expr
- | ForcedConstructorNotInstantiated Pattern
- | IllformedProjectionPatternAbstract Pattern
- | IllformedProjectionPatternConcrete Pattern
- | CannotEliminateWithPattern (Maybe Blocker) (NamedArg Pattern) Type
- | CannotEliminateWithProjection (Arg Type) Bool QName
- | WrongNumberOfConstructorArguments QName Nat Nat
- | ShouldBeEmpty Type [DeBruijnPattern]
- | ShouldBeASort Type
- | ShouldBePi Type
- | ShouldBePath Type
- | ShouldBeRecordType Type
- | ShouldBeRecordPattern DeBruijnPattern
- | NotAProjectionPattern (NamedArg Pattern)
- | NotAProperTerm
- | InvalidTypeSort Sort
- | InvalidType Term
- | SplitOnCoinductive
- | SplitOnIrrelevant (Dom Type)
- | SplitOnUnusableCohesion (Dom Type)
- | SplitOnNonVariable Term Type
- | SplitOnNonEtaRecord QName
- | SplitOnAbstract QName
- | SplitOnUnchecked QName
- | SplitOnPartial (Dom Type)
- | SplitInProp DataOrRecordE
- | DefinitionIsIrrelevant QName
- | DefinitionIsErased QName
- | ProjectionIsIrrelevant QName
- | VariableIsIrrelevant Name
- | VariableIsErased Name
- | VariableIsOfUnusableCohesion Name Cohesion
- | 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
- | UnequalFiniteness 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
- | SortOfSplitVarError (Maybe Blocker) Doc
- | BuiltinMustBeConstructor BuiltinId Expr
- | NoSuchBuiltinName String
- | DuplicateBuiltinBinding BuiltinId Term Term
- | NoBindingForBuiltin BuiltinId
- | NoBindingForPrimitive PrimitiveId
- | NoSuchPrimitiveFunction String
- | DuplicatePrimitiveBinding PrimitiveId QName QName
- | WrongArgInfoForPrimitive PrimitiveId ArgInfo ArgInfo
- | ShadowedModule Name [ModuleName]
- | BuiltinInParameterisedModule BuiltinId
- | IllegalDeclarationInDataDefinition [Declaration]
- | IllegalLetInTelescope TypedBinding
- | IllegalPatternInTelescope Binder
- | NoRHSRequiresAbsurdPattern [NamedArg Pattern]
- | TooManyFields QName [Name] [Name]
- | DuplicateFields [Name]
- | DuplicateConstructors [Name]
- | DuplicateOverlapPragma QName OverlapMode OverlapMode
- | WithOnFreeVariable Expr Term
- | UnexpectedWithPatterns [Pattern]
- | WithClausePatternMismatch Pattern (NamedArg DeBruijnPattern)
- | IllTypedPatternAfterWithAbstraction Pattern
- | FieldOutsideRecord
- | ModuleArityMismatch ModuleName Telescope [NamedArg Expr]
- | GeneralizeCyclicDependency
- | GeneralizeUnsolvedMeta
- | ReferencesFutureVariables Term (NonEmpty Int) (Arg Term) Int
- | DoesNotMentionTicks Term Type (Arg Term)
- | MismatchedProjectionsError QName QName
- | AttributeKindNotEnabled String String String
- | InvalidProjectionParameter (NamedArg Expr)
- | TacticAttributeNotAllowed
- | CannotRewriteByNonEquation Type
- | MacroResultTypeMismatch Type
- | NamedWhereModuleInRefinedContext [Term] [String]
- | CubicalPrimitiveNotFullyApplied QName
- | TooManyArgumentsToLeveledSort QName
- | TooManyArgumentsToUnivOmega QName
- | ComatchingDisabledForRecord QName
- | BuiltinMustBeIsOne Term
- | IncorrectTypeForRewriteRelation Term IncorrectTypeForRewriteRelationReason
- | UnexpectedParameter LamBinding
- | NoParameterOfName ArgName
- | UnexpectedModalityAnnotationInParameter LamBinding
- | ExpectedBindingForParameter (Dom Type) (Abs Type)
- | UnexpectedTypeSignatureForParameter (List1 (NamedArg Binder))
- | SortDoesNotAdmitDataDefinitions QName Sort
- | SortCannotDependOnItsIndex QName Type
- | UnusableAtModality WhyCheckModality Modality Term
- | SplitError SplitError
- | ImpossibleConstructor QName NegativeUnification
- | TooManyPolarities QName Int
- | RecursiveRecordNeedsInductivity QName
- | CannotSolveSizeConstraints (List1 (ProblemConstraint, HypSizeConstraint)) Doc
- | ContradictorySizeConstraint (ProblemConstraint, HypSizeConstraint)
- | EmptyTypeOfSizes Term
- | FunctionTypeInSizeUniv Term
- | LibraryError LibErrors
- | 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
- | InvalidFileName AbsolutePath InvalidFileNameReason
- | BothWithAndRHS
- | AbstractConstructorNotInScope QName
- | NotInScope [QName]
- | NoSuchModule QName
- | AmbiguousName QName AmbiguousNameReason
- | AmbiguousModule QName (List1 ModuleName)
- | AmbiguousField Name [ModuleName]
- | AmbiguousConstructor QName [QName]
- | ClashingDefinition QName QName (Maybe NiceDeclaration)
- | ClashingModule ModuleName ModuleName
- | ClashingImport Name QName
- | ClashingModuleImport Name ModuleName
- | DefinitionInDifferentModule QName
- | DuplicateImports QName [ImportedName]
- | InvalidPattern Pattern
- | RepeatedVariablesInPattern [Name]
- | GeneralizeNotSupportedHere QName
- | GeneralizedVarInLetOpenedModule QName
- | MultipleFixityDecls [(Name, [Fixity'])]
- | MultiplePolarityPragmas [Name]
- | NotAModuleExpr Expr
- | NotAnExpression Expr
- | NotAValidLetBinding NiceDeclaration
- | NotValidBeforeField NiceDeclaration
- | NothingAppliedToHiddenArg Expr
- | NothingAppliedToInstanceArg Expr
- | AsPatternInPatternSynonym
- | DotPatternInPatternSynonym
- | BadArgumentsToPatternSynonym AmbiguousQName
- | TooFewArgumentsToPatternSynonym AmbiguousQName
- | CannotResolveAmbiguousPatternSynonym (List1 (QName, PatternSynDefn))
- | IllegalInstanceVariableInPatternSynonym Name
- | PatternSynonymArgumentShadowsConstructorOrPatternSynonym LHSOrPatSyn Name (List1 AbstractName)
- | UnusedVariableInPatternSynonym Name
- | UnboundVariablesInPatternSynonym [Name]
- | NoParseForApplication (List2 Expr)
- | AmbiguousParseForApplication (List2 Expr) (List1 Expr)
- | NoParseForLHS LHSOrPatSyn [Pattern] Pattern
- | AmbiguousParseForLHS LHSOrPatSyn Pattern [Pattern]
- | AmbiguousProjection QName [QName]
- | AmbiguousOverloadedProjection (List1 QName) Doc
- | OperatorInformation [NotationSection] TypeError
- | InstanceNoCandidate Type [(Term, TCErr)]
- | UnquoteFailed UnquoteError
- | DeBruijnIndexOutOfScope Nat Telescope [Name]
- | NeedOptionCopatterns
- | NeedOptionRewriting
- | NeedOptionProp
- | NeedOptionTwoLevel
- | NonFatalErrors [TCWarning]
- | InstanceSearchDepthExhausted Term Type Int
- | TriedToCopyConstrainedPrim QName
- | CustomBackendError String Doc
- data LHSOrPatSyn
- getScope :: ReadTCState m => m ScopeInfo
- getContext :: MonadTCEnv m => m Context
- data Call
- = CheckClause Type SpineClause
- | CheckLHS SpineLHS
- | CheckPattern Pattern Telescope Type
- | CheckPatternLinearityType Name
- | CheckPatternLinearityValue Name
- | CheckLetBinding LetBinding
- | InferExpr Expr
- | CheckExprCall Comparison Expr Type
- | CheckDotPattern Expr Term
- | CheckProjection Range QName Type
- | IsTypeCall Comparison Expr Sort
- | IsType_ Expr
- | InferVar Name
- | InferDef QName
- | CheckArguments 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
- | CheckConArgFitsIn QName Bool Type Sort
- | CheckFunDefCall Range QName [Clause] Bool
- | CheckPragma Range Pragma
- | CheckPrimitive Range QName Expr
- | CheckIsEmpty Range Type
- | CheckConfluence QName QName
- | CheckModuleParameters ModuleName Telescope
- | CheckWithFunctionType Type
- | CheckSectionApplication Range Erased ModuleName ModuleApplication
- | CheckNamedWhere ModuleName
- | CheckIApplyConfluence Range QName Term Term Term Type
- | ScopeCheckExpr Expr
- | ScopeCheckDeclaration NiceDeclaration
- | ScopeCheckLHS QName Pattern
- | NoHighlighting
- | ModuleContents
- | SetRange Range
- class (Functor m, Applicative m, MonadFail m) => HasBuiltins (m :: Type -> Type) where
- getBuiltinThing :: SomeBuiltin -> m (Maybe (Builtin PrimFun))
- getBuiltinName' :: HasBuiltins m => BuiltinId -> m (Maybe QName)
- builtinProp :: BuiltinId
- builtinSet :: BuiltinId
- builtinStrictSet :: BuiltinId
- builtinPropOmega :: BuiltinId
- builtinSetOmega :: BuiltinId
- builtinSSetOmega :: BuiltinId
- getAllPatternSyns :: ReadTCState m => m PatternSynDefns
- type Context = [ContextEntry]
- setCurrentRange :: (MonadTrace m, HasRange x) => x -> m a -> m a
- notUnderOpaque :: MonadTCEnv m => m a -> m a
- registerInteractionPoint :: MonadInteractionPoints m => Bool -> Range -> Maybe Nat -> m InteractionId
- insideDotPattern :: TCM a -> TCM a
- isInsideDotPattern :: TCM Bool
- getCurrentPath :: MonadTCEnv m => m AbsolutePath
- data DisplayForm = Display {
- dfPatternVars :: Nat
- dfPats :: Elims
- dfRHS :: DisplayTerm
- constructorForm :: HasBuiltins m => Term -> m Term
- class (MonadTCEnv m, ReadTCState m, MonadError TCErr m, MonadBlock m, HasOptions m, MonadDebug m) => MonadConstraint (m :: Type -> Type) where
- addConstraint :: Blocker -> Constraint -> m ()
- addAwakeConstraint :: Blocker -> Constraint -> m ()
- solveConstraint :: Constraint -> m ()
- solveSomeAwakeConstraints :: (ProblemConstraint -> Bool) -> Bool -> m ()
- wakeConstraints :: (ProblemConstraint -> WakeUp) -> m ()
- stealConstraints :: ProblemId -> m ()
- modifyAwakeConstraints :: (Constraints -> Constraints) -> m ()
- modifySleepingConstraints :: (Constraints -> Constraints) -> m ()
- typeOfConst :: (HasConstInfo m, ReadTCState m) => QName -> m Type
- data Closure a = Closure {
- clSignature :: Signature
- clEnv :: TCEnv
- clScope :: ScopeInfo
- clModuleCheckpoints :: Map ModuleName CheckpointId
- clValue :: a
- sizeType :: (HasBuiltins m, MonadTCEnv m, ReadTCState m) => m Type
- isInstanceConstraint :: Constraint -> Bool
- shouldPostponeInstanceSearch :: (ReadTCState m, HasOptions m) => m Bool
- data RecordFieldWarning
- data BuiltinId
- = BuiltinNat
- | BuiltinSuc
- | BuiltinZero
- | BuiltinNatPlus
- | BuiltinNatMinus
- | BuiltinNatTimes
- | BuiltinNatDivSucAux
- | BuiltinNatModSucAux
- | BuiltinNatEquals
- | BuiltinNatLess
- | BuiltinWord64
- | BuiltinInteger
- | BuiltinIntegerPos
- | BuiltinIntegerNegSuc
- | BuiltinFloat
- | BuiltinChar
- | BuiltinString
- | BuiltinUnit
- | BuiltinUnitUnit
- | BuiltinSigma
- | BuiltinSigmaCon
- | BuiltinBool
- | BuiltinTrue
- | BuiltinFalse
- | BuiltinList
- | BuiltinNil
- | BuiltinCons
- | BuiltinMaybe
- | BuiltinNothing
- | BuiltinJust
- | BuiltinIO
- | BuiltinId
- | BuiltinReflId
- | BuiltinPath
- | BuiltinPathP
- | BuiltinIntervalUniv
- | BuiltinInterval
- | BuiltinIZero
- | BuiltinIOne
- | BuiltinPartial
- | BuiltinPartialP
- | BuiltinIsOne
- | BuiltinItIsOne
- | BuiltinEquiv
- | BuiltinEquivFun
- | BuiltinEquivProof
- | BuiltinTranspProof
- | BuiltinIsOne1
- | BuiltinIsOne2
- | BuiltinIsOneEmpty
- | BuiltinSub
- | BuiltinSubIn
- | BuiltinSizeUniv
- | BuiltinSize
- | BuiltinSizeLt
- | BuiltinSizeSuc
- | BuiltinSizeInf
- | BuiltinSizeMax
- | BuiltinInf
- | BuiltinSharp
- | BuiltinFlat
- | BuiltinEquality
- | BuiltinRefl
- | BuiltinRewrite
- | BuiltinLevelMax
- | BuiltinLevel
- | BuiltinLevelZero
- | BuiltinLevelSuc
- | BuiltinProp
- | BuiltinSet
- | BuiltinStrictSet
- | BuiltinPropOmega
- | BuiltinSetOmega
- | BuiltinSSetOmega
- | BuiltinLevelUniv
- | BuiltinFromNat
- | BuiltinFromNeg
- | BuiltinFromString
- | BuiltinQName
- | BuiltinAgdaSort
- | BuiltinAgdaSortSet
- | BuiltinAgdaSortLit
- | BuiltinAgdaSortProp
- | BuiltinAgdaSortPropLit
- | BuiltinAgdaSortInf
- | BuiltinAgdaSortUnsupported
- | BuiltinHiding
- | BuiltinHidden
- | BuiltinInstance
- | BuiltinVisible
- | BuiltinRelevance
- | BuiltinRelevant
- | BuiltinIrrelevant
- | BuiltinQuantity
- | BuiltinQuantity0
- | BuiltinQuantityω
- | BuiltinModality
- | BuiltinModalityConstructor
- | BuiltinAssoc
- | BuiltinAssocLeft
- | BuiltinAssocRight
- | BuiltinAssocNon
- | BuiltinPrecedence
- | BuiltinPrecRelated
- | BuiltinPrecUnrelated
- | BuiltinFixity
- | BuiltinFixityFixity
- | BuiltinArg
- | BuiltinArgInfo
- | BuiltinArgArgInfo
- | BuiltinArgArg
- | BuiltinAbs
- | BuiltinAbsAbs
- | BuiltinAgdaTerm
- | BuiltinAgdaTermVar
- | BuiltinAgdaTermLam
- | BuiltinAgdaTermExtLam
- | BuiltinAgdaTermDef
- | BuiltinAgdaTermCon
- | BuiltinAgdaTermPi
- | BuiltinAgdaTermSort
- | BuiltinAgdaTermLit
- | BuiltinAgdaTermUnsupported
- | BuiltinAgdaTermMeta
- | BuiltinAgdaErrorPart
- | BuiltinAgdaErrorPartString
- | BuiltinAgdaErrorPartTerm
- | BuiltinAgdaErrorPartPatt
- | BuiltinAgdaErrorPartName
- | BuiltinAgdaLiteral
- | BuiltinAgdaLitNat
- | BuiltinAgdaLitWord64
- | BuiltinAgdaLitFloat
- | BuiltinAgdaLitChar
- | BuiltinAgdaLitString
- | BuiltinAgdaLitQName
- | BuiltinAgdaLitMeta
- | BuiltinAgdaClause
- | BuiltinAgdaClauseClause
- | BuiltinAgdaClauseAbsurd
- | BuiltinAgdaPattern
- | BuiltinAgdaPatVar
- | BuiltinAgdaPatCon
- | BuiltinAgdaPatDot
- | BuiltinAgdaPatLit
- | BuiltinAgdaPatProj
- | BuiltinAgdaPatAbsurd
- | BuiltinAgdaDefinitionFunDef
- | BuiltinAgdaDefinitionDataDef
- | BuiltinAgdaDefinitionRecordDef
- | BuiltinAgdaDefinitionDataConstructor
- | BuiltinAgdaDefinitionPostulate
- | BuiltinAgdaDefinitionPrimitive
- | BuiltinAgdaDefinition
- | BuiltinAgdaMeta
- | BuiltinAgdaTCM
- | BuiltinAgdaTCMReturn
- | BuiltinAgdaTCMBind
- | BuiltinAgdaTCMUnify
- | BuiltinAgdaTCMTypeError
- | BuiltinAgdaTCMInferType
- | BuiltinAgdaTCMCheckType
- | BuiltinAgdaTCMNormalise
- | BuiltinAgdaTCMReduce
- | BuiltinAgdaTCMCatchError
- | BuiltinAgdaTCMGetContext
- | BuiltinAgdaTCMExtendContext
- | BuiltinAgdaTCMInContext
- | BuiltinAgdaTCMFreshName
- | BuiltinAgdaTCMDeclareDef
- | BuiltinAgdaTCMDeclarePostulate
- | BuiltinAgdaTCMDeclareData
- | BuiltinAgdaTCMDefineData
- | BuiltinAgdaTCMDefineFun
- | BuiltinAgdaTCMGetType
- | BuiltinAgdaTCMGetDefinition
- | BuiltinAgdaTCMBlock
- | BuiltinAgdaTCMCommit
- | BuiltinAgdaTCMQuoteTerm
- | BuiltinAgdaTCMUnquoteTerm
- | BuiltinAgdaTCMQuoteOmegaTerm
- | BuiltinAgdaTCMIsMacro
- | BuiltinAgdaTCMWithNormalisation
- | BuiltinAgdaTCMWithReconstructed
- | BuiltinAgdaTCMWithExpandLast
- | BuiltinAgdaTCMWithReduceDefs
- | BuiltinAgdaTCMAskNormalisation
- | BuiltinAgdaTCMAskReconstructed
- | BuiltinAgdaTCMAskExpandLast
- | BuiltinAgdaTCMAskReduceDefs
- | BuiltinAgdaTCMFormatErrorParts
- | BuiltinAgdaTCMDebugPrint
- | BuiltinAgdaTCMNoConstraints
- | BuiltinAgdaTCMWorkOnTypes
- | BuiltinAgdaTCMRunSpeculative
- | BuiltinAgdaTCMExec
- | BuiltinAgdaTCMGetInstances
- | BuiltinAgdaTCMSolveInstances
- | BuiltinAgdaTCMPragmaForeign
- | BuiltinAgdaTCMPragmaCompile
- | BuiltinAgdaBlocker
- | BuiltinAgdaBlockerAny
- | BuiltinAgdaBlockerAll
- | BuiltinAgdaBlockerMeta
- data PrimitiveId
- = PrimConId
- | PrimIdElim
- | PrimIMin
- | PrimIMax
- | PrimINeg
- | PrimPartial
- | PrimPartialP
- | PrimSubOut
- | PrimGlue
- | Prim_glue
- | Prim_unglue
- | Prim_glueU
- | Prim_unglueU
- | PrimFaceForall
- | PrimComp
- | PrimPOr
- | PrimTrans
- | PrimDepIMin
- | PrimIdFace
- | PrimIdPath
- | PrimHComp
- | PrimShowInteger
- | PrimNatPlus
- | PrimNatMinus
- | PrimNatTimes
- | PrimNatDivSucAux
- | PrimNatModSucAux
- | PrimNatEquality
- | PrimNatLess
- | PrimShowNat
- | PrimWord64FromNat
- | PrimWord64ToNat
- | PrimWord64ToNatInjective
- | PrimLevelZero
- | PrimLevelSuc
- | PrimLevelMax
- | PrimFloatEquality
- | PrimFloatInequality
- | PrimFloatLess
- | PrimFloatIsInfinite
- | PrimFloatIsNaN
- | PrimFloatIsNegativeZero
- | PrimFloatIsSafeInteger
- | PrimFloatToWord64
- | PrimFloatToWord64Injective
- | PrimNatToFloat
- | PrimIntToFloat
- | PrimFloatRound
- | PrimFloatFloor
- | PrimFloatCeiling
- | PrimFloatToRatio
- | PrimRatioToFloat
- | PrimFloatDecode
- | PrimFloatEncode
- | PrimShowFloat
- | PrimFloatPlus
- | PrimFloatMinus
- | PrimFloatTimes
- | PrimFloatNegate
- | PrimFloatDiv
- | PrimFloatPow
- | PrimFloatSqrt
- | PrimFloatExp
- | PrimFloatLog
- | PrimFloatSin
- | PrimFloatCos
- | PrimFloatTan
- | PrimFloatASin
- | PrimFloatACos
- | PrimFloatATan
- | PrimFloatATan2
- | PrimFloatSinh
- | PrimFloatCosh
- | PrimFloatTanh
- | PrimFloatASinh
- | PrimFloatACosh
- | PrimFloatATanh
- | PrimCharEquality
- | PrimIsLower
- | PrimIsDigit
- | PrimIsAlpha
- | PrimIsSpace
- | PrimIsAscii
- | PrimIsLatin1
- | PrimIsPrint
- | PrimIsHexDigit
- | PrimToUpper
- | PrimToLower
- | PrimCharToNat
- | PrimCharToNatInjective
- | PrimNatToChar
- | PrimShowChar
- | PrimStringToList
- | PrimStringToListInjective
- | PrimStringFromList
- | PrimStringFromListInjective
- | PrimStringAppend
- | PrimStringEquality
- | PrimShowString
- | PrimStringUncons
- | PrimErase
- | PrimEraseEquality
- | PrimForce
- | PrimForceLemma
- | PrimQNameEquality
- | PrimQNameLess
- | PrimShowQName
- | PrimQNameFixity
- | PrimQNameToWord64s
- | PrimQNameToWord64sInjective
- | PrimMetaEquality
- | PrimMetaLess
- | PrimShowMeta
- | PrimMetaToNat
- | PrimMetaToNatInjective
- | PrimLockUniv
- writeToCurrentLog :: (MonadDebug m, MonadTCState m, ReadTCState m) => TypeCheckAction -> m ()
- readFromCachedLog :: (MonadDebug m, MonadTCState m, ReadTCState m) => m (Maybe (TypeCheckAction, PostScopeState))
- cleanCachedLog :: (MonadDebug m, MonadTCState m) => m ()
- cacheCurrentLog :: (MonadDebug m, MonadTCState m) => m ()
- activateLoadedFileCache :: (HasOptions m, MonadDebug m, MonadTCState m) => m ()
- cachingStarts :: (MonadDebug m, MonadTCState m, ReadTCState m) => m ()
- areWeCaching :: ReadTCState m => m Bool
- localCache :: (MonadTCState m, ReadTCState m) => m a -> m a
- withoutCache :: (MonadTCState m, ReadTCState m) => m a -> m a
- restorePostScopeState :: (MonadDebug m, MonadTCState m) => PostScopeState -> m ()
- enterClosure :: (MonadTCEnv m, ReadTCState m, LensClosure c a) => c -> (a -> m b) -> m b
- solvingProblem :: MonadConstraint m => ProblemId -> m a -> m a
- type Verbosity = Maybe (Trie VerboseKeyItem VerboseLevel)
- type VerboseKey = String
- type VerboseLevel = Int
- currentModule :: MonadTCEnv m => m ModuleName
- addImport :: TopLevelModuleName -> TCM ()
- addImportCycleCheck :: TopLevelModuleName -> TCM a -> TCM a
- checkForImportCycle :: TCM ()
- dropDecodedModule :: TopLevelModuleName -> TCM ()
- getDecodedModule :: TopLevelModuleName -> TCM (Maybe ModuleInfo)
- getDecodedModules :: TCM DecodedModules
- getImportPath :: TCM [TopLevelModuleName]
- getPrettyVisitedModules :: ReadTCState m => m Doc
- getVisitedModule :: ReadTCState m => TopLevelModuleName -> m (Maybe ModuleInfo)
- getVisitedModules :: ReadTCState m => m VisitedModules
- setDecodedModules :: DecodedModules -> TCM ()
- setVisitedModules :: VisitedModules -> TCM ()
- storeDecodedModule :: ModuleInfo -> TCM ()
- visitModule :: ModuleInfo -> TCM ()
- withImportPath :: [TopLevelModuleName] -> TCM a -> TCM a
- reportSLn :: MonadDebug m => VerboseKey -> VerboseLevel -> String -> m ()
- hideAndRelParams :: (LensHiding a, LensRelevance a) => a -> a
- noMutualBlock :: TCM a -> TCM a
- makeOpen :: (ReadTCState m, MonadTCEnv m) => a -> m (Open a)
- getOpen :: (TermSubst a, MonadTCEnv m) => Open a -> m a
- tryGetOpen :: (TermSubst a, ReadTCState m, MonadTCEnv m) => (Substitution -> a -> Maybe a) -> Open a -> m (Maybe a)
- isClosed :: Open a -> Bool
- currentModuleNameHash :: ReadTCState m => m ModuleNameHash
- reportSDoc :: MonadDebug m => VerboseKey -> VerboseLevel -> TCM Doc -> m ()
- setPragmaOptions :: PragmaOptions -> TCM ()
- class MonadTCEnv m => MonadAddContext (m :: Type -> Type) where
- addCtx :: Name -> Dom Type -> m a -> m a
- addLetBinding' :: Origin -> Name -> Term -> Dom Type -> m a -> m a
- updateContext :: Substitution -> (Context -> Context) -> m a -> m a
- withFreshName :: Range -> ArgName -> (Name -> m a) -> m a
- class (Applicative m, MonadTCEnv m, ReadTCState m, HasOptions m) => MonadReduce (m :: Type -> Type) where
- liftReduce :: ReduceM a -> m a
- setHardCompileTimeModeIfErased :: Erased -> TCM a -> TCM a
- verboseS :: MonadDebug m => VerboseKey -> VerboseLevel -> m () -> m ()
- resetState :: TCM ()
- class ReadTCState m => MonadStatistics (m :: Type -> Type) where
- modifyCounter :: String -> (Integer -> Integer) -> m ()
- type Statistics = Map String Integer
- tick :: MonadStatistics m => String -> m ()
- tickN :: MonadStatistics m => String -> Integer -> m ()
- tickMax :: MonadStatistics m => String -> Integer -> m ()
- getStatistics :: ReadTCState m => m Statistics
- modifyStatistics :: (Statistics -> Statistics) -> TCM ()
- printStatistics :: (MonadDebug m, MonadTCEnv m, HasOptions m) => Maybe TopLevelModuleName -> Statistics -> m ()
- interestingCall :: Call -> Bool
- getName' :: (HasBuiltins m, IsBuiltin a) => a -> m (Maybe QName)
- builtinHComp :: PrimitiveId
- builtinConId :: PrimitiveId
- localTCState :: TCM a -> TCM a
- askR :: ReduceM ReduceEnv
- applyWhenVerboseS :: MonadDebug m => VerboseKey -> VerboseLevel -> (m a -> m a) -> m a -> m a
- defAbstract :: Definition -> IsAbstract
- data CompilerPragma = CompilerPragma Range String
- builtinSizeInf :: BuiltinId
- getBuiltin' :: HasBuiltins m => BuiltinId -> m (Maybe Term)
- unsafeModifyContext :: MonadTCEnv tcm => (Context -> Context) -> tcm a -> tcm a
- type Offset = Nat
- metaType :: ReadTCState m => MetaId -> m Type
- applyDef :: HasConstInfo m => ProjOrigin -> QName -> Arg Term -> m Term
- 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
- envUnquoteProblem :: Maybe ProblemId
- envAbstractMode :: AbstractMode
- envRelevance :: Relevance
- envQuantity :: Quantity
- envHardCompileTimeMode :: Bool
- envSplitOnStrict :: Bool
- envDisplayFormsEnabled :: Bool
- envFoldLetBindings :: Bool
- envRange :: Range
- envHighlightingRange :: Range
- envClause :: IPClause
- envCall :: Maybe (Closure Call)
- envHighlightingLevel :: HighlightingLevel
- envHighlightingMethod :: HighlightingMethod
- envExpandLast :: ExpandHidden
- envAppDef :: Maybe QName
- envSimplification :: Simplification
- envAllowedReductions :: AllowedReductions
- envReduceDefs :: ReduceDefs
- envReconstructed :: Bool
- envInjectivityDepth :: Int
- envCompareBlocked :: Bool
- envPrintDomainFreePi :: Bool
- envPrintMetasBare :: Bool
- envInsideDotPattern :: Bool
- envUnquoteFlags :: UnquoteFlags
- envInstanceDepth :: !Int
- envIsDebugPrinting :: Bool
- envPrintingPatternLambdas :: [QName]
- envCallByNeed :: Bool
- envCurrentCheckpoint :: CheckpointId
- envCheckpoints :: Map CheckpointId Substitution
- envGeneralizeMetas :: DoGeneralize
- envGeneralizedVars :: Map QName GeneralizedValue
- envActiveBackendName :: Maybe BackendName
- envConflComputingOverlap :: Bool
- envCurrentlyElaborating :: Bool
- envSyntacticEqualityFuel :: !(Maybe Int)
- envCurrentOpaqueId :: !(Maybe OpaqueId)
- getEnv :: TCM TCEnv
- data Interface = Interface {
- iSourceHash :: !Hash
- iSource :: Text
- iFileType :: FileType
- iImportedModules :: [(TopLevelModuleName, Hash)]
- iModuleName :: ModuleName
- iTopLevelModuleName :: TopLevelModuleName
- iScope :: Map ModuleName Scope
- iInsideScope :: ScopeInfo
- iSignature :: Signature
- iMetaBindings :: RemoteMetaStore
- iDisplayForms :: DisplayForms
- iUserWarnings :: Map QName Text
- iImportWarning :: Maybe Text
- iBuiltin :: BuiltinThings (PrimitiveId, QName)
- iForeignCode :: Map BackendName ForeignCodeStack
- iHighlighting :: HighlightingInfo
- iDefaultPragmaOptions :: [OptionsPragma]
- iFilePragmaOptions :: [OptionsPragma]
- iOptionsUsed :: PragmaOptions
- iPatternSyns :: PatternSynDefns
- iWarnings :: [TCWarning]
- iPartialDefs :: Set QName
- iOpaqueBlocks :: Map OpaqueId OpaqueBlock
- iOpaqueNames :: Map QName OpaqueId
- type MonadTCError (m :: Type -> Type) = (MonadTCEnv m, ReadTCState m, MonadError TCErr m)
- genericError :: (HasCallStack, MonadTCError m) => String -> m a
- getAgdaLibFiles :: AbsolutePath -> TopLevelModuleName -> TCM [AgdaLibFile]
- class ReportS a where
- reportS :: MonadDebug m => VerboseKey -> VerboseLevel -> a -> m ()
- stModuleToSource :: Lens' TCState ModuleToSource
- useTC :: ReadTCState m => Lens' TCState a -> m a
- type BackendName = String
- data TCState = TCSt {}
- class (MonadTCEnv m, ReadTCState m) => MonadInteractionPoints (m :: Type -> Type) where
- freshInteractionId :: m InteractionId
- modifyInteractionPoints :: (InteractionPoints -> InteractionPoints) -> m ()
- class (HasBuiltins m, HasConstInfo m, MonadAddContext m, MonadDebug m, MonadReduce m, MonadTCEnv m, ReadTCState m) => PureTCM (m :: Type -> Type)
- type RewriteRules = [RewriteRule]
- data SigError
- = SigUnknown String
- | SigAbstract
- | SigCubicalNotErasure
- checkpointSubstitution :: MonadTCEnv tcm => CheckpointId -> tcm Substitution
- data PrimFun = PrimFun {
- primFunName :: QName
- primFunArity :: Arity
- primFunArgOccurrences :: [Occurrence]
- primFunImplementation :: [Arg Term] -> Int -> ReduceM (Reduced MaybeReducedArgs Term)
- data RunMetaOccursCheck
- data CompareAs
- data CompareDirection
- data MetaVariable = MetaVar {}
- class (MonadConstraint m, MonadReduce m, MonadAddContext m, MonadTCEnv m, ReadTCState m, HasBuiltins m, HasConstInfo m, MonadDebug m) => MonadMetaSolver (m :: Type -> Type) where
- newMeta' :: MetaInstantiation -> Frozen -> MetaInfo -> MetaPriority -> Permutation -> Judgement a -> m MetaId
- assignV :: CompareDirection -> MetaId -> Args -> Term -> CompareAs -> m ()
- assignTerm' :: MetaId -> [Arg ArgName] -> Term -> m ()
- etaExpandMeta :: [MetaClass] -> MetaId -> m ()
- updateMetaVar :: MetaId -> (MetaVariable -> MetaVariable) -> m ()
- speculateMetas :: m () -> m KeepMetas -> m ()
- data InstanceInfo = InstanceInfo {}
- class Monad m => MonadBlock (m :: Type -> Type) where
- patternViolation :: Blocker -> m a
- catchPatternErr :: (Blocker -> m a) -> m a -> m a
- data Reduced no yes
- = NoReduction no
- | YesReduction Simplification yes
- data IsReduced
- = NotReduced
- | Reduced (Blocked ())
- data ProblemConstraint = PConstr {}
- data NLPat
- class Monad m => MonadFresh i (m :: Type -> Type) where
- fresh :: m i
- modifyBenchmark :: (Benchmark -> Benchmark) -> TCM ()
- class IsBuiltin a where
- someBuiltin :: a -> SomeBuiltin
- getBuiltinId :: a -> String
- builtinsNoDef :: [BuiltinId]
- sizeBuiltins :: [BuiltinId]
- builtinIntervalUniv :: BuiltinId
- builtinId :: BuiltinId
- builtinReflId :: BuiltinId
- builtinInterval :: BuiltinId
- builtinPartial :: BuiltinId
- builtinPartialP :: BuiltinId
- builtinIsOne :: BuiltinId
- builtinSub :: BuiltinId
- builtinIZero :: BuiltinId
- builtinIOne :: BuiltinId
- builtinLevelUniv :: BuiltinId
- builtinSizeUniv :: BuiltinId
- builtinSize :: BuiltinId
- builtinSizeLt :: BuiltinId
- builtinSizeSuc :: BuiltinId
- builtinSizeMax :: BuiltinId
- builtinNat :: BuiltinId
- builtinSuc :: BuiltinId
- builtinZero :: BuiltinId
- builtinNatPlus :: BuiltinId
- builtinNatMinus :: BuiltinId
- builtinNatTimes :: BuiltinId
- builtinNatDivSucAux :: BuiltinId
- builtinNatModSucAux :: BuiltinId
- builtinNatEquals :: BuiltinId
- builtinNatLess :: BuiltinId
- builtinInteger :: BuiltinId
- builtinIntegerPos :: BuiltinId
- builtinIntegerNegSuc :: BuiltinId
- builtinWord64 :: BuiltinId
- builtinFloat :: BuiltinId
- builtinChar :: BuiltinId
- builtinString :: BuiltinId
- builtinUnit :: BuiltinId
- builtinUnitUnit :: BuiltinId
- builtinSigma :: BuiltinId
- builtinBool :: BuiltinId
- builtinTrue :: BuiltinId
- builtinFalse :: BuiltinId
- builtinList :: BuiltinId
- builtinNil :: BuiltinId
- builtinCons :: BuiltinId
- builtinIO :: BuiltinId
- builtinMaybe :: BuiltinId
- builtinNothing :: BuiltinId
- builtinJust :: BuiltinId
- builtinPath :: BuiltinId
- builtinPathP :: BuiltinId
- builtinItIsOne :: BuiltinId
- builtinIsOne1 :: BuiltinId
- builtinIsOne2 :: BuiltinId
- builtinIsOneEmpty :: BuiltinId
- builtinSubIn :: BuiltinId
- builtinEquiv :: BuiltinId
- builtinEquivFun :: BuiltinId
- builtinEquivProof :: BuiltinId
- builtinTranspProof :: BuiltinId
- builtinInf :: BuiltinId
- builtinSharp :: BuiltinId
- builtinFlat :: BuiltinId
- builtinEquality :: BuiltinId
- builtinRefl :: BuiltinId
- builtinRewrite :: BuiltinId
- builtinLevelMax :: BuiltinId
- builtinLevel :: BuiltinId
- builtinLevelZero :: BuiltinId
- builtinLevelSuc :: BuiltinId
- builtinFromNat :: BuiltinId
- builtinFromNeg :: BuiltinId
- builtinFromString :: BuiltinId
- builtinQName :: BuiltinId
- builtinAgdaSort :: BuiltinId
- builtinAgdaSortSet :: BuiltinId
- builtinAgdaSortLit :: BuiltinId
- builtinAgdaSortProp :: BuiltinId
- builtinAgdaSortPropLit :: BuiltinId
- builtinAgdaSortInf :: BuiltinId
- builtinAgdaSortUnsupported :: BuiltinId
- builtinHiding :: BuiltinId
- builtinHidden :: BuiltinId
- builtinInstance :: BuiltinId
- builtinVisible :: BuiltinId
- builtinRelevance :: BuiltinId
- builtinRelevant :: BuiltinId
- builtinIrrelevant :: BuiltinId
- builtinQuantity :: BuiltinId
- builtinQuantity0 :: BuiltinId
- builtinQuantityω :: BuiltinId
- builtinModality :: BuiltinId
- builtinModalityConstructor :: BuiltinId
- builtinAssoc :: BuiltinId
- builtinAssocLeft :: BuiltinId
- builtinAssocRight :: BuiltinId
- builtinAssocNon :: BuiltinId
- builtinPrecedence :: BuiltinId
- builtinPrecRelated :: BuiltinId
- builtinPrecUnrelated :: BuiltinId
- builtinFixity :: BuiltinId
- builtinFixityFixity :: BuiltinId
- builtinArgInfo :: BuiltinId
- builtinArgArgInfo :: BuiltinId
- builtinArg :: BuiltinId
- builtinArgArg :: BuiltinId
- builtinAbs :: BuiltinId
- builtinAbsAbs :: BuiltinId
- builtinAgdaTerm :: BuiltinId
- builtinAgdaTermVar :: BuiltinId
- builtinAgdaTermLam :: BuiltinId
- builtinAgdaTermExtLam :: BuiltinId
- builtinAgdaTermDef :: BuiltinId
- builtinAgdaTermCon :: BuiltinId
- builtinAgdaTermPi :: BuiltinId
- builtinAgdaTermSort :: BuiltinId
- builtinAgdaTermLit :: BuiltinId
- builtinAgdaTermUnsupported :: BuiltinId
- builtinAgdaTermMeta :: BuiltinId
- builtinAgdaErrorPart :: BuiltinId
- builtinAgdaErrorPartString :: BuiltinId
- builtinAgdaErrorPartTerm :: BuiltinId
- builtinAgdaErrorPartPatt :: BuiltinId
- builtinAgdaErrorPartName :: BuiltinId
- builtinAgdaLiteral :: BuiltinId
- builtinAgdaLitNat :: BuiltinId
- builtinAgdaLitWord64 :: BuiltinId
- builtinAgdaLitFloat :: BuiltinId
- builtinAgdaLitChar :: BuiltinId
- builtinAgdaLitString :: BuiltinId
- builtinAgdaLitQName :: BuiltinId
- builtinAgdaLitMeta :: BuiltinId
- builtinAgdaClause :: BuiltinId
- builtinAgdaClauseClause :: BuiltinId
- builtinAgdaClauseAbsurd :: BuiltinId
- builtinAgdaPattern :: BuiltinId
- builtinAgdaPatVar :: BuiltinId
- builtinAgdaPatCon :: BuiltinId
- builtinAgdaPatDot :: BuiltinId
- builtinAgdaPatLit :: BuiltinId
- builtinAgdaPatProj :: BuiltinId
- builtinAgdaPatAbsurd :: BuiltinId
- builtinAgdaDefinitionFunDef :: BuiltinId
- builtinAgdaDefinitionDataDef :: BuiltinId
- builtinAgdaDefinitionRecordDef :: BuiltinId
- builtinAgdaDefinitionDataConstructor :: BuiltinId
- builtinAgdaDefinitionPostulate :: BuiltinId
- builtinAgdaDefinitionPrimitive :: BuiltinId
- builtinAgdaDefinition :: BuiltinId
- builtinAgdaMeta :: BuiltinId
- builtinAgdaTCM :: BuiltinId
- builtinAgdaTCMReturn :: BuiltinId
- builtinAgdaTCMBind :: BuiltinId
- builtinAgdaTCMUnify :: BuiltinId
- builtinAgdaTCMTypeError :: BuiltinId
- builtinAgdaTCMInferType :: BuiltinId
- builtinAgdaTCMCheckType :: BuiltinId
- builtinAgdaTCMNormalise :: BuiltinId
- builtinAgdaTCMReduce :: BuiltinId
- builtinAgdaTCMCatchError :: BuiltinId
- builtinAgdaTCMGetContext :: BuiltinId
- builtinAgdaTCMExtendContext :: BuiltinId
- builtinAgdaTCMInContext :: BuiltinId
- builtinAgdaTCMFreshName :: BuiltinId
- builtinAgdaTCMDeclareDef :: BuiltinId
- builtinAgdaTCMDeclarePostulate :: BuiltinId
- builtinAgdaTCMDeclareData :: BuiltinId
- builtinAgdaTCMDefineData :: BuiltinId
- builtinAgdaTCMDefineFun :: BuiltinId
- builtinAgdaTCMGetType :: BuiltinId
- builtinAgdaTCMGetDefinition :: BuiltinId
- builtinAgdaTCMQuoteTerm :: BuiltinId
- builtinAgdaTCMUnquoteTerm :: BuiltinId
- builtinAgdaTCMQuoteOmegaTerm :: BuiltinId
- builtinAgdaTCMCommit :: BuiltinId
- builtinAgdaTCMIsMacro :: BuiltinId
- builtinAgdaTCMBlock :: BuiltinId
- builtinAgdaBlocker :: BuiltinId
- builtinAgdaBlockerAll :: BuiltinId
- builtinAgdaBlockerAny :: BuiltinId
- builtinAgdaBlockerMeta :: BuiltinId
- builtinAgdaTCMFormatErrorParts :: BuiltinId
- builtinAgdaTCMDebugPrint :: BuiltinId
- builtinAgdaTCMWithNormalisation :: BuiltinId
- builtinAgdaTCMWithReconstructed :: BuiltinId
- builtinAgdaTCMWithExpandLast :: BuiltinId
- builtinAgdaTCMWithReduceDefs :: BuiltinId
- builtinAgdaTCMAskNormalisation :: BuiltinId
- builtinAgdaTCMAskReconstructed :: BuiltinId
- builtinAgdaTCMAskExpandLast :: BuiltinId
- builtinAgdaTCMAskReduceDefs :: BuiltinId
- builtinAgdaTCMNoConstraints :: BuiltinId
- builtinAgdaTCMWorkOnTypes :: BuiltinId
- builtinAgdaTCMRunSpeculative :: BuiltinId
- builtinAgdaTCMExec :: BuiltinId
- builtinAgdaTCMGetInstances :: BuiltinId
- builtinAgdaTCMSolveInstances :: BuiltinId
- builtinAgdaTCMPragmaForeign :: BuiltinId
- builtinAgdaTCMPragmaCompile :: BuiltinId
- builtinIdElim :: PrimitiveId
- builtinSubOut :: PrimitiveId
- builtinIMin :: PrimitiveId
- builtinIMax :: PrimitiveId
- builtinINeg :: PrimitiveId
- builtinGlue :: PrimitiveId
- builtin_glue :: PrimitiveId
- builtin_unglue :: PrimitiveId
- builtin_glueU :: PrimitiveId
- builtin_unglueU :: PrimitiveId
- builtinFaceForall :: PrimitiveId
- builtinComp :: PrimitiveId
- builtinPOr :: PrimitiveId
- builtinTrans :: PrimitiveId
- builtinDepIMin :: PrimitiveId
- builtinIdFace :: PrimitiveId
- builtinIdPath :: PrimitiveId
- builtinLockUniv :: PrimitiveId
- primitiveById :: String -> Maybe PrimitiveId
- primConId :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primIdElim :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primIMin :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primIMax :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primINeg :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primPartial :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primPartialP :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primSubOut :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primGlue :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primFaceForall :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primHComp :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primNatPlus :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primNatMinus :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primNatTimes :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primNatDivSucAux :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primNatModSucAux :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primNatEquality :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primNatLess :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primLevelZero :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primLevelSuc :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primLevelMax :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primLockUniv :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- type ContextEntry = Dom (Name, Type)
- getPolarity :: HasConstInfo m => QName -> m [Polarity]
- pattern Axiom :: Bool -> Defn
- lookupMeta :: ReadTCState m => MetaId -> m (Maybe (Either RemoteMetaVariable MetaVariable))
- isLocal :: ReadTCState m => QName -> m Bool
- newtype Fields = Fields [(Name, Type)]
- data ForeignCode = ForeignCode Range String
- defOpaque :: Definition -> IsOpaque
- data LetBinding = LetBinding {}
- newtype Section = Section {}
- data PreScopeState = PreScopeState {
- stPreTokens :: !HighlightingInfo
- stPreImports :: !Signature
- stPreImportedModules :: !(Set TopLevelModuleName)
- stPreModuleToSource :: !ModuleToSource
- stPreVisitedModules :: !VisitedModules
- stPreScope :: !ScopeInfo
- stPrePatternSyns :: !PatternSynDefns
- stPrePatternSynImports :: !PatternSynDefns
- stPreGeneralizedVars :: !(Maybe (Set QName))
- stPrePragmaOptions :: !PragmaOptions
- stPreImportedBuiltins :: !(BuiltinThings PrimFun)
- stPreImportedDisplayForms :: !DisplayForms
- stPreFreshInteractionId :: !InteractionId
- stPreImportedUserWarnings :: !(Map QName Text)
- stPreLocalUserWarnings :: !(Map QName Text)
- stPreWarningOnImport :: !(Maybe Text)
- stPreImportedPartialDefs :: !(Set QName)
- stPreProjectConfigs :: !(Map FilePath ProjectConfig)
- stPreAgdaLibFiles :: !(Map FilePath AgdaLibFile)
- stPreImportedMetaStore :: !RemoteMetaStore
- stPreCopiedNames :: !(HashMap QName QName)
- stPreNameCopies :: !(HashMap QName (HashSet QName))
- data PostScopeState = PostScopeState {
- stPostSyntaxInfo :: !HighlightingInfo
- stPostDisambiguatedNames :: !DisambiguatedNames
- stPostOpenMetaStore :: !LocalMetaStore
- stPostSolvedMetaStore :: !LocalMetaStore
- stPostInteractionPoints :: !InteractionPoints
- stPostAwakeConstraints :: !Constraints
- stPostSleepingConstraints :: !Constraints
- stPostDirty :: !Bool
- stPostOccursCheckDefs :: !(Set QName)
- stPostSignature :: !Signature
- stPostModuleCheckpoints :: !(Map ModuleName CheckpointId)
- stPostImportsDisplayForms :: !DisplayForms
- stPostForeignCode :: !(Map BackendName ForeignCodeStack)
- stPostCurrentModule :: !(Maybe (ModuleName, TopLevelModuleName))
- stPostPendingInstances :: !(Set QName)
- stPostTemporaryInstances :: !(Set QName)
- stPostConcreteNames :: !ConcreteNames
- stPostUsedNames :: !(Map RawName (DList RawName))
- stPostShadowingNames :: !(Map Name (DList RawName))
- stPostStatistics :: !Statistics
- stPostTCWarnings :: ![TCWarning]
- stPostMutualBlocks :: !(Map MutualId MutualBlock)
- stPostLocalBuiltins :: !(BuiltinThings PrimFun)
- stPostFreshMetaId :: !MetaId
- stPostFreshMutualId :: !MutualId
- stPostFreshProblemId :: !ProblemId
- stPostFreshCheckpointId :: !CheckpointId
- stPostFreshInt :: !Int
- stPostFreshNameId :: !NameId
- stPostFreshOpaqueId :: !OpaqueId
- stPostAreWeCaching :: !Bool
- stPostPostponeInstanceSearch :: !Bool
- stPostConsideringInstance :: !Bool
- stPostInstantiateBlocking :: !Bool
- stPostLocalPartialDefs :: !(Set QName)
- stPostOpaqueBlocks :: Map OpaqueId OpaqueBlock
- stPostOpaqueIds :: Map QName OpaqueId
- data PersistentTCState = PersistentTCSt {
- stDecodedModules :: !DecodedModules
- stPersistentTopLevelModuleNames :: !(BiMap RawTopLevelModuleName ModuleNameHash)
- stPersistentOptions :: CommandLineOptions
- stInteractionOutputCallback :: InteractionOutputCallback
- stBenchmark :: !Benchmark
- stAccumStatistics :: !Statistics
- stPersistLoadedFileCache :: !(Maybe LoadedFileCache)
- stPersistBackends :: [Backend_boot TCM]
- type VisitedModules = Map TopLevelModuleName ModuleInfo
- type BuiltinThings pf = Map SomeBuiltin (Builtin pf)
- type DisplayForms = HashMap QName [LocalDisplayForm]
- type RemoteMetaStore = HashMap MetaId RemoteMetaVariable
- data DisambiguatedName = DisambiguatedName NameKind QName
- type DisambiguatedNames = IntMap DisambiguatedName
- type ConcreteNames = Map Name [Name]
- type LocalMetaStore = Map MetaId MetaVariable
- type InteractionPoints = BiMap InteractionId InteractionPoint
- newtype CheckpointId = CheckpointId Int
- newtype ForeignCodeStack = ForeignCodeStack {}
- newtype MutualId = MutId Int32
- data MutualBlock = MutualBlock {
- mutualInfo :: MutualInfo
- mutualNames :: Set QName
- data OpaqueBlock = OpaqueBlock {
- opaqueId :: !OpaqueId
- opaqueUnfolding :: HashSet QName
- opaqueDecls :: HashSet QName
- opaqueParent :: Maybe OpaqueId
- opaqueRange :: Range
- type DecodedModules = Map TopLevelModuleName ModuleInfo
- data LoadedFileCache = LoadedFileCache {}
- type CachedTypeCheckLog = [(TypeCheckAction, PostScopeState)]
- type CurrentTypeCheckLog = [(TypeCheckAction, PostScopeState)]
- initPersistentState :: PersistentTCState
- initialMetaId :: MetaId
- initPreScopeState :: PreScopeState
- emptySignature :: Signature
- initPostScopeState :: PostScopeState
- stTokens :: Lens' TCState HighlightingInfo
- stImports :: Lens' TCState Signature
- stImportedModules :: Lens' TCState (Set TopLevelModuleName)
- stVisitedModules :: Lens' TCState VisitedModules
- stScope :: Lens' TCState ScopeInfo
- stPatternSyns :: Lens' TCState PatternSynDefns
- stPatternSynImports :: Lens' TCState PatternSynDefns
- stGeneralizedVars :: Lens' TCState (Maybe (Set QName))
- stPragmaOptions :: Lens' TCState PragmaOptions
- stImportedBuiltins :: Lens' TCState (BuiltinThings PrimFun)
- stForeignCode :: Lens' TCState (Map BackendName ForeignCodeStack)
- stFreshInteractionId :: Lens' TCState InteractionId
- stImportedUserWarnings :: Lens' TCState (Map QName Text)
- stLocalUserWarnings :: Lens' TCState (Map QName Text)
- getUserWarnings :: ReadTCState m => m (Map QName Text)
- useR :: ReadTCState m => Lens' TCState a -> m a
- stWarningOnImport :: Lens' TCState (Maybe Text)
- stImportedPartialDefs :: Lens' TCState (Set QName)
- stLocalPartialDefs :: Lens' TCState (Set QName)
- getPartialDefs :: ReadTCState m => m (Set QName)
- stLoadedFileCache :: Lens' TCState (Maybe LoadedFileCache)
- stBackends :: Lens' TCState [Backend_boot TCM]
- stProjectConfigs :: Lens' TCState (Map FilePath ProjectConfig)
- stAgdaLibFiles :: Lens' TCState (Map FilePath AgdaLibFile)
- stTopLevelModuleNames :: Lens' TCState (BiMap RawTopLevelModuleName ModuleNameHash)
- stImportedMetaStore :: Lens' TCState RemoteMetaStore
- stCopiedNames :: Lens' TCState (HashMap QName QName)
- stNameCopies :: Lens' TCState (HashMap QName (HashSet QName))
- stFreshNameId :: Lens' TCState NameId
- stFreshOpaqueId :: Lens' TCState OpaqueId
- stOpaqueBlocks :: Lens' TCState (Map OpaqueId OpaqueBlock)
- stOpaqueIds :: Lens' TCState (Map QName OpaqueId)
- stSyntaxInfo :: Lens' TCState HighlightingInfo
- stDisambiguatedNames :: Lens' TCState DisambiguatedNames
- stOpenMetaStore :: Lens' TCState LocalMetaStore
- stSolvedMetaStore :: Lens' TCState LocalMetaStore
- stInteractionPoints :: Lens' TCState InteractionPoints
- stAwakeConstraints :: Lens' TCState Constraints
- stSleepingConstraints :: Lens' TCState Constraints
- stDirty :: Lens' TCState Bool
- stOccursCheckDefs :: Lens' TCState (Set QName)
- stSignature :: Lens' TCState Signature
- stModuleCheckpoints :: Lens' TCState (Map ModuleName CheckpointId)
- stImportsDisplayForms :: Lens' TCState DisplayForms
- stImportedDisplayForms :: Lens' TCState DisplayForms
- stCurrentModule :: Lens' TCState (Maybe (ModuleName, TopLevelModuleName))
- stInstanceDefs :: Lens' TCState TempInstanceTable
- type TempInstanceTable = (InstanceTable, Set QName)
- sigInstances :: Lens' Signature InstanceTable
- stTemporaryInstances :: Lens' TCState (Set QName)
- stInstanceTree :: Lens' TCState (DiscrimTree QName)
- itableTree :: Lens' InstanceTable (DiscrimTree QName)
- stConcreteNames :: Lens' TCState ConcreteNames
- stUsedNames :: Lens' TCState (Map RawName (DList RawName))
- stShadowingNames :: Lens' TCState (Map Name (DList RawName))
- stStatistics :: Lens' TCState Statistics
- stTCWarnings :: Lens' TCState [TCWarning]
- stMutualBlocks :: Lens' TCState (Map MutualId MutualBlock)
- stLocalBuiltins :: Lens' TCState (BuiltinThings PrimFun)
- stFreshMetaId :: Lens' TCState MetaId
- stFreshMutualId :: Lens' TCState MutualId
- stFreshProblemId :: Lens' TCState ProblemId
- stFreshCheckpointId :: Lens' TCState CheckpointId
- stFreshInt :: Lens' TCState Int
- stAreWeCaching :: Lens' TCState Bool
- stPostponeInstanceSearch :: Lens' TCState Bool
- stConsideringInstance :: Lens' TCState Bool
- stInstantiateBlocking :: Lens' TCState Bool
- stBuiltinThings :: TCState -> BuiltinThings PrimFun
- unionBuiltin :: Builtin a -> Builtin a -> Builtin a
- class Enum i => HasFresh i where
- freshLens :: Lens' TCState i
- nextFresh' :: i -> i
- nextFresh :: HasFresh i => TCState -> (i, TCState)
- 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
- class Monad m => MonadStConcreteNames (m :: Type -> Type) where
- runStConcreteNames :: StateT ConcreteNames m a -> m a
- useConcreteNames :: m ConcreteNames
- modifyConcreteNames :: (ConcreteNames -> ConcreteNames) -> m ()
- stateTCLensM :: MonadTCState m => Lens' TCState a -> (a -> m (r, a)) -> m r
- data ModuleCheckMode
- iFullHash :: Interface -> Hash
- intSignature :: Lens' Interface Signature
- class LensClosure b a | b -> a where
- lensClosure :: Lens' b (Closure a)
- class LensTCEnv a where
- buildClosure :: (MonadTCEnv m, ReadTCState m) => a -> m (Closure a)
- data WhyCheckModality
- data IsForced
- data Candidate = Candidate {}
- fromCmp :: Comparison -> CompareDirection
- flipCmp :: CompareDirection -> CompareDirection
- dirToCmp :: (Comparison -> a -> a -> c) -> CompareDirection -> a -> a -> c
- data Judgement a
- = HasType {
- jMetaId :: a
- jComparison :: Comparison
- jMetaType :: Type
- | IsSort { }
- = HasType {
- data DoGeneralize
- data GeneralizedValue = GeneralizedValue {}
- newtype MetaPriority = MetaPriority Int
- data MetaInstantiation
- data Listener
- data Frozen
- data Instantiation = Instantiation {}
- data TypeCheckingProblem
- = CheckExpr Comparison Expr Type
- | CheckArgs Comparison ExpandHidden Range [NamedArg Expr] Type Type (ArgsCheckState CheckedTarget -> TCM Term)
- | CheckProjAppToKnownPrincipalArg Comparison Expr ProjOrigin (List1 QName) Args Type Int Term Type PrincipalArgTypeMetas
- | CheckLambda Comparison (Arg (List1 (WithHiding Name), Maybe Type)) Expr Type
- | DoQuoteTerm Comparison Term Type
- data RemoteMetaVariable = RemoteMetaVariable {}
- data CheckedTarget
- data PrincipalArgTypeMetas = PrincipalArgTypeMetas {
- patmMetas :: Args
- patmRemainder :: Type
- data ExpandHidden
- data ArgsCheckState a = ACState {}
- suffixNameSuggestion :: MetaNameSuggestion -> ArgName -> MetaNameSuggestion
- getMetaInfo :: MetaVariable -> Closure Range
- normalMetaPriority :: MetaPriority
- lowMetaPriority :: MetaPriority
- highMetaPriority :: MetaPriority
- getMetaScope :: MetaVariable -> ScopeInfo
- getMetaEnv :: MetaVariable -> TCEnv
- getMetaSig :: MetaVariable -> Signature
- metaFrozen :: Lens' MetaVariable Frozen
- _mvInfo :: Lens' MetaVariable MetaInfo
- aModeToDef :: AbstractMode -> Maybe IsAbstract
- aDefToMode :: IsAbstract -> AbstractMode
- data InteractionPoint = InteractionPoint {
- ipRange :: Range
- ipMeta :: Maybe MetaId
- ipSolved :: Bool
- ipClause :: IPClause
- ipBoundary :: IPBoundary
- data IPClause
- = IPClause {
- ipcQName :: QName
- ipcClauseNo :: Int
- ipcType :: Type
- ipcWithSub :: Maybe Substitution
- ipcClause :: SpineClause
- ipcClosure :: Closure ()
- | IPNoClause
- = IPClause {
- type IPBoundary = IPBoundary' Term
- newtype IPBoundary' t = IPBoundary {
- getBoundary :: Map (IntMap Bool) t
- data Overapplied
- type Sections = Map ModuleName Section
- type RewriteRuleMap = HashMap QName RewriteRules
- data InstanceTable = InstanceTable {
- _itableTree :: DiscrimTree QName
- _itableCounts :: Map QName Int
- sigSections :: Lens' Signature Sections
- sigDefinitions :: Lens' Signature Definitions
- sigRewriteRules :: Lens' Signature RewriteRuleMap
- type LocalDisplayForm = Open DisplayForm
- secTelescope :: Lens' Section Telescope
- data DisplayTerm
- pattern DDot :: Term -> DisplayTerm
- pattern DTerm :: Term -> DisplayTerm
- defaultDisplayForm :: QName -> [LocalDisplayForm]
- type PElims = [Elim' NLPat]
- data NLPType = NLPType {}
- data NLPSort
- = PUniv Univ NLPat
- | PInf Univ Integer
- | PSizeUniv
- | PLockUniv
- | PLevelUniv
- | PIntervalUniv
- pattern PType :: NLPat -> NLPSort
- pattern PProp :: NLPat -> NLPSort
- pattern PSSet :: NLPat -> NLPSort
- data RewriteRule = RewriteRule {}
- data Defn
- data NumGeneralizableArgs
- type CompiledRepresentation = Map BackendName [CompilerPragma]
- lensTheDef :: Lens' Definition Defn
- defaultDefn :: ArgInfo -> QName -> Type -> Language -> Defn -> Definition
- noCompiledRep :: CompiledRepresentation
- jsBackendName :: BackendName
- ghcBackendName :: BackendName
- type Face = [(Term, Bool)]
- data ExtLamInfo = ExtLamInfo {
- extLamModule :: ModuleName
- extLamAbsurd :: Bool
- 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
- defaultAxiom :: Defn
- constTranspAxiom :: Defn
- data AxiomData = AxiomData {
- _axiomConstTransp :: Bool
- data DataOrRecSigData = DataOrRecSigData {
- _datarecPars :: Int
- data FunctionData = FunctionData {
- _funClauses :: [Clause]
- _funCompiled :: Maybe CompiledClauses
- _funSplitTree :: Maybe SplitTree
- _funTreeless :: Maybe Compiled
- _funCovering :: [Clause]
- _funInv :: FunctionInverse
- _funMutual :: Maybe [QName]
- _funProjection :: Either ProjectionLikenessMissing Projection
- _funFlags :: SmallSet FunctionFlag
- _funTerminates :: Maybe Bool
- _funExtLam :: Maybe ExtLamInfo
- _funWith :: Maybe QName
- _funIsKanOp :: Maybe QName
- _funOpaque :: IsOpaque
- data DatatypeData = DatatypeData {
- _dataPars :: Nat
- _dataIxs :: Nat
- _dataClause :: Maybe Clause
- _dataCons :: [QName]
- _dataSort :: Sort
- _dataMutual :: Maybe [QName]
- _dataAbstr :: IsAbstract
- _dataPathCons :: [QName]
- _dataTranspIx :: Maybe QName
- _dataTransp :: Maybe QName
- data RecordData = RecordData {
- _recPars :: Nat
- _recClause :: Maybe Clause
- _recConHead :: ConHead
- _recNamedCon :: Bool
- _recFields :: [Dom QName]
- _recTel :: Telescope
- _recMutual :: Maybe [QName]
- _recEtaEquality' :: EtaEquality
- _recPatternMatching :: PatternOrCopattern
- _recInduction :: Maybe Induction
- _recTerminates :: Maybe Bool
- _recAbstr :: IsAbstract
- _recComp :: CompKit
- data ConstructorData = ConstructorData {
- _conPars :: Int
- _conArity :: Int
- _conSrcCon :: ConHead
- _conData :: QName
- _conAbstr :: IsAbstract
- _conComp :: CompKit
- _conProj :: Maybe [QName]
- _conForced :: [IsForced]
- _conErased :: Maybe [Bool]
- _conErasure :: !Bool
- _conInline :: !Bool
- data PrimitiveData = PrimitiveData {}
- data PrimitiveSortData = PrimitiveSortData {}
- pattern DataOrRecSig :: Int -> Defn
- pattern PrimitiveSort :: BuiltinSort -> Sort -> Defn
- axiomConstTransp :: Defn -> Bool
- datarecPars :: Defn -> Int
- data ProjectionLikenessMissing
- type FunctionInverse = FunctionInverse' Clause
- funClauses :: Defn -> [Clause]
- funCompiled :: Defn -> Maybe CompiledClauses
- funSplitTree :: Defn -> Maybe SplitTree
- funTreeless :: Defn -> Maybe Compiled
- funCovering :: Defn -> [Clause]
- funInv :: Defn -> FunctionInverse
- funMutual :: Defn -> Maybe [QName]
- funProjection :: Defn -> Either ProjectionLikenessMissing Projection
- funFlags :: Defn -> SmallSet FunctionFlag
- funTerminates :: Defn -> Maybe Bool
- funExtLam :: Defn -> Maybe ExtLamInfo
- funWith :: Defn -> Maybe QName
- funIsKanOp :: Defn -> Maybe QName
- funOpaque :: Defn -> IsOpaque
- dataPars :: Defn -> Nat
- dataIxs :: Defn -> Nat
- dataClause :: Defn -> Maybe Clause
- dataCons :: Defn -> [QName]
- dataSort :: Defn -> Sort
- dataMutual :: Defn -> Maybe [QName]
- dataAbstr :: Defn -> IsAbstract
- dataPathCons :: Defn -> [QName]
- dataTranspIx :: Defn -> Maybe QName
- dataTransp :: Defn -> Maybe QName
- recPars :: Defn -> Nat
- recClause :: Defn -> Maybe Clause
- recConHead :: Defn -> ConHead
- recNamedCon :: Defn -> Bool
- recFields :: Defn -> [Dom QName]
- recTel :: Defn -> Telescope
- recMutual :: Defn -> Maybe [QName]
- recEtaEquality' :: Defn -> EtaEquality
- recPatternMatching :: Defn -> PatternOrCopattern
- recInduction :: Defn -> Maybe Induction
- recTerminates :: Defn -> Maybe Bool
- recAbstr :: Defn -> IsAbstract
- recComp :: Defn -> CompKit
- conPars :: Defn -> Int
- conArity :: Defn -> Int
- conSrcCon :: Defn -> ConHead
- conData :: Defn -> QName
- conAbstr :: Defn -> IsAbstract
- conComp :: Defn -> CompKit
- conProj :: Defn -> Maybe [QName]
- conForced :: Defn -> [IsForced]
- conErased :: Defn -> Maybe [Bool]
- conErasure :: Defn -> Bool
- conInline :: Defn -> Bool
- primAbstr :: Defn -> IsAbstract
- primName :: Defn -> PrimitiveId
- primClauses :: Defn -> [Clause]
- primInv :: Defn -> FunctionInverse
- primCompiled :: Defn -> Maybe CompiledClauses
- primOpaque :: Defn -> IsOpaque
- data BuiltinSort
- data BuiltinDescriptor
- = BuiltinData (TCM Type) [BuiltinId]
- | BuiltinDataCons (TCM Type)
- | BuiltinPrim PrimitiveId (Term -> TCM ())
- | BuiltinSort BuiltinSort
- | BuiltinPostulate Relevance (TCM Type)
- | BuiltinUnknown (Maybe (TCM Type)) (Term -> Type -> TCM ())
- primSortName :: Defn -> BuiltinSort
- primSortSort :: Defn -> Sort
- lensFunction :: Lens' Defn FunctionData
- lensConstructor :: Lens' Defn ConstructorData
- lensRecord :: Lens' Defn RecordData
- lensRecTel :: Lens' RecordData Telescope
- lensRecEta :: Lens' RecordData EtaEquality
- data FunctionInverse' c
- = NotInjective
- | Inverse (InversionMap c)
- recRecursive :: Defn -> Bool
- recRecursive_ :: RecordData -> Bool
- recEtaEquality :: Defn -> HasEta
- _recEtaEquality :: RecordData -> HasEta
- emptyFunctionData :: HasOptions m => m FunctionData
- emptyFunctionData_ :: Bool -> FunctionData
- emptyFunction :: HasOptions m => m Defn
- emptyFunction_ :: Bool -> Defn
- funFlag_ :: FunctionFlag -> Lens' FunctionData Bool
- funFlag :: FunctionFlag -> Lens' Defn Bool
- funStatic :: Lens' Defn Bool
- funInline :: Lens' Defn Bool
- funMacro :: Lens' Defn Bool
- funMacro_ :: Lens' FunctionData Bool
- funFirstOrder :: Lens' Defn Bool
- funErasure :: Lens' Defn Bool
- funAbstract :: Lens' Defn Bool
- funAbstr :: Lens' Defn IsAbstract
- funAbstract_ :: Lens' FunctionData Bool
- funAbstr_ :: Lens' FunctionData IsAbstract
- funProj :: Lens' Defn Bool
- funProj_ :: Lens' FunctionData Bool
- isMacro :: Defn -> Bool
- isEmptyFunction :: Defn -> Bool
- isCopatternLHS :: [Clause] -> Bool
- recCon :: Defn -> QName
- defIsRecord :: Defn -> Bool
- defIsDataOrRecord :: Defn -> Bool
- defConstructors :: Defn -> [QName]
- data Simplification
- redReturn :: a -> ReduceM (Reduced a' a)
- redBind :: ReduceM (Reduced a a') -> (a -> b) -> (a' -> ReduceM (Reduced b b')) -> ReduceM (Reduced b b')
- 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 ReduceDefs
- = OnlyReduceDefs (Set QName)
- | DontReduceDefs (Set QName)
- reduceAllDefs :: ReduceDefs
- locallyReduceDefs :: MonadTCEnv m => ReduceDefs -> m a -> m a
- locallyTC :: MonadTCEnv m => Lens' TCEnv a -> (a -> a) -> m b -> m b
- eReduceDefs :: Lens' TCEnv ReduceDefs
- locallyReduceAllDefs :: MonadTCEnv m => m a -> m a
- shouldReduceDef :: MonadTCEnv m => QName -> m Bool
- asksTC :: MonadTCEnv m => (TCEnv -> a) -> m a
- toReduceDefs :: (Bool, [QName]) -> ReduceDefs
- fromReduceDefs :: ReduceDefs -> (Bool, [QName])
- locallyReconstructed :: MonadTCEnv m => m a -> m a
- eReconstructed :: Lens' TCEnv Bool
- isReconstructed :: MonadTCEnv m => m Bool
- data PrimitiveImpl = PrimImpl Type PrimFun
- 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]
- defNonterminating :: Definition -> Bool
- defTerminationUnconfirmed :: Definition -> Bool
- defForced :: Definition -> [IsForced]
- type InversionMap c = Map TermHead [c]
- data TermHead
- itableCounts :: Lens' InstanceTable (Map QName Int)
- pattern SortProp :: BuiltinSort
- pattern SortSet :: BuiltinSort
- pattern SortStrictSet :: BuiltinSort
- pattern SortPropOmega :: BuiltinSort
- pattern SortSetOmega :: BuiltinSort
- pattern SortStrictSetOmega :: BuiltinSort
- data BuiltinInfo = BuiltinInfo {}
- ifTopLevelAndHighlightingLevelIsOr :: MonadTCEnv tcm => HighlightingLevel -> Bool -> tcm () -> tcm ()
- ifTopLevelAndHighlightingLevelIs :: MonadTCEnv tcm => HighlightingLevel -> tcm () -> tcm ()
- type LetBindings = Map Name (Open LetBinding)
- data AbstractMode
- data UnquoteFlags = UnquoteFlags {
- _unquoteNormalise :: Bool
- initEnv :: TCEnv
- defaultUnquoteFlags :: UnquoteFlags
- unquoteNormalise :: Lens' UnquoteFlags Bool
- eUnquoteNormalise :: Lens' TCEnv Bool
- eUnquoteFlags :: Lens' TCEnv UnquoteFlags
- eContext :: Lens' TCEnv Context
- eLetBindings :: Lens' TCEnv LetBindings
- eCurrentModule :: Lens' TCEnv ModuleName
- eCurrentPath :: Lens' TCEnv (Maybe AbsolutePath)
- eAnonymousModules :: Lens' TCEnv [(ModuleName, Nat)]
- eImportPath :: Lens' TCEnv [TopLevelModuleName]
- eMutualBlock :: Lens' TCEnv (Maybe MutualId)
- eTerminationCheck :: Lens' TCEnv (TerminationCheck ())
- eCoverageCheck :: Lens' TCEnv CoverageCheck
- eMakeCase :: Lens' TCEnv Bool
- eSolvingConstraints :: Lens' TCEnv Bool
- eCheckingWhere :: Lens' TCEnv Bool
- eWorkingOnTypes :: Lens' TCEnv Bool
- eAssignMetas :: Lens' TCEnv Bool
- eActiveProblems :: Lens' TCEnv (Set ProblemId)
- eAbstractMode :: Lens' TCEnv AbstractMode
- eRelevance :: Lens' TCEnv Relevance
- eQuantity :: Lens' TCEnv Quantity
- eHardCompileTimeMode :: Lens' TCEnv Bool
- eSplitOnStrict :: Lens' TCEnv Bool
- eDisplayFormsEnabled :: Lens' TCEnv Bool
- eFoldLetBindings :: Lens' TCEnv Bool
- eRange :: Lens' TCEnv Range
- eHighlightingRange :: Lens' TCEnv Range
- eCall :: Lens' TCEnv (Maybe (Closure Call))
- eHighlightingLevel :: Lens' TCEnv HighlightingLevel
- eHighlightingMethod :: Lens' TCEnv HighlightingMethod
- eExpandLast :: Lens' TCEnv ExpandHidden
- eExpandLastBool :: Lens' TCEnv Bool
- isExpandLast :: ExpandHidden -> Bool
- toExpandLast :: Bool -> ExpandHidden
- eAppDef :: Lens' TCEnv (Maybe QName)
- eSimplification :: Lens' TCEnv Simplification
- eAllowedReductions :: Lens' TCEnv AllowedReductions
- eReduceDefsPair :: Lens' TCEnv (Bool, [QName])
- eInjectivityDepth :: Lens' TCEnv Int
- eCompareBlocked :: Lens' TCEnv Bool
- ePrintDomainFreePi :: Lens' TCEnv Bool
- ePrintMetasBare :: Lens' TCEnv Bool
- eInsideDotPattern :: Lens' TCEnv Bool
- eInstanceDepth :: Lens' TCEnv Int
- eIsDebugPrinting :: Lens' TCEnv Bool
- ePrintingPatternLambdas :: Lens' TCEnv [QName]
- eCallByNeed :: Lens' TCEnv Bool
- eCurrentCheckpoint :: Lens' TCEnv CheckpointId
- eCheckpoints :: Lens' TCEnv (Map CheckpointId Substitution)
- eGeneralizeMetas :: Lens' TCEnv DoGeneralize
- eGeneralizedVars :: Lens' TCEnv (Map QName GeneralizedValue)
- eActiveBackendName :: Lens' TCEnv (Maybe BackendName)
- eConflComputingOverlap :: Lens' TCEnv Bool
- eCurrentlyElaborating :: Lens' TCEnv Bool
- currentModality :: MonadTCEnv m => m Modality
- viewTC :: MonadTCEnv m => Lens' TCEnv a -> m a
- onLetBindingType :: (Dom Type -> Dom Type) -> LetBinding -> LetBinding
- isDontExpandLast :: ExpandHidden -> Bool
- data CandidateKind
- data TerminationError = TerminationError {
- termErrFunctions :: [QName]
- termErrCalls :: [CallInfo]
- data IllegalRewriteRuleReason
- = LHSNotDefinitionOrConstructor
- | VariablesNotBoundByLHS IntSet
- | VariablesBoundMoreThanOnce IntSet
- | LHSReduces Term Term
- | HeadSymbolIsProjection QName
- | HeadSymbolIsProjectionLikeFunction QName
- | HeadSymbolIsTypeConstructor QName
- | HeadSymbolContainsMetas QName
- | ConstructorParametersNotGeneral ConHead Args
- | ContainsUnsolvedMetaVariables (Set MetaId)
- | BlockedOnProblems (Set ProblemId)
- | RequiresDefinitions (Set QName)
- | DoesNotTargetRewriteRelation
- | BeforeFunctionDefinition
- | BeforeMutualFunctionDefinition QName
- | DuplicateRewriteRule
- recordFieldWarningToError :: RecordFieldWarning -> TypeError
- warningName :: Warning -> WarningName
- illegalRewriteWarningName :: IllegalRewriteRuleReason -> WarningName
- isSourceCodeWarning :: WarningName -> Bool
- tcWarningOrigin :: TCWarning -> SrcFile
- data CallInfo = CallInfo {}
- data ErasedDatatypeReason
- data SplitError
- = NotADatatype (Closure Type)
- | BlockedType Blocker (Closure Type)
- | ErasedDatatype ErasedDatatypeReason (Closure Type)
- | CoinductiveDatatype (Closure Type)
- | UnificationStuck { }
- | CosplitCatchall
- | CosplitNoTarget
- | CosplitNoRecordType (Closure Type)
- | CannotCreateMissingClause QName (Telescope, [NamedArg DeBruijnPattern]) Doc (Closure (Abs Type))
- | GenericSplitError String
- data UnificationFailure
- data NegativeUnification
- data UnquoteError
- = BadVisibility String (Arg Term)
- | ConInsteadOfDef QName String String
- | DefInsteadOfCon QName String String
- | NonCanonical String Term
- | BlockedOnMeta TCState Blocker
- | PatLamWithoutClauses Term
- | UnquotePanic String
- type DataOrRecordE = DataOrRecord' InductionAndEta
- data IncorrectTypeForRewriteRelationReason
- data InvalidFileNameReason
- data InductionAndEta = InductionAndEta {}
- sizedTypesOption :: HasOptions m => m Bool
- guardednessOption :: HasOptions m => m Bool
- withoutKOption :: HasOptions m => m Bool
- cubicalOption :: HasOptions m => m (Maybe Cubical)
- cubicalCompatibleOption :: HasOptions m => m Bool
- 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' ReduceEnv TCEnv
- reduceSt :: Lens' ReduceEnv TCState
- unKleisli :: (a -> ReduceM b) -> ReduceM (a -> b)
- onReduceEnv :: (ReduceEnv -> ReduceEnv) -> ReduceM a -> ReduceM a
- fmapReduce :: (a -> b) -> ReduceM a -> ReduceM b
- apReduce :: ReduceM (a -> b) -> ReduceM a -> ReduceM b
- thenReduce :: ReduceM a -> ReduceM b -> ReduceM b
- beforeReduce :: ReduceM a -> ReduceM b -> ReduceM a
- bindReduce :: ReduceM a -> (a -> ReduceM b) -> ReduceM b
- runReduceF :: (a -> ReduceM b) -> TCM (a -> b)
- localR :: (ReduceEnv -> ReduceEnv) -> ReduceM a -> ReduceM a
- newtype BlockT (m :: Type -> Type) a = BlockT {}
- getsTC :: ReadTCState m => (TCState -> a) -> m a
- modifyTC' :: MonadTCState m => (TCState -> TCState) -> m ()
- setTCLens :: MonadTCState m => Lens' TCState a -> a -> m ()
- setTCLens' :: MonadTCState m => Lens' TCState a -> a -> m ()
- modifyTCLens :: MonadTCState m => Lens' TCState a -> (a -> a) -> m ()
- modifyTCLens' :: MonadTCState m => Lens' TCState a -> (a -> a) -> m ()
- modifyTCLensM :: MonadTCState m => Lens' TCState a -> (a -> m a) -> m ()
- stateTCLens :: MonadTCState m => Lens' TCState a -> (a -> (r, a)) -> m r
- runBlocked :: Monad m => BlockT m a -> m (Either Blocker a)
- pureTCM :: forall (m :: Type -> Type) a. MonadIO m => (TCState -> TCEnv -> a) -> TCMT m a
- returnTCMT :: forall (m :: Type -> Type) a. Applicative m => a -> TCMT m a
- bindTCMT :: forall (m :: Type -> Type) a b. Monad m => TCMT m a -> (a -> TCMT m b) -> TCMT m b
- thenTCMT :: forall (m :: Type -> Type) a b. Applicative m => TCMT m a -> TCMT m b -> TCMT m b
- fmapTCMT :: forall (m :: Type -> Type) a b. Functor m => (a -> b) -> TCMT m a -> TCMT m b
- apTCMT :: forall (m :: Type -> Type) a b. Applicative m => TCMT m (a -> b) -> TCMT m a -> TCMT m b
- internalError :: (HasCallStack, MonadTCM tcm) => String -> tcm a
- catchError_ :: TCM a -> (TCErr -> TCM a) -> TCM a
- finally_ :: TCM a -> TCM b -> TCM a
- typeError' :: MonadTCError m => CallStack -> TypeError -> m a
- locatedTypeError :: MonadTCError m => (a -> TypeError) -> HasCallStack => a -> m b
- genericDocError :: (HasCallStack, MonadTCError m) => Doc -> m a
- typeError'_ :: (MonadTCEnv m, ReadTCState m) => CallStack -> TypeError -> m TCErr
- typeError_ :: (HasCallStack, 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
- doExpandLast :: TCM a -> TCM a
- workOnTypes :: (MonadTCEnv m, HasOptions m, MonadDebug m) => m a -> m a
- withCurrentModule :: MonadTCEnv m => ModuleName -> m a -> m a
- getAnonymousVariables :: MonadTCEnv m => ModuleName -> m Nat
- withAnonymousModule :: ModuleName -> Nat -> TCM a -> TCM a
- withEnv :: MonadTCEnv m => TCEnv -> m a -> m a
- withHighlightingLevel :: HighlightingLevel -> TCM a -> TCM a
- dontExpandLast :: TCM a -> TCM a
- reallyDontExpandLast :: TCM a -> TCM a
- performedSimplification :: MonadTCEnv m => m a -> m a
- performedSimplification' :: MonadTCEnv m => Simplification -> m a -> m a
- getSimplification :: MonadTCEnv m => m Simplification
- updateAllowedReductions :: (AllowedReductions -> AllowedReductions) -> TCEnv -> TCEnv
- modifyAllowedReductions :: MonadTCEnv m => (AllowedReductions -> AllowedReductions) -> m a -> m a
- putAllowedReductions :: MonadTCEnv m => AllowedReductions -> m a -> m a
- onlyReduceProjections :: MonadTCEnv m => m a -> m a
- allowAllReductions :: MonadTCEnv m => m a -> m a
- allowNonTerminatingReductions :: MonadTCEnv m => m a -> m a
- onlyReduceTypes :: MonadTCEnv m => m a -> m a
- typeLevelReductions :: MonadTCEnv m => m a -> m a
- callByName :: TCM a -> TCM a
- dontFoldLetBindings :: MonadTCEnv m => m a -> m a
- removeLetBinding :: MonadTCEnv m => Name -> m a -> m a
- newtype BuiltinAccess a = BuiltinAccess {
- unBuiltinAccess :: TCState -> a
- runBuiltinAccess :: TCState -> BuiltinAccess a -> a
- litType :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => Literal -> m Type
- primZero :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primSuc :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primNat :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primWord64 :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primFloat :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primChar :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primString :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primQName :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaMeta :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- setBuiltinThings :: BuiltinThings PrimFun -> TCM ()
- bindBuiltinName :: BuiltinId -> Term -> TCM ()
- bindPrimitive :: PrimitiveId -> PrimFun -> TCM ()
- bindBuiltinRewriteRelation :: QName -> TCM ()
- getBuiltinRewriteRelations :: (HasBuiltins m, MonadTCError m) => m (Set QName)
- getBuiltinRewriteRelations' :: HasBuiltins m => m (Maybe (Set QName))
- getBuiltin :: (HasBuiltins m, MonadTCError m) => BuiltinId -> m Term
- getPrimitive' :: HasBuiltins m => PrimitiveId -> m (Maybe PrimFun)
- getPrimitive :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => PrimitiveId -> m PrimFun
- getPrimitiveTerm :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => PrimitiveId -> m Term
- getPrimitiveTerm' :: HasBuiltins m => PrimitiveId -> m (Maybe Term)
- getPrimitiveName' :: HasBuiltins m => PrimitiveId -> m (Maybe QName)
- getTerm' :: (HasBuiltins m, IsBuiltin a) => a -> m (Maybe Term)
- getTerm :: (HasBuiltins m, IsBuiltin a) => String -> a -> m Term
- constructorForm' :: Applicative m => m Term -> m Term -> Term -> m Term
- primInteger :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primIntegerPos :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primIntegerNegSuc :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primUnit :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primUnitUnit :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primBool :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primTrue :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primFalse :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primSigma :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primList :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primNil :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primCons :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primIO :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primMaybe :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primNothing :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primJust :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primPath :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primPathP :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primIntervalUniv :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primInterval :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primIZero :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primIOne :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primIsOne :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primItIsOne :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primIsOne1 :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primIsOne2 :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primIsOneEmpty :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primSub :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primSubIn :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primTrans :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primId :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primEquiv :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primEquivFun :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primEquivProof :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primTranspProof :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- prim_glue :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- prim_unglue :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- prim_glueU :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- prim_unglueU :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primSizeUniv :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primSize :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primSizeLt :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primSizeSuc :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primSizeInf :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primSizeMax :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primInf :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primSharp :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primFlat :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primEquality :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primRefl :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primLevel :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primLevelUniv :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primProp :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primSet :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primStrictSet :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primPropOmega :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primSetOmega :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primSSetOmega :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primFromNat :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primFromNeg :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primFromString :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primArgInfo :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primArgArgInfo :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primArg :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primArgArg :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAbs :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAbsAbs :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaTerm :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaTermVar :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaTermLam :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaTermExtLam :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaTermDef :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaTermCon :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaTermPi :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaTermSort :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaTermLit :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaTermUnsupported :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaTermMeta :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaErrorPart :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaErrorPartString :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaErrorPartTerm :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaErrorPartPatt :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaErrorPartName :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primHiding :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primHidden :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primInstance :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primVisible :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primRelevance :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primRelevant :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primIrrelevant :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primQuantity :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primQuantity0 :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primQuantityω :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primModality :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primModalityConstructor :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAssoc :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAssocLeft :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAssocRight :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAssocNon :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primPrecedence :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primPrecRelated :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primPrecUnrelated :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primFixity :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primFixityFixity :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaLiteral :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaLitNat :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaLitWord64 :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaLitFloat :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaLitString :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaLitChar :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaLitQName :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaLitMeta :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaSort :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaSortSet :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaSortLit :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaSortProp :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaSortPropLit :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaSortInf :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaSortUnsupported :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaDefinition :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaDefinitionFunDef :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaDefinitionDataDef :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaDefinitionRecordDef :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaDefinitionPostulate :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaDefinitionPrimitive :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaDefinitionDataConstructor :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaClause :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaClauseClause :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaClauseAbsurd :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaPattern :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaPatCon :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaPatVar :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaPatDot :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaPatLit :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaPatProj :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaPatAbsurd :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaBlocker :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaBlockerAny :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaBlockerAll :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaBlockerMeta :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaTCM :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaTCMReturn :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaTCMBind :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaTCMUnify :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaTCMTypeError :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaTCMInferType :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaTCMCheckType :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaTCMNormalise :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaTCMReduce :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaTCMCatchError :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaTCMGetContext :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaTCMExtendContext :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaTCMInContext :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaTCMFreshName :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaTCMDeclareDef :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaTCMDeclarePostulate :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaTCMDeclareData :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaTCMDefineData :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaTCMDefineFun :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaTCMGetType :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaTCMGetDefinition :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaTCMQuoteTerm :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaTCMUnquoteTerm :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaTCMQuoteOmegaTerm :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaTCMCommit :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaTCMIsMacro :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaTCMBlock :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaTCMFormatErrorParts :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaTCMDebugPrint :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaTCMWithNormalisation :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaTCMWithReconstructed :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaTCMWithExpandLast :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaTCMWithReduceDefs :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaTCMAskNormalisation :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaTCMAskReconstructed :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaTCMAskExpandLast :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaTCMAskReduceDefs :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaTCMNoConstraints :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaTCMWorkOnTypes :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaTCMRunSpeculative :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaTCMExec :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaTCMGetInstances :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaTCMSolveInstances :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaTCMPragmaForeign :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- primAgdaTCMPragmaCompile :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term
- data CoinductionKit = CoinductionKit {
- nameOfInf :: QName
- nameOfSharp :: QName
- nameOfFlat :: QName
- coinductionKit' :: TCM CoinductionKit
- coinductionKit :: TCM (Maybe CoinductionKit)
- data SortKit = SortKit {
- nameOfUniv :: UnivSize -> Univ -> QName
- isNameOfUniv :: QName -> Maybe (UnivSize, Univ)
- mkSortKit :: QName -> QName -> QName -> QName -> QName -> QName -> SortKit
- sortKit :: (HasBuiltins m, MonadTCError m, HasOptions m) => m SortKit
- infallibleSortKit :: HasBuiltins m => m SortKit
- getPrimName :: Term -> QName
- isPrimitive :: HasBuiltins m => PrimitiveId -> QName -> m Bool
- intervalSort :: Sort
- intervalView' :: HasBuiltins m => m (Term -> IntervalView)
- intervalView :: HasBuiltins m => Term -> m IntervalView
- intervalUnview :: HasBuiltins m => IntervalView -> m Term
- intervalUnview' :: HasBuiltins m => m (IntervalView -> Term)
- pathView :: HasBuiltins m => Type -> m PathView
- pathView' :: HasBuiltins m => m (Type -> PathView)
- idViewAsPath :: HasBuiltins m => Type -> m PathView
- boldPathView :: Type -> PathView
- pathUnview :: PathView -> Type
- conidView' :: HasBuiltins m => m (Term -> Term -> Maybe (Arg Term, Arg Term))
- primEqualityName :: TCM QName
- equalityView :: Type -> TCM EqualityView
- class EqualityUnview a where
- equalityUnview :: a -> Type
- constrainedPrims :: [PrimitiveId]
- getNameOfConstrained :: HasBuiltins m => PrimitiveId -> m (Maybe QName)
- isInteractionMeta :: ReadTCState m => MetaId -> m (Maybe InteractionId)
- notSoPrettySigCubicalNotErasure :: QName -> String
- inFreshModuleIfFreeParams :: TCM a -> TCM a
- lookupSection :: (Functor m, ReadTCState m) => ModuleName -> m Telescope
- defaultGetVerbosity :: HasOptions m => m Verbosity
- defaultGetProfileOptions :: HasOptions m => m ProfileOptions
- defaultIsDebugPrinting :: MonadTCEnv m => m Bool
- defaultNowDebugPrinting :: MonadTCEnv m => m a -> m a
- displayDebugMessage :: MonadDebug m => VerboseKey -> VerboseLevel -> String -> m ()
- catchAndPrintImpossible :: (CatchImpossible m, Monad m) => VerboseKey -> VerboseLevel -> m String -> m String
- hasVerbosity :: MonadDebug m => VerboseKey -> VerboseLevel -> m Bool
- alwaysReportSLn :: MonadDebug m => VerboseKey -> VerboseLevel -> String -> m ()
- reportResult :: MonadDebug m => VerboseKey -> VerboseLevel -> (a -> TCM Doc) -> m a -> m a
- alwaysReportSDoc :: MonadDebug m => VerboseKey -> VerboseLevel -> TCM Doc -> m ()
- unlessDebugPrinting :: MonadDebug m => m () -> m ()
- class TraceS a where
- traceS :: MonadDebug m => VerboseKey -> VerboseLevel -> a -> m c -> m c
- traceSDoc :: MonadDebug m => VerboseKey -> VerboseLevel -> TCM Doc -> m a -> m a
- traceSLn :: MonadDebug m => VerboseKey -> VerboseLevel -> String -> m a -> m a
- openVerboseBracket :: MonadDebug m => VerboseKey -> VerboseLevel -> String -> m ()
- closeVerboseBracket :: MonadDebug m => VerboseKey -> VerboseLevel -> m ()
- closeVerboseBracketException :: MonadDebug m => VerboseKey -> VerboseLevel -> m ()
- hasExactVerbosity :: MonadDebug m => VerboseKey -> VerboseLevel -> m Bool
- whenExactVerbosity :: MonadDebug m => VerboseKey -> VerboseLevel -> m () -> m ()
- __CRASH_WHEN__ :: (HasCallStack, MonadTCM m, MonadDebug m) => VerboseKey -> VerboseLevel -> m ()
- hasProfileOption :: MonadDebug m => ProfileOption -> m Bool
- whenProfile :: MonadDebug m => ProfileOption -> m () -> m ()
- resetAllState :: TCM ()
- updatePersistentState :: (PersistentTCState -> PersistentTCState) -> TCState -> TCState
- localTCStateSaving :: TCM a -> TCM (a, TCState)
- localTCStateSavingWarnings :: TCM a -> TCM a
- data SpeculateResult
- speculateTCState :: TCM (a, SpeculateResult) -> TCM a
- speculateTCState_ :: TCM SpeculateResult -> TCM ()
- freshTCM :: TCM a -> TCM (Either TCErr a)
- lensPersistentState :: Lens' TCState PersistentTCState
- modifyPersistentState :: (PersistentTCState -> PersistentTCState) -> TCM ()
- lensAccumStatisticsP :: Lens' PersistentTCState Statistics
- lensAccumStatistics :: Lens' TCState Statistics
- setScope :: ScopeInfo -> TCM ()
- modifyScope :: MonadTCState m => (ScopeInfo -> ScopeInfo) -> m ()
- modifyScope_ :: MonadTCState m => (ScopeInfo -> ScopeInfo) -> m ()
- useScope :: ReadTCState m => Lens' ScopeInfo a -> m a
- locallyScope :: ReadTCState m => Lens' ScopeInfo a -> (a -> a) -> m b -> m b
- withScope_ :: ReadTCState m => ScopeInfo -> m a -> m a
- localScope :: TCM a -> TCM a
- notInScopeError :: QName -> TCM a
- printScope :: String -> Int -> String -> TCM ()
- notInScopeWarning :: QName -> TCM ()
- modifySignature :: MonadTCState m => (Signature -> Signature) -> m ()
- modifyImportedSignature :: MonadTCState m => (Signature -> Signature) -> m ()
- getSignature :: ReadTCState m => m Signature
- modifyGlobalDefinition :: MonadTCState m => QName -> (Definition -> Definition) -> m ()
- updateDefinition :: QName -> (Definition -> Definition) -> Signature -> Signature
- setSignature :: MonadTCState m => Signature -> m ()
- withSignature :: (ReadTCState m, MonadTCState m) => Signature -> m a -> m a
- addRewriteRulesFor :: QName -> RewriteRules -> [QName] -> Signature -> Signature
- updateTheDef :: (Defn -> Defn) -> Definition -> Definition
- setMatchableSymbols :: QName -> [QName] -> Signature -> Signature
- updateDefCopatternLHS :: (Bool -> Bool) -> Definition -> Definition
- modifyRecEta :: MonadTCState m => QName -> (EtaEquality -> EtaEquality) -> m ()
- lookupDefinition :: QName -> Signature -> Maybe Definition
- updateDefinitions :: (Definitions -> Definitions) -> Signature -> Signature
- updateDefType :: (Type -> Type) -> Definition -> Definition
- updateDefArgOccurrences :: ([Occurrence] -> [Occurrence]) -> Definition -> Definition
- updateDefPolarity :: ([Polarity] -> [Polarity]) -> Definition -> Definition
- updateDefCompiledRep :: (CompiledRepresentation -> CompiledRepresentation) -> Definition -> Definition
- addCompilerPragma :: BackendName -> CompilerPragma -> Definition -> Definition
- updateFunClauses :: ([Clause] -> [Clause]) -> Defn -> Defn
- updateCovering :: ([Clause] -> [Clause]) -> Defn -> Defn
- updateCompiledClauses :: (Maybe CompiledClauses -> Maybe CompiledClauses) -> Defn -> Defn
- updateDefBlocked :: (Blocked_ -> Blocked_) -> Definition -> Definition
- setTopLevelModule :: TopLevelModuleName -> TCM ()
- currentTopLevelModule :: (MonadTCEnv m, ReadTCState m) => m (Maybe TopLevelModuleName)
- withTopLevelModule :: TopLevelModuleName -> TCM a -> TCM a
- addForeignCode :: BackendName -> String -> TCM ()
- getInteractionOutputCallback :: ReadTCState m => m InteractionOutputCallback
- appInteractionOutputCallback :: Response -> TCM ()
- setInteractionOutputCallback :: InteractionOutputCallback -> TCM ()
- getPatternSyns :: ReadTCState m => m PatternSynDefns
- setPatternSyns :: PatternSynDefns -> TCM ()
- modifyPatternSyns :: (PatternSynDefns -> PatternSynDefns) -> TCM ()
- getPatternSynImports :: ReadTCState m => m PatternSynDefns
- lookupPatternSyn :: AmbiguousQName -> TCM PatternSynDefn
- lookupSinglePatternSyn :: QName -> TCM PatternSynDefn
- theBenchmark :: TCState -> Benchmark
- updateBenchmark :: (Benchmark -> Benchmark) -> TCState -> TCState
- updateInstanceDefs :: (TempInstanceTable -> TempInstanceTable) -> TCState -> TCState
- modifyInstanceDefs :: (TempInstanceTable -> TempInstanceTable) -> TCM ()
- getAllInstanceDefs :: TCM TempInstanceTable
- getAnonInstanceDefs :: TCM (Set QName)
- clearUnknownInstance :: QName -> TCM ()
- addUnknownInstance :: QName -> TCM ()
- getCurrentRange :: MonadTCEnv m => m Range
- data BoundedSize
- class IsSizeType a where
- isSizeType :: (HasOptions m, HasBuiltins m) => a -> m (Maybe BoundedSize)
- isSizeTypeTest :: (HasOptions m, HasBuiltins m) => m (Term -> Maybe BoundedSize)
- getBuiltinSize :: HasBuiltins m => m (Maybe QName, Maybe QName)
- getBuiltinDefName :: HasBuiltins m => BuiltinId -> m (Maybe QName)
- isSizeNameTest :: (HasOptions m, HasBuiltins m) => m (QName -> Bool)
- isSizeNameTestRaw :: (HasOptions m, HasBuiltins m) => m (QName -> Bool)
- haveSizedTypes :: TCM Bool
- haveSizeLt :: TCM Bool
- builtinSizeHook :: BuiltinId -> QName -> Type -> TCM ()
- sizeSort :: Sort
- sizeUniv :: Type
- sizeType_ :: QName -> Type
- sizeSucName :: (HasBuiltins m, HasOptions m) => m (Maybe QName)
- sizeSuc :: HasBuiltins m => Nat -> Term -> m Term
- sizeSuc_ :: QName -> Term -> Term
- sizeMax :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => List1 Term -> m Term
- data SizeView
- sizeView :: (HasBuiltins m, MonadTCEnv m, ReadTCState m) => Term -> m SizeView
- data ProjectedVar = ProjectedVar {
- pvIndex :: Int
- prProjs :: [(ProjOrigin, QName)]
- viewProjectedVar :: Term -> Maybe ProjectedVar
- unviewProjectedVar :: ProjectedVar -> Term
- data DeepSizeView
- data SizeViewComparable a
- sizeViewComparable :: DeepSizeView -> DeepSizeView -> SizeViewComparable ()
- sizeViewSuc_ :: QName -> DeepSizeView -> DeepSizeView
- sizeViewPred :: Nat -> DeepSizeView -> DeepSizeView
- sizeViewOffset :: DeepSizeView -> Maybe Offset
- removeSucs :: (DeepSizeView, DeepSizeView) -> (DeepSizeView, DeepSizeView)
- unSizeView :: SizeView -> TCM Term
- unDeepSizeView :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => DeepSizeView -> m Term
- type SizeMaxView = List1 DeepSizeView
- type SizeMaxView' = [DeepSizeView]
- maxViewMax :: SizeMaxView -> SizeMaxView -> SizeMaxView
- maxViewCons :: DeepSizeView -> SizeMaxView -> SizeMaxView
- sizeViewComparableWithMax :: DeepSizeView -> SizeMaxView -> SizeViewComparable SizeMaxView'
- maxViewSuc_ :: QName -> SizeMaxView -> SizeMaxView
- unMaxView :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => SizeMaxView -> m Term
- modifyContextInfo :: MonadTCEnv tcm => (forall e. Dom e -> Dom e) -> tcm a -> tcm a
- unsafeInTopContext :: (MonadTCEnv m, ReadTCState m) => m a -> m a
- unsafeEscapeContext :: MonadTCM tcm => Int -> tcm a -> tcm a
- escapeContext :: MonadAddContext m => Impossible -> Int -> m a -> m a
- checkpoint :: (MonadDebug tcm, MonadTCM tcm, MonadFresh CheckpointId tcm, ReadTCState tcm) => Substitution -> tcm a -> tcm a
- getContextTelescope :: (Applicative m, MonadTCEnv m) => m Telescope
- checkpointSubstitution' :: MonadTCEnv tcm => CheckpointId -> tcm (Maybe Substitution)
- getModuleParameterSub :: (MonadTCEnv m, ReadTCState m) => ModuleName -> m (Maybe Substitution)
- defaultAddCtx :: MonadAddContext m => Name -> Dom Type -> m a -> m a
- withFreshName_ :: MonadAddContext m => ArgName -> (Name -> m a) -> m a
- withShadowingNameTCM :: Name -> TCM b -> TCM b
- defaultAddLetBinding' :: (ReadTCState m, MonadTCEnv m) => Origin -> Name -> Term -> Dom Type -> m a -> m a
- addRecordNameContext :: (MonadAddContext m, MonadFresh NameId m) => Dom Type -> m b -> m b
- class AddContext b where
- addContext :: MonadAddContext m => b -> m a -> m a
- contextSize :: b -> Nat
- newtype KeepNames a = KeepNames a
- underAbstraction' :: (Subst a, MonadAddContext m, AddContext (name, Dom Type)) => (String -> name) -> Dom Type -> Abs a -> (a -> m b) -> m b
- underAbstraction :: (Subst a, MonadAddContext m) => Dom Type -> Abs a -> (a -> m b) -> m b
- underAbstractionAbs' :: (Subst a, MonadAddContext m, AddContext (name, Dom Type)) => (String -> name) -> Dom Type -> Abs a -> (a -> m b) -> m b
- underAbstractionAbs :: (Subst a, MonadAddContext m) => Dom Type -> Abs a -> (a -> m b) -> m b
- underAbstraction_ :: (Subst a, MonadAddContext m) => Abs a -> (a -> m b) -> m b
- mapAbstraction :: (Subst a, Subst b, MonadAddContext m) => Dom Type -> (a -> m b) -> Abs a -> m (Abs b)
- mapAbstraction_ :: (Subst a, Subst b, MonadAddContext m) => (a -> m b) -> Abs a -> m (Abs b)
- getLetBindings :: MonadTCEnv tcm => tcm [(Name, LetBinding)]
- addLetBinding :: MonadAddContext m => ArgInfo -> Origin -> Name -> Term -> Type -> m a -> m a
- removeLetBindingsFrom :: MonadTCEnv m => Name -> m a -> m a
- getContextSize :: (Applicative m, MonadTCEnv m) => m Nat
- getContextArgs :: (Applicative m, MonadTCEnv m) => m Args
- getContextTerms :: (Applicative m, MonadTCEnv m) => m [Term]
- getContextNames :: (Applicative m, MonadTCEnv m) => m [Name]
- lookupBV_ :: Nat -> Context -> Maybe ContextEntry
- lookupBV' :: MonadTCEnv m => Nat -> m (Maybe ContextEntry)
- lookupBV :: (MonadFail m, MonadTCEnv m) => Nat -> m (Dom (Name, Type))
- domOfBV :: (Applicative m, MonadFail m, MonadTCEnv m) => Nat -> m (Dom Type)
- typeOfBV :: (Applicative m, MonadFail m, MonadTCEnv m) => Nat -> m Type
- nameOfBV' :: (Applicative m, MonadFail m, MonadTCEnv m) => Nat -> m (Maybe Name)
- nameOfBV :: (Applicative m, MonadFail m, MonadTCEnv m) => Nat -> m Name
- getVarInfo :: (MonadFail m, MonadTCEnv m) => Name -> m (Term, Dom Type)
- workOnTypes' :: MonadTCEnv m => Bool -> m a -> m a
- applyQuantityToJudgement :: (MonadTCEnv tcm, LensQuantity q) => q -> tcm a -> tcm a
- applyRelevanceToContext :: (MonadTCEnv tcm, LensRelevance r) => r -> tcm a -> tcm a
- applyRelevanceToContextOnly :: MonadTCEnv tcm => Relevance -> tcm a -> tcm a
- applyRelevanceToJudgementOnly :: MonadTCEnv tcm => Relevance -> tcm a -> tcm a
- applyRelevanceToContextFunBody :: (MonadTCM tcm, LensRelevance r) => r -> tcm a -> tcm a
- applyCohesionToContext :: (MonadTCEnv tcm, LensCohesion m) => m -> tcm a -> tcm a
- applyCohesionToContextOnly :: MonadTCEnv tcm => Cohesion -> tcm a -> tcm a
- splittableCohesion :: (HasOptions m, LensCohesion a) => a -> m Bool
- applyModalityToContext :: (MonadTCEnv tcm, LensModality m) => m -> tcm a -> tcm a
- applyModalityToContextOnly :: MonadTCEnv tcm => Modality -> tcm a -> tcm a
- applyModalityToJudgementOnly :: MonadTCEnv tcm => Modality -> tcm a -> tcm a
- applyModalityToContextFunBody :: (MonadTCM tcm, LensModality r) => r -> tcm a -> tcm a
- wakeIrrelevantVars :: MonadTCEnv tcm => tcm a -> tcm a
- withClosure :: (MonadTCEnv m, ReadTCState m) => Closure a -> (a -> m b) -> m (Closure b)
- mapClosure :: (MonadTCEnv m, ReadTCState m) => (a -> m b) -> Closure a -> m (Closure b)
- solvingProblems :: MonadConstraint m => Set ProblemId -> m a -> m a
- isProblemSolved :: (MonadTCEnv m, ReadTCState m) => ProblemId -> m Bool
- isProblemSolved' :: (MonadTCEnv m, ReadTCState m) => Bool -> ProblemId -> m Bool
- isProblemCompletelySolved :: (MonadTCEnv m, ReadTCState m) => ProblemId -> m Bool
- getAllConstraints :: ReadTCState m => m Constraints
- isBlockingConstraint :: Constraint -> Bool
- getConstraintsForProblem :: ReadTCState m => ProblemId -> m Constraints
- getAwakeConstraints :: ReadTCState m => m Constraints
- dropConstraints :: MonadConstraint m => (ProblemConstraint -> Bool) -> m ()
- takeConstraints :: MonadConstraint m => (ProblemConstraint -> Bool) -> m Constraints
- putConstraintsToSleep :: MonadConstraint m => (ProblemConstraint -> Bool) -> m ()
- putAllConstraintsToSleep :: MonadConstraint m => m ()
- data ConstraintStatus
- holdConstraints :: (ConstraintStatus -> ProblemConstraint -> Bool) -> TCM a -> TCM a
- takeAwakeConstraint :: MonadConstraint m => m (Maybe ProblemConstraint)
- takeAwakeConstraint' :: MonadConstraint m => (ProblemConstraint -> Bool) -> m (Maybe ProblemConstraint)
- withConstraint :: MonadConstraint m => (Constraint -> m a) -> ProblemConstraint -> m a
- buildProblemConstraint :: (MonadTCEnv m, ReadTCState m) => Set ProblemId -> Blocker -> Constraint -> m ProblemConstraint
- buildProblemConstraint_ :: (MonadTCEnv m, ReadTCState m) => Blocker -> Constraint -> m ProblemConstraint
- buildConstraint :: Blocker -> Constraint -> TCM ProblemConstraint
- addConstraint' :: Blocker -> Constraint -> TCM ()
- addConstraintTo :: Lens' TCState Constraints -> Blocker -> Constraint -> TCM ()
- addAwakeConstraint' :: Blocker -> Constraint -> TCM ()
- nowSolvingConstraints :: MonadTCEnv m => m a -> m a
- isSolvingConstraints :: MonadTCEnv m => m Bool
- catchConstraint :: MonadConstraint m => Constraint -> m () -> m ()
- canDropRecursiveInstance :: (ReadTCState m, HasOptions m) => m Bool
- wakeConstraints' :: MonadConstraint m => (ProblemConstraint -> WakeUp) -> m ()
- mapAwakeConstraints :: (Constraints -> Constraints) -> TCState -> TCState
- mapSleepingConstraints :: (Constraints -> Constraints) -> TCState -> TCState
- inMutualBlock :: (MutualId -> TCM a) -> TCM a
- setMutualBlockInfo :: MutualId -> MutualInfo -> TCM ()
- insertMutualBlockInfo :: MutualId -> MutualInfo -> TCM ()
- setMutualBlock :: MutualId -> QName -> TCM ()
- currentOrFreshMutualBlock :: TCM MutualId
- lookupMutualBlock :: ReadTCState tcm => MutualId -> tcm MutualBlock
- mutualBlockOf :: QName -> TCM MutualId
- checkPragmaOptionConsistency :: PragmaOptions -> PragmaOptions -> TCM ()
- setCommandLineOptions' :: AbsolutePath -> CommandLineOptions -> TCM ()
- setLibraryPaths :: AbsolutePath -> CommandLineOptions -> TCM CommandLineOptions
- setIncludeDirs :: [FilePath] -> AbsolutePath -> TCM ()
- getAgdaLibFilesWithoutTopLevelModuleName :: AbsolutePath -> TCM [AgdaLibFile]
- checkLibraryFileNotTooFarDown :: TopLevelModuleName -> AgdaLibFile -> TCM ()
- getLibraryOptions :: AbsolutePath -> TopLevelModuleName -> TCM [OptionsPragma]
- setLibraryIncludes :: CommandLineOptions -> TCM CommandLineOptions
- addDefaultLibraries :: AbsolutePath -> CommandLineOptions -> TCM CommandLineOptions
- addTrustedExecutables :: CommandLineOptions -> TCM CommandLineOptions
- setOptionsFromPragma :: OptionsPragma -> TCM ()
- setOptionsFromPragma' :: Bool -> OptionsPragma -> TCM ()
- checkAndSetOptionsFromPragma :: OptionsPragma -> TCM ()
- enableDisplayForms :: MonadTCEnv m => m a -> m a
- disableDisplayForms :: MonadTCEnv m => m a -> m a
- displayFormsEnabled :: MonadTCEnv m => m Bool
- isPropEnabled :: HasOptions m => m Bool
- isLevelUniverseEnabled :: HasOptions m => m Bool
- isTwoLevelEnabled :: HasOptions m => m Bool
- hasUniversePolymorphism :: HasOptions m => m Bool
- showImplicitArguments :: HasOptions m => m Bool
- showGeneralizedArguments :: HasOptions m => m Bool
- showIrrelevantArguments :: HasOptions m => m Bool
- showIdentitySubstitutions :: HasOptions m => m Bool
- withShowAllArguments :: ReadTCState m => m a -> m a
- withShowAllArguments' :: ReadTCState m => Bool -> m a -> m a
- withPragmaOptions :: ReadTCState m => (PragmaOptions -> PragmaOptions) -> m a -> m a
- withoutPrintingGeneralization :: ReadTCState m => m a -> m a
- positivityCheckEnabled :: HasOptions m => m Bool
- typeInType :: HasOptions m => m Bool
- etaEnabled :: HasOptions m => m Bool
- maxInstanceSearchDepth :: HasOptions m => m Int
- maxInversionDepth :: HasOptions m => m Int
- getLanguage :: HasOptions m => m Language
- setHardCompileTimeModeIfErased' :: LensQuantity q => q -> TCM a -> TCM a
- setRunTimeModeUnlessInHardCompileTimeMode :: TCM a -> TCM a
- setModeUnlessInHardCompileTimeMode :: Erased -> TCM a -> TCM a
- warnForPlentyInHardCompileTimeMode :: Erased -> TCM ()
- addConstant :: QName -> Definition -> TCM ()
- addConstant' :: QName -> ArgInfo -> QName -> Type -> Defn -> TCM ()
- setTerminates :: MonadTCState m => QName -> Bool -> m ()
- setCompiledClauses :: QName -> CompiledClauses -> TCM ()
- setSplitTree :: QName -> SplitTree -> TCM ()
- modifyFunClauses :: QName -> ([Clause] -> [Clause]) -> TCM ()
- addClauses :: (MonadConstraint m, MonadTCState m) => QName -> [Clause] -> m ()
- mkPragma :: String -> TCM CompilerPragma
- addPragma :: BackendName -> QName -> String -> TCM ()
- getUniqueCompilerPragma :: BackendName -> QName -> TCM (Maybe CompilerPragma)
- setFunctionFlag :: FunctionFlag -> Bool -> QName -> TCM ()
- markStatic :: QName -> TCM ()
- markInline :: Bool -> QName -> TCM ()
- markInjective :: QName -> TCM ()
- markFirstOrder :: QName -> TCM ()
- unionSignatures :: [Signature] -> Signature
- addSection :: ModuleName -> TCM ()
- getSection :: (Functor m, ReadTCState m) => ModuleName -> m (Maybe Section)
- setModuleCheckpoint :: ModuleName -> TCM ()
- addDisplayForms :: QName -> TCM ()
- projectionArgs :: Definition -> Int
- addDisplayForm :: QName -> DisplayForm -> TCM ()
- applySection :: ModuleName -> Telescope -> ModuleName -> Args -> ScopeCopyInfo -> TCM ()
- applySection' :: ModuleName -> Telescope -> ModuleName -> Args -> ScopeCopyInfo -> TCM ()
- hasLoopingDisplayForm :: QName -> TCM Bool
- getDisplayForms :: (HasConstInfo m, ReadTCState m) => QName -> m [LocalDisplayForm]
- hasDisplayForms :: (HasConstInfo m, ReadTCState m) => QName -> m Bool
- chaseDisplayForms :: QName -> TCM (Set QName)
- canonicalName :: HasConstInfo m => QName -> m QName
- sameDef :: HasConstInfo m => QName -> QName -> m (Maybe QName)
- singleConstructorType :: QName -> TCM Bool
- prettySigCubicalNotErasure :: MonadPretty m => QName -> m Doc
- sigError :: (HasCallStack, MonadDebug m) => m a -> SigError -> m a
- getOriginalConstInfo :: (ReadTCState m, HasConstInfo m) => QName -> m Definition
- defaultGetRewriteRulesFor :: (ReadTCState m, MonadTCEnv m) => QName -> m RewriteRules
- getOriginalProjection :: HasConstInfo m => QName -> m QName
- isProjection :: HasConstInfo m => QName -> m (Maybe Projection)
- defaultGetConstInfo :: (HasOptions m, MonadDebug m, MonadTCEnv m) => TCState -> TCEnv -> QName -> m (Either SigError Definition)
- alwaysMakeAbstract :: Definition -> Maybe Definition
- getConInfo :: HasConstInfo m => ConHead -> m Definition
- getPolarity' :: HasConstInfo m => Comparison -> QName -> m [Polarity]
- setPolarity :: (MonadTCState m, MonadDebug m) => QName -> [Polarity] -> m ()
- getForcedArgs :: HasConstInfo m => QName -> m [IsForced]
- getArgOccurrence :: QName -> Nat -> TCM Occurrence
- setArgOccurrences :: MonadTCState m => QName -> [Occurrence] -> m ()
- modifyArgOccurrences :: MonadTCState m => QName -> ([Occurrence] -> [Occurrence]) -> m ()
- setTreeless :: QName -> TTerm -> TCM ()
- setCompiledArgUse :: QName -> [ArgUsage] -> TCM ()
- getCompiled :: HasConstInfo m => QName -> m (Maybe Compiled)
- setErasedConArgs :: QName -> [Bool] -> TCM ()
- getCompiledArgUse :: HasConstInfo m => QName -> m (Maybe [ArgUsage])
- addDataCons :: QName -> [QName] -> TCM ()
- getMutual :: QName -> TCM (Maybe [QName])
- getMutual_ :: Defn -> Maybe [QName]
- setMutual :: QName -> [QName] -> TCM ()
- mutuallyRecursive :: QName -> QName -> TCM Bool
- definitelyNonRecursive_ :: Defn -> Bool
- getCurrentModuleFreeVars :: TCM Nat
- getDefModule :: HasConstInfo m => QName -> m (Either SigError ModuleName)
- getDefFreeVars :: (Functor m, Applicative m, ReadTCState m, MonadTCEnv m) => QName -> m Nat
- getModuleFreeVars :: (Functor m, Applicative m, MonadTCEnv m, ReadTCState m) => ModuleName -> m Nat
- freeVarsToApply :: (Functor m, HasConstInfo m, HasOptions m, ReadTCState m, MonadTCEnv m, MonadDebug m) => QName -> m Args
- moduleParamsToApply :: (Functor m, Applicative m, HasOptions m, MonadTCEnv m, ReadTCState m, MonadDebug m) => ModuleName -> m Args
- instantiateDef :: (Functor m, HasConstInfo m, HasOptions m, ReadTCState m, MonadTCEnv m, MonadDebug m) => Definition -> m Definition
- instantiateRewriteRule :: (Functor m, HasConstInfo m, HasOptions m, ReadTCState m, MonadTCEnv m, MonadDebug m) => RewriteRule -> m RewriteRule
- instantiateRewriteRules :: (Functor m, HasConstInfo m, HasOptions m, ReadTCState m, MonadTCEnv m, MonadDebug m) => RewriteRules -> m RewriteRules
- inAbstractMode :: MonadTCEnv m => m a -> m a
- inConcreteMode :: MonadTCEnv m => m a -> m a
- ignoreAbstractMode :: MonadTCEnv m => m a -> m a
- underOpaqueId :: MonadTCEnv m => OpaqueId -> m a -> m a
- inConcreteOrAbstractMode :: (MonadTCEnv m, HasConstInfo m) => QName -> (Definition -> m a) -> m a
- relOfConst :: HasConstInfo m => QName -> m Relevance
- modalityOfConst :: HasConstInfo m => QName -> m Modality
- droppedPars :: Definition -> Int
- isProjection_ :: Defn -> Maybe Projection
- isRelevantProjection :: HasConstInfo m => QName -> m (Maybe Projection)
- isRelevantProjection_ :: Definition -> Maybe Projection
- isStaticFun :: Defn -> Bool
- isInlineFun :: Defn -> Bool
- isProperProjection :: Defn -> Bool
- usesCopatterns :: (HasConstInfo m, HasBuiltins m) => QName -> m Bool
- allMetaClasses :: [MetaClass]
- data KeepMetas
- dontAssignMetas :: (MonadTCEnv m, HasOptions m, MonadDebug m) => m a -> m a
- isRemoteMeta :: ReadTCState m => m (MetaId -> Bool)
- nextLocalMeta :: ReadTCState m => m MetaId
- data LocalMetaStores = LocalMetaStores {}
- metasCreatedBy :: ReadTCState m => m a -> m (a, LocalMetaStores)
- lookupLocalMeta' :: ReadTCState m => MetaId -> m (Maybe MetaVariable)
- lookupLocalMeta :: (HasCallStack, MonadDebug m, ReadTCState m) => MetaId -> m MetaVariable
- lookupMetaInstantiation :: ReadTCState m => MetaId -> m MetaInstantiation
- lookupMetaJudgement :: ReadTCState m => MetaId -> m (Judgement MetaId)
- lookupMetaModality :: ReadTCState m => MetaId -> m Modality
- updateMetaVarTCM :: HasCallStack => MetaId -> (MetaVariable -> MetaVariable) -> TCM ()
- isOpenMeta :: MetaInstantiation -> Bool
- insertMetaVar :: MetaId -> MetaVariable -> TCM ()
- getMetaPriority :: (HasCallStack, MonadDebug m, ReadTCState m) => MetaId -> m MetaPriority
- isSortMeta :: ReadTCState m => MetaId -> m Bool
- isSortJudgement :: Judgement a -> Bool
- isSortMeta_ :: MetaVariable -> Bool
- metaInstantiationToMetaKind :: MetaInstantiation -> MetaKind
- getMetaType :: ReadTCState m => MetaId -> m Type
- getMetaContextArgs :: MonadTCEnv m => MetaVariable -> m Args
- getMetaTypeInContext :: (HasBuiltins m, HasCallStack, MonadDebug m, MonadReduce m, MonadTCEnv m, ReadTCState m) => MetaId -> m Type
- isGeneralizableMeta :: (HasCallStack, MonadDebug m, ReadTCState m) => MetaId -> m DoGeneralize
- class IsInstantiatedMeta a where
- isInstantiatedMeta :: (MonadFail m, ReadTCState m) => a -> m Bool
- isInstantiatedMeta' :: (MonadFail m, ReadTCState m) => MetaId -> m (Maybe Term)
- constraintMetas :: Constraint -> TCM (Set MetaId)
- getMetaListeners :: (HasCallStack, MonadDebug m, ReadTCState m) => MetaId -> m [Listener]
- createMetaInfo :: (MonadTCEnv m, ReadTCState m) => m MetaInfo
- createMetaInfo' :: (MonadTCEnv m, ReadTCState m) => RunMetaOccursCheck -> m MetaInfo
- setValueMetaName :: MonadMetaSolver m => Term -> MetaNameSuggestion -> m ()
- setMetaNameSuggestion :: MonadMetaSolver m => MetaId -> MetaNameSuggestion -> m ()
- getMetaNameSuggestion :: (HasCallStack, MonadDebug m, ReadTCState m) => MetaId -> m MetaNameSuggestion
- setMetaGeneralizableArgInfo :: MonadMetaSolver m => MetaId -> ArgInfo -> m ()
- updateMetaVarRange :: MonadMetaSolver m => MetaId -> Range -> m ()
- setMetaOccursCheck :: MonadMetaSolver m => MetaId -> RunMetaOccursCheck -> m ()
- findInteractionPoint_ :: Range -> InteractionPoints -> Maybe InteractionId
- connectInteractionPoint :: MonadInteractionPoints m => InteractionId -> MetaId -> m ()
- removeInteractionPoint :: MonadInteractionPoints m => InteractionId -> m ()
- getInteractionPoints :: ReadTCState m => m [InteractionId]
- getInteractionMetas :: ReadTCState m => m [MetaId]
- getUniqueMetasRanges :: (HasCallStack, MonadDebug m, ReadTCState m) => [MetaId] -> m [Range]
- getUnsolvedMetas :: (HasCallStack, MonadDebug m, ReadTCState m) => m [Range]
- getOpenMetas :: ReadTCState m => m [MetaId]
- getUnsolvedInteractionMetas :: (HasCallStack, MonadDebug m, ReadTCState m) => m [Range]
- getInteractionIdsAndMetas :: ReadTCState m => m [(InteractionId, MetaId)]
- lookupInteractionPoint :: (MonadFail m, ReadTCState m, MonadError TCErr m) => InteractionId -> m InteractionPoint
- lookupInteractionId :: (MonadFail m, ReadTCState m, MonadError TCErr m, MonadTCEnv m) => InteractionId -> m MetaId
- lookupInteractionMeta :: ReadTCState m => InteractionId -> m (Maybe MetaId)
- lookupInteractionMeta_ :: InteractionId -> InteractionPoints -> Maybe MetaId
- newMeta :: MonadMetaSolver m => Frozen -> MetaInfo -> MetaPriority -> Permutation -> Judgement a -> m MetaId
- newMetaTCM' :: MetaInstantiation -> Frozen -> MetaInfo -> MetaPriority -> Permutation -> Judgement a -> TCM MetaId
- getInteractionScope :: (MonadDebug m, MonadFail m, ReadTCState m, MonadError TCErr m, MonadTCEnv m) => InteractionId -> m ScopeInfo
- withMetaInfo' :: (MonadTCEnv m, ReadTCState m, MonadTrace m) => MetaVariable -> m a -> m a
- withMetaInfo :: (MonadTCEnv m, ReadTCState m, MonadTrace m) => Closure Range -> m a -> m a
- listenToMeta :: MonadMetaSolver m => Listener -> MetaId -> m ()
- unlistenToMeta :: MonadMetaSolver m => Listener -> MetaId -> m ()
- clearMetaListeners :: MonadMetaSolver m => MetaId -> m ()
- etaExpandMetaSafe :: MonadMetaSolver m => MetaId -> m ()
- etaExpandListeners :: MonadMetaSolver m => MetaId -> m ()
- wakeupListener :: MonadMetaSolver m => Listener -> m ()
- solveAwakeConstraints :: MonadConstraint m => m ()
- solveAwakeConstraints' :: MonadConstraint m => Bool -> m ()
- freezeMetas :: MonadTCState m => LocalMetaStore -> m (Set MetaId)
- unfreezeMetas :: TCM ()
- isFrozen :: (HasCallStack, MonadDebug m, ReadTCState m) => MetaId -> m Bool
- withFrozenMetas :: (MonadMetaSolver m, MonadTCState m) => m a -> m a
- class UnFreezeMeta a where
- unfreezeMeta :: MonadMetaSolver m => a -> m ()
- data CheckResult where
- pattern CheckResult :: Interface -> [TCWarning] -> ModuleCheckMode -> Source -> CheckResult
- crInterface :: CheckResult -> Interface
- crWarnings :: CheckResult -> [TCWarning]
- crMode :: CheckResult -> ModuleCheckMode
- activeBackendMayEraseType :: QName -> TCM Bool
- backendInteraction :: AbsolutePath -> [Backend] -> TCM () -> (AbsolutePath -> TCM CheckResult) -> TCM ()
- parseBackendOptions :: [Backend] -> [String] -> CommandLineOptions -> OptM ([Backend], CommandLineOptions)
- callBackend :: String -> IsMain -> CheckResult -> TCM ()
- lookupBackend :: BackendName -> TCM (Maybe Backend)
- activeBackend :: TCM (Maybe Backend)
Documentation
module Agda.Compiler.Backend.Base
type Backend = Backend_boot TCM Source #
type Backend' opts env menv mod def = Backend'_boot TCM opts env menv mod def Source #
type Flag opts = opts -> OptM opts Source #
f :: Flag opts
is an action on the option record that results from
parsing an option. f opts
produces either an error message or an
updated options record
toTreeless :: EvaluationStrategy -> QName -> TCM (Maybe TTerm) Source #
module Agda.Syntax.Treeless
Type-checking errors.
TypeError | |
| |
Exception Range Doc | |
IOException TCState Range IOException | The first argument is the state in which the error was raised. |
PatternErr Blocker | The exception which is usually caught.
Raised for pattern violations during unification ( |
Instances
data Definition Source #
Defn | |
|
Instances
LensArgInfo Definition Source # | |||||
Defined in Agda.TypeChecking.Monad.Base getArgInfo :: Definition -> ArgInfo Source # setArgInfo :: ArgInfo -> Definition -> Definition Source # mapArgInfo :: (ArgInfo -> ArgInfo) -> Definition -> Definition Source # | |||||
LensModality Definition Source # | |||||
Defined in Agda.TypeChecking.Monad.Base getModality :: Definition -> Modality Source # setModality :: Modality -> Definition -> Definition Source # mapModality :: (Modality -> Modality) -> Definition -> Definition Source # | |||||
LensQuantity Definition Source # | |||||
Defined in Agda.TypeChecking.Monad.Base getQuantity :: Definition -> Quantity Source # setQuantity :: Quantity -> Definition -> Definition Source # mapQuantity :: (Quantity -> Quantity) -> Definition -> Definition Source # | |||||
LensRelevance Definition Source # | |||||
Defined in Agda.TypeChecking.Monad.Base getRelevance :: Definition -> Relevance Source # setRelevance :: Relevance -> Definition -> Definition Source # mapRelevance :: (Relevance -> Relevance) -> Definition -> Definition Source # | |||||
Pretty Definition Source # | |||||
Defined in Agda.TypeChecking.Monad.Base pretty :: Definition -> Doc Source # prettyPrec :: Int -> Definition -> Doc Source # prettyList :: [Definition] -> Doc Source # | |||||
NamesIn Definition Source # | |||||
Defined in Agda.Syntax.Internal.Names namesAndMetasIn' :: Monoid m => (Either QName MetaId -> m) -> Definition -> m Source # | |||||
KillRange Definition Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
KillRange Definitions Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
InstantiateFull Definition Source # | |||||
Defined in Agda.TypeChecking.Reduce | |||||
EmbPrj Definition Source # | |||||
Defined in Agda.TypeChecking.Serialise.Instances.Internal icode :: Definition -> S Int32 Source # icod_ :: Definition -> S Int32 Source # value :: Int32 -> R Definition Source # | |||||
Abstract Definition Source # | |||||
Defined in Agda.TypeChecking.Substitute abstract :: Telescope -> Definition -> Definition Source # | |||||
Apply Definition Source # | |||||
Defined in Agda.TypeChecking.Substitute apply :: Definition -> Args -> Definition Source # applyE :: Definition -> Elims -> Definition Source # | |||||
NFData Definition Source # | |||||
Defined in Agda.TypeChecking.Monad.Base rnf :: Definition -> () | |||||
Generic Definition Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
from :: Definition -> Rep Definition x to :: Rep Definition x -> Definition | |||||
Show Definition Source # | |||||
Defined in Agda.TypeChecking.Monad.Base showsPrec :: Int -> Definition -> ShowS show :: Definition -> String showList :: [Definition] -> ShowS | |||||
type Rep Definition Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep Definition = D1 ('MetaData "Definition" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "Defn" 'PrefixI 'True) ((((S1 ('MetaSel ('Just "defArgInfo") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ArgInfo) :*: S1 ('MetaSel ('Just "defName") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :*: (S1 ('MetaSel ('Just "defType") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type) :*: S1 ('MetaSel ('Just "defPolarity") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Polarity]))) :*: ((S1 ('MetaSel ('Just "defArgOccurrences") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Occurrence]) :*: S1 ('MetaSel ('Just "defArgGeneralizable") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 NumGeneralizableArgs)) :*: (S1 ('MetaSel ('Just "defGeneralizedParams") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Maybe Name]) :*: (S1 ('MetaSel ('Just "defDisplay") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [LocalDisplayForm]) :*: S1 ('MetaSel ('Just "defMutual") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 MutualId))))) :*: (((S1 ('MetaSel ('Just "defCompiledRep") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 CompiledRepresentation) :*: S1 ('MetaSel ('Just "defInstance") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe InstanceInfo))) :*: (S1 ('MetaSel ('Just "defCopy") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool) :*: (S1 ('MetaSel ('Just "defMatchable") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Set QName)) :*: S1 ('MetaSel ('Just "defNoCompilation") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool)))) :*: ((S1 ('MetaSel ('Just "defInjective") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool) :*: S1 ('MetaSel ('Just "defCopatternLHS") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool)) :*: (S1 ('MetaSel ('Just "defBlocked") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Blocked_) :*: (S1 ('MetaSel ('Just "defLanguage") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Language) :*: S1 ('MetaSel ('Just "theDef") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Defn))))))) |
Builtin Term | |
Prim pf | |
BuiltinRewriteRelations (Set QName) |
|
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 # | |||||
NamesIn a => NamesIn (Builtin a) Source # | |||||
Defined in Agda.Syntax.Internal.Names | |||||
InstantiateFull a => InstantiateFull (Builtin a) Source # | |||||
Defined in Agda.TypeChecking.Reduce | |||||
EmbPrj a => EmbPrj (Builtin a) Source # | |||||
NFData pf => NFData (Builtin pf) Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Generic (Builtin pf) Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
| |||||
Show pf => Show (Builtin pf) Source # | |||||
type Rep (Builtin pf) Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep (Builtin pf) = D1 ('MetaData "Builtin" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "Builtin" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)) :+: (C1 ('MetaCons "Prim" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 pf)) :+: C1 ('MetaCons "BuiltinRewriteRelations" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Set QName))))) |
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
NamesIn System Source # | |||||
Defined in Agda.Syntax.Internal.Names | |||||
KillRange System Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
InstantiateFull System Source # | |||||
Defined in Agda.TypeChecking.Reduce | |||||
EmbPrj System Source # | |||||
Abstract System Source # | |||||
Apply System Source # | |||||
NFData System Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Generic System Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
| |||||
Show System Source # | |||||
Reify (QNamed System) Source # | |||||
Defined in Agda.Syntax.Translation.InternalToAbstract
| |||||
type Rep System Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep System = D1 ('MetaData "System" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "System" 'PrefixI 'True) (S1 ('MetaSel ('Just "systemTel") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Telescope) :*: S1 ('MetaSel ('Just "systemClauses") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [(Face, Term)]))) | |||||
type ReifiesTo (QNamed System) Source # | |||||
Defined in Agda.Syntax.Translation.InternalToAbstract |
pattern Function :: [Clause] -> Maybe CompiledClauses -> Maybe SplitTree -> Maybe Compiled -> [Clause] -> FunctionInverse -> Maybe [QName] -> Either ProjectionLikenessMissing Projection -> SmallSet FunctionFlag -> Maybe Bool -> Maybe ExtLamInfo -> Maybe QName -> Maybe QName -> IsOpaque -> Defn Source #
newtype TCMT (m :: Type -> Type) a Source #
Instances
MonadFixityError ScopeM Source # | |||||
Defined in Agda.Syntax.Scope.Monad throwMultipleFixityDecls :: [(Name, [Fixity'])] -> ScopeM a Source # throwMultiplePolarityPragmas :: [Name] -> ScopeM a Source # warnUnknownNamesInFixityDecl :: [Name] -> ScopeM () Source # warnUnknownNamesInPolarityPragmas :: [Name] -> ScopeM () Source # warnUnknownFixityInMixfixDecl :: [Name] -> ScopeM () Source # warnPolarityPragmasButNotPostulates :: [Name] -> ScopeM () Source # | |||||
MonadBlock TCM Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
MonadStConcreteNames TCM Source # | |||||
Defined in Agda.TypeChecking.Monad.Base runStConcreteNames :: StateT ConcreteNames TCM a -> TCM a Source # useConcreteNames :: TCM ConcreteNames Source # modifyConcreteNames :: (ConcreteNames -> ConcreteNames) -> TCM () Source # | |||||
MonadConstraint TCM Source # | |||||
Defined in Agda.TypeChecking.Constraints addConstraint :: Blocker -> Constraint -> TCM () Source # addAwakeConstraint :: Blocker -> Constraint -> TCM () Source # solveConstraint :: Constraint -> TCM () Source # solveSomeAwakeConstraints :: (ProblemConstraint -> Bool) -> Bool -> TCM () Source # wakeConstraints :: (ProblemConstraint -> WakeUp) -> TCM () Source # stealConstraints :: ProblemId -> TCM () Source # modifyAwakeConstraints :: (Constraints -> Constraints) -> TCM () Source # modifySleepingConstraints :: (Constraints -> Constraints) -> TCM () Source # | |||||
MonadAddContext TCM Source # | |||||
Defined in Agda.TypeChecking.Monad.Context | |||||
MonadDebug TCM Source # | |||||
Defined in Agda.TypeChecking.Monad.Debug formatDebugMessage :: VerboseKey -> VerboseLevel -> TCM Doc -> TCM String Source # traceDebugMessage :: VerboseKey -> VerboseLevel -> String -> TCM a -> TCM a Source # verboseBracket :: VerboseKey -> VerboseLevel -> String -> TCM a -> TCM a Source # getVerbosity :: TCM Verbosity Source # getProfileOptions :: TCM ProfileOptions Source # isDebugPrinting :: TCM Bool Source # nowDebugPrinting :: TCM a -> TCM a Source # | |||||
MonadInteractionPoints TCM Source # | |||||
Defined in Agda.TypeChecking.Monad.MetaVars | |||||
MonadMetaSolver TCM Source # | |||||
Defined in Agda.TypeChecking.MetaVars newMeta' :: MetaInstantiation -> Frozen -> MetaInfo -> MetaPriority -> Permutation -> Judgement a -> TCM MetaId Source # assignV :: CompareDirection -> MetaId -> Args -> Term -> CompareAs -> TCM () Source # assignTerm' :: MetaId -> [Arg ArgName] -> Term -> TCM () Source # etaExpandMeta :: [MetaClass] -> MetaId -> TCM () Source # updateMetaVar :: MetaId -> (MetaVariable -> MetaVariable) -> TCM () Source # speculateMetas :: TCM () -> TCM KeepMetas -> TCM () Source # | |||||
PureTCM TCM Source # | |||||
Defined in Agda.TypeChecking.Monad.Pure | |||||
MonadStatistics TCM Source # | |||||
Defined in Agda.TypeChecking.Monad.Statistics modifyCounter :: String -> (Integer -> Integer) -> TCM () Source # | |||||
MonadTrace TCM Source # | |||||
Defined in Agda.TypeChecking.Monad.Trace traceCall :: Call -> TCM a -> TCM a Source # traceCallM :: TCM Call -> TCM a -> TCM a Source # traceCallCPS :: Call -> ((a -> TCM b) -> TCM b) -> (a -> TCM b) -> TCM b Source # traceClosureCall :: Closure Call -> TCM a -> TCM a Source # printHighlightingInfo :: RemoveTokenBasedHighlighting -> HighlightingInfo -> TCM () Source # | |||||
MonadWarning TCM Source # | |||||
Defined in Agda.TypeChecking.Warnings addWarning :: TCWarning -> TCM () Source # | |||||
MonadBench TCM Source # | We store benchmark statistics in an IORef. This enables benchmarking pure computation, see Agda.Benchmarking. | ||||
Defined in Agda.TypeChecking.Monad.Base
getBenchmark :: TCM (Benchmark (BenchPhase TCM)) Source # putBenchmark :: Benchmark (BenchPhase TCM) -> TCM () Source # modifyBenchmark :: (Benchmark (BenchPhase TCM) -> Benchmark (BenchPhase TCM)) -> TCM () Source # | |||||
CatchImpossible TCM Source # | Like The intended use is to catch internal errors during debug printing. In debug printing, we are not expecting state changes. | ||||
Defined in Agda.TypeChecking.Monad.Base catchImpossible :: TCM a -> (Impossible -> TCM a) -> TCM a Source # catchImpossibleJust :: (Impossible -> Maybe b) -> TCM a -> (b -> TCM a) -> TCM a Source # handleImpossible :: (Impossible -> TCM a) -> TCM a -> TCM a Source # handleImpossibleJust :: (Impossible -> Maybe b) -> (b -> TCM a) -> TCM a -> TCM a Source # | |||||
MonadTrans TCMT Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
HasFresh i => MonadFresh i TCM Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
MonadError TCErr TCM Source # | |||||
Defined in Agda.TypeChecking.Monad.Base throwError :: TCErr -> TCM a catchError :: TCM a -> (TCErr -> TCM a) -> TCM a | |||||
MonadIO m => HasOptions (TCMT m) Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
MonadIO m => MonadReduce (TCMT m) Source # | |||||
Defined in Agda.TypeChecking.Monad.Base liftReduce :: ReduceM a -> TCMT m a Source # | |||||
MonadIO m => MonadTCEnv (TCMT m) Source # | |||||
MonadIO m => MonadTCM (TCMT m) Source # | |||||
MonadIO m => MonadTCState (TCMT m) Source # | |||||
MonadIO m => ReadTCState (TCMT m) Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
MonadIO m => HasBuiltins (TCMT m) Source # | |||||
Defined in Agda.TypeChecking.Monad.Builtin getBuiltinThing :: SomeBuiltin -> TCMT m (Maybe (Builtin PrimFun)) Source # | |||||
ReportS (TCM Doc) Source # | |||||
Defined in Agda.TypeChecking.Monad.Debug reportS :: MonadDebug m => VerboseKey -> VerboseLevel -> TCM Doc -> m () Source # | |||||
ReportS [TCM Doc] Source # | |||||
Defined in Agda.TypeChecking.Monad.Debug reportS :: MonadDebug m => VerboseKey -> VerboseLevel -> [TCM Doc] -> m () Source # | |||||
TraceS (TCM Doc) Source # | |||||
Defined in Agda.TypeChecking.Monad.Debug traceS :: MonadDebug m => VerboseKey -> VerboseLevel -> TCM Doc -> m c -> m c Source # | |||||
TraceS [TCM Doc] Source # | |||||
Defined in Agda.TypeChecking.Monad.Debug traceS :: MonadDebug m => VerboseKey -> VerboseLevel -> [TCM Doc] -> m c -> m c Source # | |||||
HasConstInfo (TCMT IO) Source # | |||||
Defined in Agda.TypeChecking.Monad.Signature getConstInfo :: QName -> TCMT IO Definition Source # getConstInfo' :: QName -> TCMT IO (Either SigError Definition) Source # getRewriteRulesFor :: QName -> TCMT IO RewriteRules Source # | |||||
Null (TCM Doc) Source # | |||||
MonadIO m => MonadIO (TCMT m) Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Applicative m => Applicative (TCMT m) Source # | |||||
Functor m => Functor (TCMT m) Source # | |||||
Monad m => Monad (TCMT m) Source # | |||||
Semigroup (TCM Doc) Source # | This instance is more specific than a generic instance
| ||||
MonadIO m => MonadFail (TCMT m) Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
MonadFix m => MonadFix (TCMT m) Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
(MonadIO m, Null a) => Null (TCMT m a) Source # | |||||
(MonadIO m, Semigroup a, Monoid a) => Monoid (TCMT m a) Source # | Strict (non-shortcut) monoid. | ||||
(MonadIO m, Semigroup a) => Semigroup (TCMT m a) Source # | Strict (non-shortcut) semigroup. Note that there might be a lazy alternative, e.g.,
for TCM All we might want | ||||
(IsString a, MonadIO m) => IsString (TCMT m a) Source # | |||||
Defined in Agda.TypeChecking.Monad.Base fromString :: String -> TCMT m a | |||||
type BenchPhase TCM Source # | |||||
Defined in Agda.TypeChecking.Monad.Base |
class (Functor m, Applicative m, MonadFail m, HasOptions m, MonadDebug m, MonadTCEnv m) => HasConstInfo (m :: Type -> Type) where Source #
Nothing
getConstInfo :: QName -> m Definition Source #
Lookup the definition of a name. The result is a closed thing, all free variables have been abstracted over.
getConstInfo' :: QName -> m (Either SigError Definition) Source #
Version that reports exceptions:
default getConstInfo' :: forall (n :: Type -> Type) (t :: (Type -> Type) -> Type -> Type). (HasConstInfo n, MonadTrans t, m ~ t n) => QName -> m (Either SigError Definition) Source #
getRewriteRulesFor :: QName -> m RewriteRules Source #
Lookup the rewrite rules with the given head symbol.
default getRewriteRulesFor :: forall (n :: Type -> Type) (t :: (Type -> Type) -> Type -> Type). (HasConstInfo n, MonadTrans t, m ~ t n) => QName -> m RewriteRules Source #
Instances
HasConstInfo AbsToCon Source # | |
Defined in Agda.Syntax.Translation.AbstractToConcrete getConstInfo :: QName -> AbsToCon Definition Source # getConstInfo' :: QName -> AbsToCon (Either SigError Definition) Source # getRewriteRulesFor :: QName -> AbsToCon RewriteRules Source # | |
HasConstInfo TerM Source # | |
Defined in Agda.Termination.Monad getConstInfo :: QName -> TerM Definition Source # getConstInfo' :: QName -> TerM (Either SigError Definition) Source # | |
HasConstInfo ReduceM Source # | |
Defined in Agda.TypeChecking.Reduce.Monad getConstInfo :: QName -> ReduceM Definition Source # getConstInfo' :: QName -> ReduceM (Either SigError Definition) Source # getRewriteRulesFor :: QName -> ReduceM RewriteRules Source # | |
HasConstInfo NLM Source # | |
Defined in Agda.TypeChecking.Rewriting.NonLinMatch getConstInfo :: QName -> NLM Definition Source # getConstInfo' :: QName -> NLM (Either SigError Definition) Source # | |
HasConstInfo m => HasConstInfo (PureConversionT m) Source # | |
Defined in Agda.TypeChecking.Conversion.Pure getConstInfo :: QName -> PureConversionT m Definition Source # getConstInfo' :: QName -> PureConversionT m (Either SigError Definition) Source # getRewriteRulesFor :: QName -> PureConversionT m RewriteRules Source # | |
HasConstInfo m => HasConstInfo (BlockT m) Source # | |
Defined in Agda.TypeChecking.Monad.Signature getConstInfo :: QName -> BlockT m Definition Source # getConstInfo' :: QName -> BlockT m (Either SigError Definition) Source # getRewriteRulesFor :: QName -> BlockT m RewriteRules Source # | |
HasConstInfo (TCMT IO) Source # | |
Defined in Agda.TypeChecking.Monad.Signature getConstInfo :: QName -> TCMT IO Definition Source # getConstInfo' :: QName -> TCMT IO (Either SigError Definition) Source # getRewriteRulesFor :: QName -> TCMT IO RewriteRules Source # | |
HasConstInfo m => HasConstInfo (NamesT m) Source # | |
Defined in Agda.TypeChecking.Names getConstInfo :: QName -> NamesT m Definition Source # getConstInfo' :: QName -> NamesT m (Either SigError Definition) Source # getRewriteRulesFor :: QName -> NamesT m RewriteRules Source # | |
HasConstInfo m => HasConstInfo (ListT m) Source # | |
Defined in Agda.TypeChecking.Monad.Signature getConstInfo :: QName -> ListT m Definition Source # getConstInfo' :: QName -> ListT m (Either SigError Definition) Source # getRewriteRulesFor :: QName -> ListT m RewriteRules Source # | |
HasConstInfo m => HasConstInfo (ChangeT m) Source # | |
Defined in Agda.TypeChecking.Monad.Signature getConstInfo :: QName -> ChangeT m Definition Source # getConstInfo' :: QName -> ChangeT m (Either SigError Definition) Source # getRewriteRulesFor :: QName -> ChangeT m RewriteRules Source # | |
HasConstInfo m => HasConstInfo (MaybeT m) Source # | |
Defined in Agda.TypeChecking.Monad.Signature getConstInfo :: QName -> MaybeT m Definition Source # getConstInfo' :: QName -> MaybeT m (Either SigError Definition) Source # getRewriteRulesFor :: QName -> MaybeT m RewriteRules Source # | |
HasConstInfo m => HasConstInfo (ExceptT err m) Source # | |
Defined in Agda.TypeChecking.Monad.Signature getConstInfo :: QName -> ExceptT err m Definition Source # getConstInfo' :: QName -> ExceptT err m (Either SigError Definition) Source # getRewriteRulesFor :: QName -> ExceptT err m RewriteRules Source # | |
HasConstInfo m => HasConstInfo (IdentityT m) Source # | |
Defined in Agda.TypeChecking.Monad.Signature getConstInfo :: QName -> IdentityT m Definition Source # getConstInfo' :: QName -> IdentityT m (Either SigError Definition) Source # getRewriteRulesFor :: QName -> IdentityT m RewriteRules Source # | |
HasConstInfo m => HasConstInfo (ReaderT r m) Source # | |
Defined in Agda.TypeChecking.Monad.Signature getConstInfo :: QName -> ReaderT r m Definition Source # getConstInfo' :: QName -> ReaderT r m (Either SigError Definition) Source # getRewriteRulesFor :: QName -> ReaderT r m RewriteRules Source # | |
HasConstInfo m => HasConstInfo (StateT s m) Source # | |
Defined in Agda.TypeChecking.Monad.Signature getConstInfo :: QName -> StateT s m Definition Source # getConstInfo' :: QName -> StateT s m (Either SigError Definition) Source # getRewriteRulesFor :: QName -> StateT s m RewriteRules Source # | |
(Monoid w, HasConstInfo m) => HasConstInfo (WriterT w m) Source # | |
Defined in Agda.TypeChecking.Monad.Signature getConstInfo :: QName -> WriterT w m Definition Source # getConstInfo' :: QName -> WriterT w m (Either SigError Definition) Source # getRewriteRulesFor :: QName -> WriterT w m RewriteRules Source # |
getErasedConArgs :: HasConstInfo m => QName -> m [Bool] Source #
Returns a list of length conArity
.
If no erasure analysis has been performed yet, this will be a list of False
s.
getTreeless :: HasConstInfo m => QName -> m (Maybe TTerm) Source #
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 !Erased !ModuleName !Telescope | |
LeaveSection !ModuleName | |
Decl !Declaration | Never a Section or ScopeDecl |
Pragmas !PragmaOptions |
Instances
NFData TypeCheckAction Source # | |||||
Defined in Agda.TypeChecking.Monad.Base rnf :: TypeCheckAction -> () | |||||
Generic TypeCheckAction Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
from :: TypeCheckAction -> Rep TypeCheckAction x to :: Rep TypeCheckAction x -> TypeCheckAction | |||||
type Rep TypeCheckAction Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep TypeCheckAction = D1 ('MetaData "TypeCheckAction" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) ((C1 ('MetaCons "EnterSection" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Erased) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 ModuleName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Telescope))) :+: C1 ('MetaCons "LeaveSection" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 ModuleName))) :+: (C1 ('MetaCons "Decl" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Declaration)) :+: C1 ('MetaCons "Pragmas" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 PragmaOptions)))) |
pattern Primitive :: IsAbstract -> PrimitiveId -> [Clause] -> FunctionInverse -> Maybe CompiledClauses -> IsOpaque -> Defn Source #
Various classes of metavariables.
Records | Meta variables of record type. |
SingletonRecords | Meta variables of "hereditarily singleton" record type. |
Levels | Meta variables of level type, if type-in-type is activated. |
Instances
Bounded MetaClass Source # | |
Defined in Agda.TypeChecking.Monad.MetaVars | |
Enum MetaClass Source # | |
Defined in Agda.TypeChecking.Monad.MetaVars | |
Show MetaClass Source # | |
Eq MetaClass Source # | |
Instances
HasOptions ReduceM Source # | |
MonadReduce ReduceM Source # | |
Defined in Agda.TypeChecking.Monad.Base liftReduce :: ReduceM a -> ReduceM a Source # | |
MonadTCEnv ReduceM Source # | |
ReadTCState ReduceM Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
HasBuiltins ReduceM Source # | |
Defined in Agda.TypeChecking.Reduce.Monad getBuiltinThing :: SomeBuiltin -> ReduceM (Maybe (Builtin PrimFun)) Source # | |
MonadAddContext ReduceM Source # | |
Defined in Agda.TypeChecking.Reduce.Monad addCtx :: Name -> Dom Type -> ReduceM a -> ReduceM a Source # addLetBinding' :: Origin -> Name -> Term -> Dom Type -> ReduceM a -> ReduceM a Source # updateContext :: Substitution -> (Context -> Context) -> ReduceM a -> ReduceM a Source # withFreshName :: Range -> ArgName -> (Name -> ReduceM a) -> ReduceM a Source # | |
MonadDebug ReduceM Source # | |
Defined in Agda.TypeChecking.Reduce.Monad formatDebugMessage :: VerboseKey -> VerboseLevel -> TCM Doc -> ReduceM String Source # traceDebugMessage :: VerboseKey -> VerboseLevel -> String -> ReduceM a -> ReduceM a Source # verboseBracket :: VerboseKey -> VerboseLevel -> String -> ReduceM a -> ReduceM a Source # getVerbosity :: ReduceM Verbosity Source # getProfileOptions :: ReduceM ProfileOptions Source # isDebugPrinting :: ReduceM Bool Source # nowDebugPrinting :: ReduceM a -> ReduceM a Source # | |
PureTCM ReduceM Source # | |
Defined in Agda.TypeChecking.Reduce.Monad | |
HasConstInfo ReduceM Source # | |
Defined in Agda.TypeChecking.Reduce.Monad getConstInfo :: QName -> ReduceM Definition Source # getConstInfo' :: QName -> ReduceM (Either SigError Definition) Source # getRewriteRulesFor :: QName -> ReduceM RewriteRules Source # | |
Applicative ReduceM Source # | |
Functor ReduceM Source # | |
Monad ReduceM Source # | |
MonadFail ReduceM Source # | |
Defined in Agda.TypeChecking.Monad.Base |
runReduceM :: ReduceM a -> TCM a Source #
class (Functor m, Applicative m, Monad m) => MonadDebug (m :: Type -> Type) where Source #
Nothing
formatDebugMessage :: VerboseKey -> VerboseLevel -> TCM Doc -> m String Source #
default formatDebugMessage :: forall (t :: (Type -> Type) -> Type -> Type) (n :: Type -> Type). (MonadTrans t, MonadDebug n, m ~ t n) => VerboseKey -> VerboseLevel -> TCM Doc -> m String Source #
traceDebugMessage :: VerboseKey -> VerboseLevel -> String -> m a -> m a Source #
default traceDebugMessage :: forall (t :: (Type -> Type) -> Type -> Type) (n :: Type -> Type) a. (MonadTransControl t, MonadDebug n, m ~ t n) => VerboseKey -> VerboseLevel -> String -> m a -> m a Source #
verboseBracket :: VerboseKey -> VerboseLevel -> String -> m a -> m a Source #
Print brackets around debug messages issued by a computation.
default verboseBracket :: forall (t :: (Type -> Type) -> Type -> Type) (n :: Type -> Type) a. (MonadTransControl t, MonadDebug n, m ~ t n) => VerboseKey -> VerboseLevel -> String -> m a -> m a Source #
getVerbosity :: m Verbosity Source #
default getVerbosity :: forall (t :: (Type -> Type) -> Type -> Type) (n :: Type -> Type). (MonadTrans t, MonadDebug n, m ~ t n) => m Verbosity Source #
getProfileOptions :: m ProfileOptions Source #
default getProfileOptions :: forall (t :: (Type -> Type) -> Type -> Type) (n :: Type -> Type). (MonadTrans t, MonadDebug n, m ~ t n) => m ProfileOptions Source #
isDebugPrinting :: m Bool Source #
Check whether we are currently debug printing.
default isDebugPrinting :: forall (t :: (Type -> Type) -> Type -> Type) (n :: Type -> Type). (MonadTrans t, MonadDebug n, m ~ t n) => m Bool Source #
nowDebugPrinting :: m a -> m a Source #
Flag in a computation that we are currently debug printing.
default nowDebugPrinting :: forall (t :: (Type -> Type) -> Type -> Type) (n :: Type -> Type) a. (MonadTransControl t, MonadDebug n, m ~ t n) => m a -> m a Source #
Instances
MonadDebug AbsToCon Source # | |
Defined in Agda.Syntax.Translation.AbstractToConcrete formatDebugMessage :: VerboseKey -> VerboseLevel -> TCM Doc -> AbsToCon String Source # traceDebugMessage :: VerboseKey -> VerboseLevel -> String -> AbsToCon a -> AbsToCon a Source # verboseBracket :: VerboseKey -> VerboseLevel -> String -> AbsToCon a -> AbsToCon a Source # getVerbosity :: AbsToCon Verbosity Source # getProfileOptions :: AbsToCon ProfileOptions Source # isDebugPrinting :: AbsToCon Bool Source # nowDebugPrinting :: AbsToCon a -> AbsToCon a Source # | |
MonadDebug TerM Source # | |
Defined in Agda.Termination.Monad formatDebugMessage :: VerboseKey -> VerboseLevel -> TCM Doc -> TerM String Source # traceDebugMessage :: VerboseKey -> VerboseLevel -> String -> TerM a -> TerM a Source # verboseBracket :: VerboseKey -> VerboseLevel -> String -> TerM a -> TerM a Source # getVerbosity :: TerM Verbosity Source # getProfileOptions :: TerM ProfileOptions Source # isDebugPrinting :: TerM Bool Source # nowDebugPrinting :: TerM a -> TerM a Source # | |
MonadDebug ReduceM Source # | |
Defined in Agda.TypeChecking.Reduce.Monad formatDebugMessage :: VerboseKey -> VerboseLevel -> TCM Doc -> ReduceM String Source # traceDebugMessage :: VerboseKey -> VerboseLevel -> String -> ReduceM a -> ReduceM a Source # verboseBracket :: VerboseKey -> VerboseLevel -> String -> ReduceM a -> ReduceM a Source # getVerbosity :: ReduceM Verbosity Source # getProfileOptions :: ReduceM ProfileOptions Source # isDebugPrinting :: ReduceM Bool Source # nowDebugPrinting :: ReduceM a -> ReduceM a Source # | |
MonadDebug TCM Source # | |
Defined in Agda.TypeChecking.Monad.Debug formatDebugMessage :: VerboseKey -> VerboseLevel -> TCM Doc -> TCM String Source # traceDebugMessage :: VerboseKey -> VerboseLevel -> String -> TCM a -> TCM a Source # verboseBracket :: VerboseKey -> VerboseLevel -> String -> TCM a -> TCM a Source # getVerbosity :: TCM Verbosity Source # getProfileOptions :: TCM ProfileOptions Source # isDebugPrinting :: TCM Bool Source # nowDebugPrinting :: TCM a -> TCM a Source # | |
MonadDebug NLM Source # | |
Defined in Agda.TypeChecking.Rewriting.NonLinMatch formatDebugMessage :: VerboseKey -> VerboseLevel -> TCM Doc -> NLM String Source # traceDebugMessage :: VerboseKey -> VerboseLevel -> String -> NLM a -> NLM a Source # verboseBracket :: VerboseKey -> VerboseLevel -> String -> NLM a -> NLM a Source # getVerbosity :: NLM Verbosity Source # getProfileOptions :: NLM ProfileOptions Source # isDebugPrinting :: NLM Bool Source # nowDebugPrinting :: NLM a -> NLM a Source # | |
MonadDebug m => MonadDebug (PureConversionT m) Source # | |
Defined in Agda.TypeChecking.Conversion.Pure formatDebugMessage :: VerboseKey -> VerboseLevel -> TCM Doc -> PureConversionT m String Source # traceDebugMessage :: VerboseKey -> VerboseLevel -> String -> PureConversionT m a -> PureConversionT m a Source # verboseBracket :: VerboseKey -> VerboseLevel -> String -> PureConversionT m a -> PureConversionT m a Source # getVerbosity :: PureConversionT m Verbosity Source # getProfileOptions :: PureConversionT m ProfileOptions Source # isDebugPrinting :: PureConversionT m Bool Source # nowDebugPrinting :: PureConversionT m a -> PureConversionT m a Source # | |
MonadDebug m => MonadDebug (BlockT m) Source # | |
Defined in Agda.TypeChecking.Monad.Debug formatDebugMessage :: VerboseKey -> VerboseLevel -> TCM Doc -> BlockT m String Source # traceDebugMessage :: VerboseKey -> VerboseLevel -> String -> BlockT m a -> BlockT m a Source # verboseBracket :: VerboseKey -> VerboseLevel -> String -> BlockT m a -> BlockT m a Source # getVerbosity :: BlockT m Verbosity Source # getProfileOptions :: BlockT m ProfileOptions Source # isDebugPrinting :: BlockT m Bool Source # nowDebugPrinting :: BlockT m a -> BlockT m a Source # | |
MonadDebug m => MonadDebug (NamesT m) Source # | |
Defined in Agda.TypeChecking.Names formatDebugMessage :: VerboseKey -> VerboseLevel -> TCM Doc -> NamesT m String Source # traceDebugMessage :: VerboseKey -> VerboseLevel -> String -> NamesT m a -> NamesT m a Source # verboseBracket :: VerboseKey -> VerboseLevel -> String -> NamesT m a -> NamesT m a Source # getVerbosity :: NamesT m Verbosity Source # getProfileOptions :: NamesT m ProfileOptions Source # isDebugPrinting :: NamesT m Bool Source # nowDebugPrinting :: NamesT m a -> NamesT m a Source # | |
MonadDebug m => MonadDebug (ListT m) Source # | |
Defined in Agda.TypeChecking.Monad.Debug formatDebugMessage :: VerboseKey -> VerboseLevel -> TCM Doc -> ListT m String Source # traceDebugMessage :: VerboseKey -> VerboseLevel -> String -> ListT m a -> ListT m a Source # verboseBracket :: VerboseKey -> VerboseLevel -> String -> ListT m a -> ListT m a Source # getVerbosity :: ListT m Verbosity Source # getProfileOptions :: ListT m ProfileOptions Source # isDebugPrinting :: ListT m Bool Source # nowDebugPrinting :: ListT m a -> ListT m a Source # | |
MonadDebug m => MonadDebug (ChangeT m) Source # | |
Defined in Agda.TypeChecking.Monad.Debug formatDebugMessage :: VerboseKey -> VerboseLevel -> TCM Doc -> ChangeT m String Source # traceDebugMessage :: VerboseKey -> VerboseLevel -> String -> ChangeT m a -> ChangeT m a Source # verboseBracket :: VerboseKey -> VerboseLevel -> String -> ChangeT m a -> ChangeT m a Source # getVerbosity :: ChangeT m Verbosity Source # getProfileOptions :: ChangeT m ProfileOptions Source # isDebugPrinting :: ChangeT m Bool Source # nowDebugPrinting :: ChangeT m a -> ChangeT m a Source # | |
MonadDebug m => MonadDebug (MaybeT m) Source # | |
Defined in Agda.TypeChecking.Monad.Debug formatDebugMessage :: VerboseKey -> VerboseLevel -> TCM Doc -> MaybeT m String Source # traceDebugMessage :: VerboseKey -> VerboseLevel -> String -> MaybeT m a -> MaybeT m a Source # verboseBracket :: VerboseKey -> VerboseLevel -> String -> MaybeT m a -> MaybeT m a Source # getVerbosity :: MaybeT m Verbosity Source # getProfileOptions :: MaybeT m ProfileOptions Source # isDebugPrinting :: MaybeT m Bool Source # nowDebugPrinting :: MaybeT m a -> MaybeT m a Source # | |
MonadDebug m => MonadDebug (ExceptT e m) Source # | |
Defined in Agda.TypeChecking.Monad.Debug formatDebugMessage :: VerboseKey -> VerboseLevel -> TCM Doc -> ExceptT e m String Source # traceDebugMessage :: VerboseKey -> VerboseLevel -> String -> ExceptT e m a -> ExceptT e m a Source # verboseBracket :: VerboseKey -> VerboseLevel -> String -> ExceptT e m a -> ExceptT e m a Source # getVerbosity :: ExceptT e m Verbosity Source # getProfileOptions :: ExceptT e m ProfileOptions Source # isDebugPrinting :: ExceptT e m Bool Source # nowDebugPrinting :: ExceptT e m a -> ExceptT e m a Source # | |
MonadDebug m => MonadDebug (IdentityT m) Source # | |
Defined in Agda.TypeChecking.Monad.Debug formatDebugMessage :: VerboseKey -> VerboseLevel -> TCM Doc -> IdentityT m String Source # traceDebugMessage :: VerboseKey -> VerboseLevel -> String -> IdentityT m a -> IdentityT m a Source # verboseBracket :: VerboseKey -> VerboseLevel -> String -> IdentityT m a -> IdentityT m a Source # getVerbosity :: IdentityT m Verbosity Source # getProfileOptions :: IdentityT m ProfileOptions Source # isDebugPrinting :: IdentityT m Bool Source # nowDebugPrinting :: IdentityT m a -> IdentityT m a Source # | |
MonadDebug m => MonadDebug (ReaderT r m) Source # | |
Defined in Agda.TypeChecking.Monad.Debug formatDebugMessage :: VerboseKey -> VerboseLevel -> TCM Doc -> ReaderT r m String Source # traceDebugMessage :: VerboseKey -> VerboseLevel -> String -> ReaderT r m a -> ReaderT r m a Source # verboseBracket :: VerboseKey -> VerboseLevel -> String -> ReaderT r m a -> ReaderT r m a Source # getVerbosity :: ReaderT r m Verbosity Source # getProfileOptions :: ReaderT r m ProfileOptions Source # isDebugPrinting :: ReaderT r m Bool Source # nowDebugPrinting :: ReaderT r m a -> ReaderT r m a Source # | |
MonadDebug m => MonadDebug (StateT s m) Source # | |
Defined in Agda.TypeChecking.Monad.Debug formatDebugMessage :: VerboseKey -> VerboseLevel -> TCM Doc -> StateT s m String Source # traceDebugMessage :: VerboseKey -> VerboseLevel -> String -> StateT s m a -> StateT s m a Source # verboseBracket :: VerboseKey -> VerboseLevel -> String -> StateT s m a -> StateT s m a Source # getVerbosity :: StateT s m Verbosity Source # getProfileOptions :: StateT s m ProfileOptions Source # isDebugPrinting :: StateT s m Bool Source # nowDebugPrinting :: StateT s m a -> StateT s m a Source # | |
(MonadDebug m, Monoid w) => MonadDebug (WriterT w m) Source # | |
Defined in Agda.TypeChecking.Monad.Debug formatDebugMessage :: VerboseKey -> VerboseLevel -> TCM Doc -> WriterT w m String Source # traceDebugMessage :: VerboseKey -> VerboseLevel -> String -> WriterT w m a -> WriterT w m a Source # verboseBracket :: VerboseKey -> VerboseLevel -> String -> WriterT w m a -> WriterT w m a Source # getVerbosity :: WriterT w m Verbosity Source # getProfileOptions :: WriterT w m ProfileOptions Source # isDebugPrinting :: WriterT w m Bool Source # nowDebugPrinting :: WriterT w m a -> WriterT w m a Source # |
__IMPOSSIBLE_VERBOSE__ :: (HasCallStack, MonadDebug m) => String -> m a Source #
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
NFData HighlightingLevel Source # | |||||
Defined in Agda.TypeChecking.Monad.Base rnf :: HighlightingLevel -> () | |||||
Generic HighlightingLevel Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
from :: HighlightingLevel -> Rep HighlightingLevel x to :: Rep HighlightingLevel x -> HighlightingLevel | |||||
Read HighlightingLevel Source # | |||||
Defined in Agda.TypeChecking.Monad.Base readsPrec :: Int -> ReadS HighlightingLevel readList :: ReadS [HighlightingLevel] readPrec :: ReadPrec HighlightingLevel readListPrec :: ReadPrec [HighlightingLevel] | |||||
Show HighlightingLevel Source # | |||||
Defined in Agda.TypeChecking.Monad.Base showsPrec :: Int -> HighlightingLevel -> ShowS show :: HighlightingLevel -> String showList :: [HighlightingLevel] -> ShowS | |||||
Eq HighlightingLevel Source # | |||||
Defined in Agda.TypeChecking.Monad.Base (==) :: HighlightingLevel -> HighlightingLevel -> Bool (/=) :: HighlightingLevel -> HighlightingLevel -> Bool | |||||
Ord HighlightingLevel Source # | |||||
Defined in Agda.TypeChecking.Monad.Base compare :: HighlightingLevel -> HighlightingLevel -> Ordering (<) :: HighlightingLevel -> HighlightingLevel -> Bool (<=) :: HighlightingLevel -> HighlightingLevel -> Bool (>) :: HighlightingLevel -> HighlightingLevel -> Bool (>=) :: HighlightingLevel -> HighlightingLevel -> Bool max :: HighlightingLevel -> HighlightingLevel -> HighlightingLevel min :: HighlightingLevel -> HighlightingLevel -> HighlightingLevel | |||||
type Rep HighlightingLevel Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep HighlightingLevel = D1 ('MetaData "HighlightingLevel" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "None" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "NonInteractive" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "Interactive" 'PrefixI 'False) (U1 :: Type -> Type))) |
data HighlightingMethod Source #
How should highlighting be sent to the user interface?
Instances
NFData HighlightingMethod Source # | |||||
Defined in Agda.TypeChecking.Monad.Base rnf :: HighlightingMethod -> () | |||||
Generic HighlightingMethod Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
from :: HighlightingMethod -> Rep HighlightingMethod x to :: Rep HighlightingMethod x -> HighlightingMethod | |||||
Read HighlightingMethod Source # | |||||
Defined in Agda.TypeChecking.Monad.Base readsPrec :: Int -> ReadS HighlightingMethod readList :: ReadS [HighlightingMethod] readPrec :: ReadPrec HighlightingMethod readListPrec :: ReadPrec [HighlightingMethod] | |||||
Show HighlightingMethod Source # | |||||
Defined in Agda.TypeChecking.Monad.Base showsPrec :: Int -> HighlightingMethod -> ShowS show :: HighlightingMethod -> String showList :: [HighlightingMethod] -> ShowS | |||||
Eq HighlightingMethod Source # | |||||
Defined in Agda.TypeChecking.Monad.Base (==) :: HighlightingMethod -> HighlightingMethod -> Bool (/=) :: HighlightingMethod -> HighlightingMethod -> Bool | |||||
type Rep HighlightingMethod Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep HighlightingMethod = D1 ('MetaData "HighlightingMethod" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "Direct" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "Indirect" 'PrefixI 'False) (U1 :: Type -> Type)) |
data Comparison Source #
Instances
Pretty Comparison Source # | |||||
Defined in Agda.TypeChecking.Monad.Base pretty :: Comparison -> Doc Source # prettyPrec :: Int -> Comparison -> Doc Source # prettyList :: [Comparison] -> Doc Source # | |||||
PrettyTCM Comparison Source # | |||||
Defined in Agda.TypeChecking.Pretty prettyTCM :: MonadPretty m => Comparison -> m Doc Source # | |||||
EmbPrj Comparison Source # | |||||
Defined in Agda.TypeChecking.Serialise.Instances.Internal icode :: Comparison -> S Int32 Source # icod_ :: Comparison -> S Int32 Source # value :: Int32 -> R Comparison Source # | |||||
NFData Comparison Source # | |||||
Defined in Agda.TypeChecking.Monad.Base rnf :: Comparison -> () | |||||
Generic Comparison Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
from :: Comparison -> Rep Comparison x to :: Rep Comparison x -> Comparison | |||||
Show Comparison Source # | |||||
Defined in Agda.TypeChecking.Monad.Base showsPrec :: Int -> Comparison -> ShowS show :: Comparison -> String showList :: [Comparison] -> ShowS | |||||
Eq Comparison Source # | |||||
Defined in Agda.TypeChecking.Monad.Base (==) :: Comparison -> Comparison -> Bool (/=) :: Comparison -> Comparison -> Bool | |||||
type Rep Comparison Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep Comparison = D1 ('MetaData "Comparison" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "CmpEq" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "CmpLeq" 'PrefixI 'False) (U1 :: Type -> Type)) |
Polarity for equality and subtype checking.
Covariant | monotone |
Contravariant | antitone |
Invariant | no information (mixed variance) |
Nonvariant | constant |
Instances
Pretty Polarity Source # | |||||
KillRange Polarity Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
PrettyTCM Polarity Source # | |||||
Defined in Agda.TypeChecking.Pretty | |||||
EmbPrj Polarity Source # | |||||
NFData Polarity Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Generic Polarity Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
| |||||
Show Polarity Source # | |||||
Eq Polarity Source # | |||||
Abstract [Polarity] Source # | |||||
Apply [Polarity] Source # | |||||
type Rep Polarity Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep Polarity = D1 ('MetaData "Polarity" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) ((C1 ('MetaCons "Covariant" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "Contravariant" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "Invariant" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "Nonvariant" 'PrefixI 'False) (U1 :: Type -> Type))) |
A thing tagged with the context it came from. Also keeps the substitution from previous checkpoints. This lets us handle the case when an open thing was created in a context that we have since exited. Remember which module it's from to make sure we don't get confused by checkpoints from other files.
Instances
Decoration Open Source # | |||||
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 # | |||||
Pretty a => Pretty (Open a) Source # | |||||
NamesIn a => NamesIn (Open a) Source # | |||||
Defined in Agda.Syntax.Internal.Names | |||||
KillRange a => KillRange (Open a) Source # | |||||
Defined in Agda.TypeChecking.Monad.Base killRange :: KillRangeT (Open a) Source # | |||||
InstantiateFull t => InstantiateFull (Open t) Source # | |||||
Defined in Agda.TypeChecking.Reduce | |||||
EmbPrj a => EmbPrj (Open a) Source # | |||||
NFData a => NFData (Open a) Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Generic (Open a) Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
| |||||
Show a => Show (Open a) Source # | |||||
type Rep (Open a) Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep (Open a) = D1 ('MetaData "Open" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "OpenThing" 'PrefixI 'True) ((S1 ('MetaSel ('Just "openThingCheckpoint") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 CheckpointId) :*: S1 ('MetaSel ('Just "openThingCheckpointMap") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Map CheckpointId Substitution))) :*: (S1 ('MetaSel ('Just "openThingModule") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ModuleNameHash) :*: S1 ('MetaSel ('Just "openThing") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 a)))) |
MetaInfo
is cloned from one meta to the next during pruning.
MetaInfo | |
|
Instances
LensIsAbstract MetaInfo Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
LensModality MetaInfo Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
LensQuantity MetaInfo Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
LensRelevance MetaInfo Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
HasRange MetaInfo Source # | |||||
SetRange MetaInfo Source # | |||||
NFData MetaInfo Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Generic MetaInfo Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
| |||||
LensClosure MetaInfo Range Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
type Rep MetaInfo Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep MetaInfo = D1 ('MetaData "MetaInfo" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "MetaInfo" 'PrefixI 'True) ((S1 ('MetaSel ('Just "miClosRange") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Closure Range)) :*: S1 ('MetaSel ('Just "miModality") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Modality)) :*: (S1 ('MetaSel ('Just "miMetaOccursCheck") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 RunMetaOccursCheck) :*: (S1 ('MetaSel ('Just "miNameSuggestion") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 MetaNameSuggestion) :*: S1 ('MetaSel ('Just "miGeneralizable") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Arg DoGeneralize)))))) |
type Constraints = [ProblemConstraint] Source #
data Constraint Source #
ValueCmp Comparison CompareAs Term Term | |
ValueCmpOnFace Comparison Term Type Term Term | |
ElimCmp [Polarity] [IsForced] Type Term [Elim] [Elim] | |
SortCmp Comparison Sort Sort | |
LevelCmp Comparison Level Level | |
HasBiggerSort Sort | |
HasPTSRule (Dom Type) (Abs Sort) | |
CheckDataSort QName Sort | Check that the sort |
CheckMetaInst MetaId | |
CheckType Type | |
UnBlock MetaId | Meta created for a term blocked by a postponed type checking problem or unsolved
constraints. The |
IsEmpty Range Type | The range is the one of the absurd pattern. |
CheckSizeLtSat Term | Check that the |
FindInstance MetaId (Maybe [Candidate]) | the first argument is the instance argument and the second one is the list of candidates (or Nothing if we haven’t determined the list of candidates yet) |
ResolveInstanceHead QName | Resolve the head symbol of the type that the given instance targets |
CheckFunDef DefInfo QName [Clause] TCErr | Last argument is the error causing us to postpone. |
UnquoteTactic Term Term Type | First argument is computation and the others are hole and goal type |
CheckLockedVars Term Type (Arg Term) Type |
|
UsableAtModality WhyCheckModality (Maybe Sort) Modality Term | Is the term usable at the given modality?
This check should run if the |
Instances
TermLike Constraint Source # | |||||
Defined in Agda.TypeChecking.Monad.Base traverseTermM :: Monad m => (Term -> m Term) -> Constraint -> m Constraint Source # foldTerm :: Monoid m => (Term -> m) -> Constraint -> m Source # | |||||
AllMetas Constraint Source # | |||||
Defined in Agda.TypeChecking.Monad.Base allMetas :: Monoid m => (MetaId -> m) -> Constraint -> m Source # | |||||
HasRange Constraint Source # | |||||
Defined in Agda.TypeChecking.Monad.Base getRange :: Constraint -> Range Source # | |||||
Reify Constraint Source # | |||||
Defined in Agda.Interaction.BasicOps
reify :: MonadReify m => Constraint -> m (ReifiesTo Constraint) Source # reifyWhen :: MonadReify m => Bool -> Constraint -> m (ReifiesTo Constraint) Source # | |||||
Free Constraint Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
MentionsMeta Constraint Source # | |||||
Defined in Agda.TypeChecking.MetaVars.Mention mentionsMetas :: HashSet MetaId -> Constraint -> Bool Source # | |||||
PrettyTCM Constraint Source # | |||||
Defined in Agda.TypeChecking.Pretty.Constraint prettyTCM :: MonadPretty m => Constraint -> m Doc Source # | |||||
Instantiate Constraint Source # | |||||
Defined in Agda.TypeChecking.Reduce | |||||
InstantiateFull Constraint Source # | |||||
Defined in Agda.TypeChecking.Reduce | |||||
Normalise Constraint Source # | |||||
Defined in Agda.TypeChecking.Reduce normalise' :: Constraint -> ReduceM Constraint Source # | |||||
Reduce Constraint Source # | |||||
Defined in Agda.TypeChecking.Reduce reduce' :: Constraint -> ReduceM Constraint Source # reduceB' :: Constraint -> ReduceM (Blocked Constraint) Source # | |||||
Simplify Constraint Source # | |||||
Defined in Agda.TypeChecking.Reduce simplify' :: Constraint -> ReduceM Constraint | |||||
Subst Constraint Source # | |||||
Defined in Agda.TypeChecking.Substitute
applySubst :: Substitution' (SubstArg Constraint) -> Constraint -> Constraint Source # | |||||
NFData Constraint Source # | |||||
Defined in Agda.TypeChecking.Monad.Base rnf :: Constraint -> () | |||||
Generic Constraint Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
from :: Constraint -> Rep Constraint x to :: Rep Constraint x -> Constraint | |||||
Show Constraint Source # | |||||
Defined in Agda.TypeChecking.Monad.Base showsPrec :: Int -> Constraint -> ShowS show :: Constraint -> String showList :: [Constraint] -> ShowS | |||||
type ReifiesTo Constraint Source # | |||||
Defined in Agda.Interaction.BasicOps | |||||
type SubstArg Constraint Source # | |||||
Defined in Agda.TypeChecking.Substitute | |||||
type Rep Constraint Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep Constraint = D1 ('MetaData "Constraint" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) ((((C1 ('MetaCons "ValueCmp" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 CompareAs)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term))) :+: C1 ('MetaCons "ValueCmpOnFace" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term))))) :+: (C1 ('MetaCons "ElimCmp" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Polarity]) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [IsForced]) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type))) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Elim]) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Elim])))) :+: C1 ('MetaCons "SortCmp" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Sort) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Sort))))) :+: ((C1 ('MetaCons "LevelCmp" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Level) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Level))) :+: C1 ('MetaCons "HasBiggerSort" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Sort))) :+: (C1 ('MetaCons "HasPTSRule" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Dom Type)) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Abs Sort))) :+: (C1 ('MetaCons "CheckDataSort" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Sort)) :+: C1 ('MetaCons "CheckMetaInst" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 MetaId)))))) :+: (((C1 ('MetaCons "CheckType" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)) :+: C1 ('MetaCons "UnBlock" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 MetaId))) :+: (C1 ('MetaCons "IsEmpty" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)) :+: (C1 ('MetaCons "CheckSizeLtSat" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)) :+: C1 ('MetaCons "FindInstance" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 MetaId) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe [Candidate])))))) :+: ((C1 ('MetaCons "ResolveInstanceHead" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "CheckFunDef" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 DefInfo) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Clause]) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 TCErr)))) :+: (C1 ('MetaCons "UnquoteTactic" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type))) :+: (C1 ('MetaCons "CheckLockedVars" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Arg Term)) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type))) :+: C1 ('MetaCons "UsableAtModality" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 WhyCheckModality) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe Sort))) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Modality) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)))))))) |
data WarningsAndNonFatalErrors Source #
Assorted warnings and errors to be displayed to the user
Sig | |
|
Instances
KillRange Signature Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
InstantiateFull Signature Source # | |||||
Defined in Agda.TypeChecking.Reduce | |||||
EmbPrj Signature Source # | |||||
NFData Signature Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Generic Signature Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
| |||||
Show Signature Source # | |||||
type Rep Signature Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep Signature = D1 ('MetaData "Signature" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "Sig" 'PrefixI 'True) ((S1 ('MetaSel ('Just "_sigSections") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Sections) :*: S1 ('MetaSel ('Just "_sigDefinitions") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Definitions)) :*: (S1 ('MetaSel ('Just "_sigRewriteRules") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 RewriteRuleMap) :*: S1 ('MetaSel ('Just "_sigInstances") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 InstanceTable)))) |
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] | |
InlineNoExactSplit QName Clause |
|
NotStrictlyPositive QName (Seq OccursWhere) | |
ConstructorDoesNotFitInData QName Sort Sort TCErr | Checking whether constructor |
CoinductiveEtaRecord QName | A record type declared as both |
UnsolvedMetaVariables [Range] | Do not use directly with |
UnsolvedInteractionMetas [Range] | Do not use directly with |
UnsolvedConstraints Constraints | Do not use directly with |
InteractionMetaBoundaries [Range] | Do not use directly with |
CantGeneralizeOverSorts [MetaId] | |
AbsurdPatternRequiresNoRHS [NamedArg DeBruijnPattern] | |
OldBuiltin BuiltinId BuiltinId | In `OldBuiltin old new`, the BUILTIN old has been replaced by new. |
BuiltinDeclaresIdentifier BuiltinId | The builtin declares a new identifier, so it should not be in scope. |
DuplicateRecordDirective RecordDirective | The given record directive is conflicting with a prior one in the same record declaration. |
EmptyRewritePragma | If the user wrote just |
EmptyWhere | An empty |
IllformedAsClause String | If the user wrote something other than an unqualified name
in the |
InvalidCharacterLiteral Char | A character literal Agda does not support, e.g. surrogate code points. |
ClashesViaRenaming NameOrModule [Name] | If a |
UselessPatternDeclarationForRecord String | The 'pattern' declaration is useless in the presence
of either |
UselessPragma Range Doc | Warning when pragma is useless and thus ignored.
|
UselessPublic | If the user opens a module public before the module header. (See issue #2377.) |
UselessHiding [ImportedName] | Names in |
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. |
NoGuardednessFlag QName | A coinductive record was declared but neither --guardedness nor --sized-types is enabled. |
SafeFlagPostulate Name | |
SafeFlagPragma [String] | Unsafe OPTIONS. |
SafeFlagWithoutKFlagPrimEraseEquality | |
WithoutKFlagPrimEraseEquality | |
ConflictingPragmaOptions String String | `ConflictingPragmaOptions a b`: Inconsistent options `--a` and `--no-b`, since `--a` implies `--b`. Ignoring `--no-b`. |
OptionWarning OptionWarning | |
ParseWarning ParseWarning | |
LibraryWarning LibWarning | |
DeprecationWarning String String String | `DeprecationWarning old new version`:
|
UserWarning Text | User-defined warning (e.g. to mention that a name is deprecated) |
DuplicateUsing (List1 ImportedName) | Duplicate mentions of the same name in |
FixityInRenamingModule (List1 Range) | Fixity of modules cannot be changed via renaming (since modules have no fixity). |
ModuleDoesntExport QName [Name] [Name] [ImportedName] | Some imported names are not actually exported by the source module. The second argument is the names that could be exported. The third argument is the module names that could be exported. |
InfectiveImport Doc | Importing a file using an infective option into one which doesn't |
CoInfectiveImport Doc | Importing a file not using a coinfective option from one which does |
ConfluenceCheckingIncompleteBecauseOfMeta QName | Confluence checking incomplete because the definition of the |
ConfluenceForCubicalNotSupported | Confluence checking with |
IllegalRewriteRule QName IllegalRewriteRuleReason | |
RewriteNonConfluent Term Term Term Doc | Confluence checker found critical pair and equality checking resulted in a type error |
RewriteMaybeNonConfluent Term Term [Doc] | Confluence checker got stuck on computing overlap between two rewrite rules |
RewriteAmbiguousRules Term Term Term | The global confluence checker found a term |
RewriteMissingRule Term Term Term | The global confluence checker found a term |
PragmaCompileErased BackendName QName | COMPILE directive for an erased symbol. |
PragmaCompileList |
|
PragmaCompileMaybe |
|
NoMain TopLevelModuleName | Compiler run on module that does not have a |
NotInScopeW [QName] | Out of scope error we can recover from. |
UnsupportedIndexedMatch Doc | Was not able to compute a full equivalence when splitting. |
AsPatternShadowsConstructorOrPatternSynonym LHSOrPatSyn | The as-name in an as-pattern may not shadow a constructor ( |
PatternShadowsConstructor Name QName | A pattern variable has the name of a constructor (data constructor or matchable record constructor). |
PlentyInHardCompileTimeMode QωOrigin | Explicit use of @ |
RecordFieldWarning RecordFieldWarning | |
MissingTypeSignatureForOpaque QName IsOpaque | An |
NotAffectedByOpaque | |
UnfoldTransparentName QName | |
UselessOpaque | |
FaceConstraintCannotBeHidden ArgInfo | Face constraint patterns |
FaceConstraintCannotBeNamed NamedName | Face constraint patterns |
DuplicateInterfaceFiles AbsolutePath AbsolutePath | `DuplicateInterfaceFiles selectedInterfaceFile ignoredInterfaceFile` |
CustomBackendWarning String Doc | Used for backend-specific warnings. The string is the backend name. |
Instances
EmbPrj Warning Source # | |||||
NFData Warning Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Generic Warning Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
| |||||
Show Warning Source # | |||||
type Rep Warning Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep Warning = D1 ('MetaData "Warning" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) ((((((C1 ('MetaCons "NicifierIssue" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 DeclarationWarning)) :+: C1 ('MetaCons "TerminationIssue" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [TerminationError]))) :+: (C1 ('MetaCons "UnreachableClauses" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Range])) :+: C1 ('MetaCons "CoverageIssue" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [(Telescope, [NamedArg DeBruijnPattern])])))) :+: ((C1 ('MetaCons "CoverageNoExactSplit" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Clause])) :+: C1 ('MetaCons "InlineNoExactSplit" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Clause))) :+: (C1 ('MetaCons "NotStrictlyPositive" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Seq OccursWhere))) :+: (C1 ('MetaCons "ConstructorDoesNotFitInData" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Sort)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Sort) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 TCErr))) :+: C1 ('MetaCons "CoinductiveEtaRecord" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)))))) :+: (((C1 ('MetaCons "UnsolvedMetaVariables" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Range])) :+: C1 ('MetaCons "UnsolvedInteractionMetas" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Range]))) :+: (C1 ('MetaCons "UnsolvedConstraints" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Constraints)) :+: C1 ('MetaCons "InteractionMetaBoundaries" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Range])))) :+: ((C1 ('MetaCons "CantGeneralizeOverSorts" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [MetaId])) :+: C1 ('MetaCons "AbsurdPatternRequiresNoRHS" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [NamedArg DeBruijnPattern]))) :+: (C1 ('MetaCons "OldBuiltin" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 BuiltinId) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 BuiltinId)) :+: (C1 ('MetaCons "BuiltinDeclaresIdentifier" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 BuiltinId)) :+: C1 ('MetaCons "DuplicateRecordDirective" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 RecordDirective))))))) :+: ((((C1 ('MetaCons "EmptyRewritePragma" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "EmptyWhere" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "IllformedAsClause" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String)) :+: C1 ('MetaCons "InvalidCharacterLiteral" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Char)))) :+: ((C1 ('MetaCons "ClashesViaRenaming" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 NameOrModule) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Name])) :+: C1 ('MetaCons "UselessPatternDeclarationForRecord" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String))) :+: (C1 ('MetaCons "UselessPragma" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Doc)) :+: (C1 ('MetaCons "UselessPublic" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "UselessHiding" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [ImportedName])))))) :+: (((C1 ('MetaCons "UselessInline" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "WrongInstanceDeclaration" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "InstanceWithExplicitArg" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: (C1 ('MetaCons "InstanceNoOutputTypeName" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Doc)) :+: C1 ('MetaCons "InstanceArgWithExplicitArg" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Doc))))) :+: ((C1 ('MetaCons "InversionDepthReached" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "NoGuardednessFlag" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName))) :+: (C1 ('MetaCons "SafeFlagPostulate" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Name)) :+: (C1 ('MetaCons "SafeFlagPragma" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [String])) :+: C1 ('MetaCons "SafeFlagWithoutKFlagPrimEraseEquality" 'PrefixI 'False) (U1 :: Type -> Type))))))) :+: (((((C1 ('MetaCons "WithoutKFlagPrimEraseEquality" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "ConflictingPragmaOptions" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String))) :+: (C1 ('MetaCons "OptionWarning" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 OptionWarning)) :+: C1 ('MetaCons "ParseWarning" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ParseWarning)))) :+: ((C1 ('MetaCons "LibraryWarning" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 LibWarning)) :+: C1 ('MetaCons "DeprecationWarning" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String)))) :+: (C1 ('MetaCons "UserWarning" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Text)) :+: (C1 ('MetaCons "DuplicateUsing" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 ImportedName))) :+: C1 ('MetaCons "FixityInRenamingModule" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 Range))))))) :+: (((C1 ('MetaCons "ModuleDoesntExport" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Name])) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Name]) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [ImportedName]))) :+: C1 ('MetaCons "InfectiveImport" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Doc))) :+: (C1 ('MetaCons "CoInfectiveImport" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Doc)) :+: C1 ('MetaCons "ConfluenceCheckingIncompleteBecauseOfMeta" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)))) :+: ((C1 ('MetaCons "ConfluenceForCubicalNotSupported" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "IllegalRewriteRule" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 IllegalRewriteRuleReason))) :+: (C1 ('MetaCons "RewriteNonConfluent" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Doc))) :+: (C1 ('MetaCons "RewriteMaybeNonConfluent" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Doc]))) :+: C1 ('MetaCons "RewriteAmbiguousRules" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)))))))) :+: ((((C1 ('MetaCons "RewriteMissingRule" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term))) :+: C1 ('MetaCons "PragmaCompileErased" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 BackendName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName))) :+: (C1 ('MetaCons "PragmaCompileList" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "PragmaCompileMaybe" 'PrefixI 'False) (U1 :: Type -> Type))) :+: ((C1 ('MetaCons "NoMain" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 TopLevelModuleName)) :+: C1 ('MetaCons "NotInScopeW" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [QName]))) :+: (C1 ('MetaCons "UnsupportedIndexedMatch" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Doc)) :+: (C1 ('MetaCons "AsPatternShadowsConstructorOrPatternSynonym" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 LHSOrPatSyn)) :+: C1 ('MetaCons "PatternShadowsConstructor" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Name) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)))))) :+: (((C1 ('MetaCons "PlentyInHardCompileTimeMode" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QωOrigin)) :+: C1 ('MetaCons "RecordFieldWarning" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 RecordFieldWarning))) :+: (C1 ('MetaCons "MissingTypeSignatureForOpaque" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 IsOpaque)) :+: (C1 ('MetaCons "NotAffectedByOpaque" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "UnfoldTransparentName" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName))))) :+: ((C1 ('MetaCons "UselessOpaque" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "FaceConstraintCannotBeHidden" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ArgInfo))) :+: (C1 ('MetaCons "FaceConstraintCannotBeNamed" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 NamedName)) :+: (C1 ('MetaCons "DuplicateInterfaceFiles" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 AbsolutePath) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 AbsolutePath)) :+: C1 ('MetaCons "CustomBackendWarning" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Doc))))))))) |
inTopContext :: (MonadTCEnv tcm, ReadTCState tcm) => tcm a -> tcm a Source #
Change to top (=empty) context. Resets the checkpoints.
TCWarning | |
|
Instances
EncodeTCM DisplayInfo Source # | |||||
Defined in Agda.Interaction.JSONTop | |||||
EncodeTCM Info_Error Source # | |||||
Defined in Agda.Interaction.JSONTop | |||||
EncodeTCM Response Source # | |||||
EncodeTCM TCWarning Source # | |||||
HasRange TCWarning Source # | |||||
PrettyTCM TCWarning Source # | |||||
Defined in Agda.TypeChecking.Pretty.Warning | |||||
EmbPrj TCWarning Source # | |||||
NFData TCWarning Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Generic TCWarning Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
| |||||
Show TCWarning Source # | |||||
Eq TCWarning Source # | |||||
type Rep TCWarning Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep TCWarning = D1 ('MetaData "TCWarning" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "TCWarning" 'PrefixI 'True) ((S1 ('MetaSel ('Just "tcWarningLocation") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 CallStack) :*: S1 ('MetaSel ('Just "tcWarningRange") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range)) :*: (S1 ('MetaSel ('Just "tcWarning") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Warning) :*: (S1 ('MetaSel ('Just "tcWarningPrintedWarning") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Doc) :*: S1 ('MetaSel ('Just "tcWarningCached") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool))))) |
For printing, we couple a meta with its name suggestion.
Instances
EncodeTCM NamedMeta Source # | |||||
Pretty NamedMeta Source # | |||||
ToConcrete NamedMeta Source # | |||||
Defined in Agda.Syntax.Translation.AbstractToConcrete
| |||||
PrettyTCM NamedMeta Source # | |||||
Defined in Agda.TypeChecking.Pretty | |||||
type ConOfAbs NamedMeta Source # | |||||
Defined in Agda.Syntax.Translation.AbstractToConcrete |
withInteractionId :: (MonadDebug m, MonadFail m, ReadTCState m, MonadError TCErr m, MonadTCEnv m, MonadTrace m) => InteractionId -> m a -> m a Source #
getInteractionRange :: (MonadInteractionPoints m, MonadFail m, MonadError TCErr m) => InteractionId -> m Range Source #
Get the Range
for an interaction point.
getMetaRange :: (HasCallStack, MonadDebug m, ReadTCState m) => MetaId -> m Range Source #
Get the Range
for a local meta-variable.
withMetaId :: (HasCallStack, MonadDebug m, MonadTCEnv m, MonadTrace m, ReadTCState m) => MetaId -> m a -> m a Source #
getIncludeDirs :: HasOptions m => m [AbsolutePath] Source #
Gets the include directories.
Precondition: optAbsoluteIncludePaths
must be nonempty (i.e.
setCommandLineOptions
must have run).
topLevelModuleName :: RawTopLevelModuleName -> TCM TopLevelModuleName Source #
Tries to convert a raw top-level module name to a top-level module name.
type ModuleToSource = Map TopLevelModuleName AbsolutePath Source #
Maps top-level module names to the corresponding source file names.
data ModuleInfo Source #
ModuleInfo | |
|
Instances
NFData ModuleInfo Source # | |||||
Defined in Agda.TypeChecking.Monad.Base rnf :: ModuleInfo -> () | |||||
Generic ModuleInfo Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
from :: ModuleInfo -> Rep ModuleInfo x to :: Rep ModuleInfo x -> ModuleInfo | |||||
type Rep ModuleInfo Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep ModuleInfo = D1 ('MetaData "ModuleInfo" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "ModuleInfo" 'PrefixI 'True) ((S1 ('MetaSel ('Just "miInterface") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Interface) :*: S1 ('MetaSel ('Just "miWarnings") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [TCWarning])) :*: (S1 ('MetaSel ('Just "miPrimitive") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool) :*: S1 ('MetaSel ('Just "miMode") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ModuleCheckMode)))) |
pattern Constructor :: Int -> Int -> ConHead -> QName -> IsAbstract -> CompKit -> Maybe [QName] -> [IsForced] -> Maybe [Bool] -> Bool -> Bool -> Defn Source #
pattern Record :: Nat -> Maybe Clause -> ConHead -> Bool -> [Dom QName] -> Telescope -> Maybe [QName] -> EtaEquality -> PatternOrCopattern -> Maybe Induction -> Maybe Bool -> IsAbstract -> CompKit -> Defn Source #
pattern Datatype :: Nat -> Nat -> Maybe Clause -> [QName] -> Sort -> Maybe [QName] -> IsAbstract -> [QName] -> Maybe QName -> Maybe QName -> Defn Source #
class (MonadTCEnv m, ReadTCState m) => MonadTrace (m :: Type -> Type) where Source #
traceCall :: Call -> m a -> m a Source #
Record a function call in the trace.
traceCallM :: m Call -> m a -> m a Source #
traceCallCPS :: Call -> ((a -> m b) -> m b) -> (a -> m b) -> m b Source #
Like traceCall
, but resets envCall
and the current ranges to the
previous values in the continuation.
traceClosureCall :: Closure Call -> m a -> m a Source #
printHighlightingInfo :: RemoveTokenBasedHighlighting -> HighlightingInfo -> m () Source #
Lispify and print the given highlighting information.
default printHighlightingInfo :: forall (t :: (Type -> Type) -> Type -> Type) (n :: Type -> Type). (MonadTrans t, MonadTrace n, t n ~ m) => RemoveTokenBasedHighlighting -> HighlightingInfo -> m () Source #
Instances
MonadTrace TCM Source # | |
Defined in Agda.TypeChecking.Monad.Trace traceCall :: Call -> TCM a -> TCM a Source # traceCallM :: TCM Call -> TCM a -> TCM a Source # traceCallCPS :: Call -> ((a -> TCM b) -> TCM b) -> (a -> TCM b) -> TCM b Source # traceClosureCall :: Closure Call -> TCM a -> TCM a Source # printHighlightingInfo :: RemoveTokenBasedHighlighting -> HighlightingInfo -> TCM () Source # | |
MonadTrace m => MonadTrace (ExceptT e m) Source # | |
Defined in Agda.TypeChecking.Monad.Trace traceCall :: Call -> ExceptT e m a -> ExceptT e m a Source # traceCallM :: ExceptT e m Call -> ExceptT e m a -> ExceptT e m a Source # traceCallCPS :: Call -> ((a -> ExceptT e m b) -> ExceptT e m b) -> (a -> ExceptT e m b) -> ExceptT e m b Source # traceClosureCall :: Closure Call -> ExceptT e m a -> ExceptT e m a Source # printHighlightingInfo :: RemoveTokenBasedHighlighting -> HighlightingInfo -> ExceptT e m () Source # | |
MonadTrace m => MonadTrace (IdentityT m) Source # | |
Defined in Agda.TypeChecking.Monad.Trace traceCall :: Call -> IdentityT m a -> IdentityT m a Source # traceCallM :: IdentityT m Call -> IdentityT m a -> IdentityT m a Source # traceCallCPS :: Call -> ((a -> IdentityT m b) -> IdentityT m b) -> (a -> IdentityT m b) -> IdentityT m b Source # traceClosureCall :: Closure Call -> IdentityT m a -> IdentityT m a Source # printHighlightingInfo :: RemoveTokenBasedHighlighting -> HighlightingInfo -> IdentityT m () Source # | |
MonadTrace m => MonadTrace (ReaderT r m) Source # | |
Defined in Agda.TypeChecking.Monad.Trace traceCall :: Call -> ReaderT r m a -> ReaderT r m a Source # traceCallM :: ReaderT r m Call -> ReaderT r m a -> ReaderT r m a Source # traceCallCPS :: Call -> ((a -> ReaderT r m b) -> ReaderT r m b) -> (a -> ReaderT r m b) -> ReaderT r m b Source # traceClosureCall :: Closure Call -> ReaderT r m a -> ReaderT r m a Source # printHighlightingInfo :: RemoveTokenBasedHighlighting -> HighlightingInfo -> ReaderT r m () Source # | |
MonadTrace m => MonadTrace (StateT s m) Source # | |
Defined in Agda.TypeChecking.Monad.Trace traceCall :: Call -> StateT s m a -> StateT s m a Source # traceCallM :: StateT s m Call -> StateT s m a -> StateT s m a Source # traceCallCPS :: Call -> ((a -> StateT s m b) -> StateT s m b) -> (a -> StateT s m b) -> StateT s m b Source # traceClosureCall :: Closure Call -> StateT s m a -> StateT s m a Source # printHighlightingInfo :: RemoveTokenBasedHighlighting -> HighlightingInfo -> StateT s m () Source # | |
(MonadTrace m, Monoid w) => MonadTrace (WriterT w m) Source # | |
Defined in Agda.TypeChecking.Monad.Trace traceCall :: Call -> WriterT w m a -> WriterT w m a Source # traceCallM :: WriterT w m Call -> WriterT w m a -> WriterT w m a Source # traceCallCPS :: Call -> ((a -> WriterT w m b) -> WriterT w m b) -> (a -> WriterT w m b) -> WriterT w m b Source # traceClosureCall :: Closure Call -> WriterT w m a -> WriterT w m a Source # printHighlightingInfo :: RemoveTokenBasedHighlighting -> HighlightingInfo -> WriterT w m () Source # |
highlightAsTypeChecked Source #
:: MonadTrace m | |
=> Range | rPre |
-> Range | r |
-> m a | |
-> m a |
highlightAsTypeChecked rPre r m
runs m
and returns its
result. Additionally, some code may be highlighted:
- If
r
is non-empty and not a sub-range ofrPre
(aftercontinuousPerLine
has been applied to both):r
is highlighted as being type-checked whilem
is running (this highlighting is removed ifm
completes successfully). - Otherwise: Highlighting is removed for
rPre - r
beforem
runs, and ifm
completes successfully, thenrPre - r
is highlighted as being type-checked.
type Definitions = HashMap QName Definition Source #
setCommandLineOptions :: CommandLineOptions -> TCM () Source #
Sets the command line options (both persistent and pragma options are updated).
Relative include directories are made absolute with respect to the
current working directory. If the include directories have changed
then the state is reset (partly, see setIncludeDirs
).
An empty list of relative include directories (
) is
interpreted as Left
[]["."]
.
withScope :: ReadTCState m => ScopeInfo -> m a -> m (a, ScopeInfo) Source #
Run a computation in a local scope.
type InteractionOutputCallback = Response_boot TCErr TCWarning WarningsAndNonFatalErrors -> TCM () Source #
Callback fuction to call when there is a response to give to the interactive frontend.
Note that the response is given in pieces and incrementally, so the user can have timely response even during long computations.
Typical InteractionOutputCallback
functions:
- Convert the response into a
String
representation and print it on standard output (suitable for inter-process communication). - Put the response into a mutable variable stored in the
closure of the
InteractionOutputCallback
function. (suitable for intra-process communication).
defaultInteractionOutputCallback :: InteractionOutputCallback Source #
The default InteractionOutputCallback
function prints certain
things to stdout (other things generate internal errors).
Datatype representing a single boundary condition: x_0 = u_0, ... ,x_n = u_n ⊢ t = ?n es
class (Functor m, Applicative m, Monad m) => HasOptions (m :: Type -> Type) where Source #
Nothing
pragmaOptions :: m PragmaOptions Source #
Returns the pragma options which are currently in effect.
default pragmaOptions :: forall (n :: Type -> Type) (t :: (Type -> Type) -> Type -> Type). (HasOptions n, MonadTrans t, m ~ t n) => m PragmaOptions Source #
commandLineOptions :: m CommandLineOptions Source #
Returns the command line options which are currently in effect.
default commandLineOptions :: forall (n :: Type -> Type) (t :: (Type -> Type) -> Type -> Type). (HasOptions n, MonadTrans t, m ~ t n) => m CommandLineOptions Source #
Instances
HasOptions IM Source # | |
Defined in Agda.Interaction.Monad | |
HasOptions AbsToCon Source # | |
HasOptions TerM Source # | |
HasOptions ReduceM Source # | |
HasOptions NLM Source # | |
HasOptions m => HasOptions (PureConversionT m) Source # | |
HasOptions m => HasOptions (BlockT m) Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
MonadIO m => HasOptions (TCMT m) Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
HasOptions m => HasOptions (NamesT m) Source # | |
Defined in Agda.TypeChecking.Names | |
HasOptions m => HasOptions (ListT m) Source # | |
HasOptions m => HasOptions (ChangeT m) Source # | |
HasOptions m => HasOptions (MaybeT m) Source # | |
Defined in Agda.Interaction.Options.HasOptions pragmaOptions :: MaybeT m PragmaOptions Source # commandLineOptions :: MaybeT m CommandLineOptions Source # | |
HasOptions m => HasOptions (ExceptT e m) Source # | |
Defined in Agda.Interaction.Options.HasOptions pragmaOptions :: ExceptT e m PragmaOptions Source # commandLineOptions :: ExceptT e m CommandLineOptions Source # | |
HasOptions m => HasOptions (IdentityT m) Source # | |
Defined in Agda.Interaction.Options.HasOptions pragmaOptions :: IdentityT m PragmaOptions Source # commandLineOptions :: IdentityT m CommandLineOptions Source # | |
HasOptions m => HasOptions (ReaderT r m) Source # | |
Defined in Agda.Interaction.Options.HasOptions pragmaOptions :: ReaderT r m PragmaOptions Source # commandLineOptions :: ReaderT r m CommandLineOptions Source # | |
HasOptions m => HasOptions (StateT s m) Source # | |
Defined in Agda.Interaction.Options.HasOptions pragmaOptions :: StateT s m PragmaOptions Source # commandLineOptions :: StateT s m CommandLineOptions Source # | |
(HasOptions m, Monoid w) => HasOptions (WriterT w m) Source # | |
Defined in Agda.Interaction.Options.HasOptions pragmaOptions :: WriterT w m PragmaOptions Source # commandLineOptions :: WriterT w m CommandLineOptions Source # |
class Monad m => MonadTCEnv (m :: Type -> Type) where Source #
MonadTCEnv
made into its own dedicated service class.
This allows us to use MonadReader
for ReaderT
extensions of TCM
.
Nothing
default askTC :: forall (t :: (Type -> Type) -> Type -> Type) (n :: Type -> Type). (MonadTrans t, MonadTCEnv n, t n ~ m) => m TCEnv Source #
Instances
MonadTCEnv IM Source # | |
MonadTCEnv AbsToCon Source # | |
MonadTCEnv TerM Source # | |
MonadTCEnv ReduceM Source # | |
MonadTCEnv NLM 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 (BlockT m) Source # | |
MonadIO m => MonadTCEnv (TCMT m) Source # | |
MonadTCEnv m => MonadTCEnv (NamesT m) Source # | |
MonadTCEnv m => MonadTCEnv (ListT m) Source # | |
MonadTCEnv m => MonadTCEnv (ChangeT m) Source # | |
MonadTCEnv m => MonadTCEnv (MaybeT m) Source # | |
MonadTCEnv m => MonadTCEnv (ExceptT err m) Source # | |
MonadTCEnv m => MonadTCEnv (IdentityT m) Source # | |
MonadTCEnv m => MonadTCEnv (ReaderT r m) Source # | |
MonadTCEnv m => MonadTCEnv (StateT s m) Source # | |
(Monoid w, MonadTCEnv m) => MonadTCEnv (WriterT w m) Source # | |
class (Applicative tcm, MonadIO tcm, MonadTCEnv tcm, MonadTCState tcm, HasOptions tcm) => MonadTCM (tcm :: Type -> Type) where Source #
Embedding a TCM computation.
Nothing
Instances
MonadTCM IM Source # | |
MonadTCM TerM Source # | |
MonadTCM m => MonadTCM (BlockT m) Source # | |
MonadIO m => MonadTCM (TCMT m) Source # | |
MonadTCM m => MonadTCM (NamesT m) Source # | |
MonadTCM tcm => MonadTCM (ListT tcm) Source # | |
MonadTCM tcm => MonadTCM (ChangeT tcm) Source # | |
MonadTCM tcm => MonadTCM (MaybeT tcm) Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
MonadTCM tcm => MonadTCM (ExceptT err tcm) Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
MonadTCM tcm => MonadTCM (IdentityT tcm) Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
MonadTCM tcm => MonadTCM (ReaderT r tcm) Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
MonadTCM tcm => MonadTCM (StateT s tcm) Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
(Monoid w, MonadTCM tcm) => MonadTCM (WriterT w tcm) Source # | |
Defined in Agda.TypeChecking.Monad.Base |
class Monad m => MonadTCState (m :: Type -> Type) where Source #
MonadTCState
made into its own dedicated service class.
This allows us to use MonadState
for StateT
extensions of TCM
.
Nothing
default getTC :: forall (t :: (Type -> Type) -> Type -> Type) (n :: Type -> Type). (MonadTrans t, MonadTCState n, t n ~ m) => m TCState Source #
putTC :: TCState -> m () Source #
default putTC :: forall (t :: (Type -> Type) -> Type -> Type) (n :: Type -> Type). (MonadTrans t, MonadTCState n, t n ~ m) => TCState -> m () Source #
Instances
MonadTCState IM Source # | |
MonadTCState TerM Source # | |
MonadTCState m => MonadTCState (BlockT m) Source # | |
MonadIO m => MonadTCState (TCMT m) Source # | |
MonadTCState m => MonadTCState (NamesT m) Source # | |
MonadTCState m => MonadTCState (ListT m) Source # | |
MonadTCState m => MonadTCState (ChangeT m) Source # | |
MonadTCState m => MonadTCState (MaybeT m) Source # | |
MonadTCState m => MonadTCState (ExceptT err m) Source # | |
MonadTCState m => MonadTCState (IdentityT m) Source # | |
MonadTCState m => MonadTCState (ReaderT r m) Source # | |
MonadTCState m => MonadTCState (StateT s m) Source # | |
(Monoid w, MonadTCState m) => MonadTCState (WriterT w m) Source # | |
class Monad m => ReadTCState (m :: Type -> Type) where Source #
Nothing
getTCState :: m TCState Source #
default getTCState :: forall (t :: (Type -> Type) -> Type -> Type) (n :: Type -> Type). (MonadTrans t, ReadTCState n, t n ~ m) => m TCState Source #
locallyTCState :: Lens' TCState a -> (a -> a) -> m b -> m b Source #
default locallyTCState :: forall (t :: (Type -> Type) -> Type -> Type) (n :: Type -> Type) a b. (MonadTransControl t, ReadTCState n, t n ~ m) => Lens' TCState a -> (a -> a) -> m b -> m b Source #
withTCState :: (TCState -> TCState) -> m a -> m a Source #
Instances
ReadTCState IM Source # | |
Defined in Agda.Interaction.Monad | |
ReadTCState AbsToCon Source # | |
Defined in Agda.Syntax.Translation.AbstractToConcrete | |
ReadTCState TerM Source # | |
Defined in Agda.Termination.Monad | |
ReadTCState ReduceM Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
ReadTCState NLM Source # | |
Defined in Agda.TypeChecking.Rewriting.NonLinMatch | |
ReadTCState m => ReadTCState (PureConversionT m) Source # | |
Defined in Agda.TypeChecking.Conversion.Pure getTCState :: PureConversionT m TCState Source # locallyTCState :: Lens' TCState a -> (a -> a) -> PureConversionT m b -> PureConversionT m b Source # withTCState :: (TCState -> TCState) -> PureConversionT m a -> PureConversionT m a Source # | |
ReadTCState m => ReadTCState (BlockT m) Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
MonadIO m => ReadTCState (TCMT m) Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
ReadTCState m => ReadTCState (NamesT m) Source # | |
Defined in Agda.TypeChecking.Names | |
ReadTCState m => ReadTCState (ListT m) Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
ReadTCState m => ReadTCState (ChangeT m) Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
ReadTCState m => ReadTCState (MaybeT m) Source # | |
Defined in Agda.TypeChecking.Monad.Base getTCState :: MaybeT m TCState Source # locallyTCState :: Lens' TCState a -> (a -> a) -> MaybeT m b -> MaybeT m b Source # withTCState :: (TCState -> TCState) -> MaybeT m a -> MaybeT m a Source # | |
ReadTCState m => ReadTCState (ExceptT err m) Source # | |
Defined in Agda.TypeChecking.Monad.Base getTCState :: ExceptT err m TCState Source # locallyTCState :: Lens' TCState a -> (a -> a) -> ExceptT err m b -> ExceptT err m b Source # withTCState :: (TCState -> TCState) -> ExceptT err m a -> ExceptT err m a Source # | |
ReadTCState m => ReadTCState (IdentityT m) Source # | |
Defined in Agda.TypeChecking.Monad.Base getTCState :: IdentityT m TCState Source # locallyTCState :: Lens' TCState a -> (a -> a) -> IdentityT m b -> IdentityT m b Source # withTCState :: (TCState -> TCState) -> IdentityT m a -> IdentityT m a Source # | |
ReadTCState m => ReadTCState (ReaderT r m) Source # | |
Defined in Agda.TypeChecking.Monad.Base getTCState :: ReaderT r m TCState Source # locallyTCState :: Lens' TCState a -> (a -> a) -> ReaderT r m b -> ReaderT r m b Source # withTCState :: (TCState -> TCState) -> ReaderT r m a -> ReaderT r m a Source # | |
ReadTCState m => ReadTCState (StateT s m) Source # | |
Defined in Agda.TypeChecking.Monad.Base getTCState :: StateT s m TCState Source # locallyTCState :: Lens' TCState a -> (a -> a) -> StateT s m b -> StateT s m b Source # withTCState :: (TCState -> TCState) -> StateT s m a -> StateT s m a Source # | |
(Monoid w, ReadTCState m) => ReadTCState (WriterT w m) Source # | |
Defined in Agda.TypeChecking.Monad.Base getTCState :: WriterT w m TCState Source # locallyTCState :: Lens' TCState a -> (a -> a) -> WriterT w m b -> WriterT w m b Source # withTCState :: (TCState -> TCState) -> WriterT w m a -> WriterT w m a Source # |
data SomeBuiltin Source #
Either a BuiltinId
or PrimitiveId
, used for some lookups.
Instances
IsBuiltin SomeBuiltin Source # | |||||
Defined in Agda.Syntax.Builtin someBuiltin :: SomeBuiltin -> SomeBuiltin Source # getBuiltinId :: SomeBuiltin -> String Source # | |||||
EmbPrj SomeBuiltin Source # | |||||
Defined in Agda.TypeChecking.Serialise.Instances.Common icode :: SomeBuiltin -> S Int32 Source # icod_ :: SomeBuiltin -> S Int32 Source # value :: Int32 -> R SomeBuiltin Source # | |||||
NFData SomeBuiltin Source # | |||||
Defined in Agda.Syntax.Builtin rnf :: SomeBuiltin -> () | |||||
Generic SomeBuiltin Source # | |||||
Defined in Agda.Syntax.Builtin
from :: SomeBuiltin -> Rep SomeBuiltin x to :: Rep SomeBuiltin x -> SomeBuiltin | |||||
Show SomeBuiltin Source # | |||||
Defined in Agda.Syntax.Builtin showsPrec :: Int -> SomeBuiltin -> ShowS show :: SomeBuiltin -> String showList :: [SomeBuiltin] -> ShowS | |||||
Eq SomeBuiltin Source # | |||||
Defined in Agda.Syntax.Builtin (==) :: SomeBuiltin -> SomeBuiltin -> Bool (/=) :: SomeBuiltin -> SomeBuiltin -> Bool | |||||
Ord SomeBuiltin Source # | |||||
Defined in Agda.Syntax.Builtin compare :: SomeBuiltin -> SomeBuiltin -> Ordering (<) :: SomeBuiltin -> SomeBuiltin -> Bool (<=) :: SomeBuiltin -> SomeBuiltin -> Bool (>) :: SomeBuiltin -> SomeBuiltin -> Bool (>=) :: SomeBuiltin -> SomeBuiltin -> Bool max :: SomeBuiltin -> SomeBuiltin -> SomeBuiltin min :: SomeBuiltin -> SomeBuiltin -> SomeBuiltin | |||||
Hashable SomeBuiltin Source # | |||||
Defined in Agda.Syntax.Builtin hashWithSalt :: Int -> SomeBuiltin -> Int hash :: SomeBuiltin -> Int | |||||
type Rep SomeBuiltin Source # | |||||
Defined in Agda.Syntax.Builtin type Rep SomeBuiltin = D1 ('MetaData "SomeBuiltin" "Agda.Syntax.Builtin" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "BuiltinName" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 BuiltinId)) :+: C1 ('MetaCons "PrimitiveName" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 PrimitiveId))) |
builtinById :: String -> Maybe BuiltinId Source #
Lookup a builtin by the string used in the BUILTIN
pragma.
isBuiltinNoDef :: BuiltinId -> Bool Source #
Builtins that come without a definition in Agda syntax. These are giving names to Agda internal concepts which cannot be assigned an Agda type.
An example would be a user-defined name for Set
.
{-# BUILTIN TYPE Type #-}
The type of Type
would be Type : Level → Setω
which is not valid Agda.
typeError :: (HasCallStack, MonadTCError m) => TypeError -> m a Source #
InternalError String | |
NotImplemented String | |
NotSupported String | |
CompilationError String | |
PropMustBeSingleton | |
DataMustEndInSort Term | |
ShouldEndInApplicationOfTheDatatype Type | The target of a constructor isn't an application of its
datatype. The |
ShouldBeAppliedToTheDatatypeParameters Term Term | The target of a constructor isn't its datatype applied to something that isn't the parameters. First term is the correct target and the second term is the actual target. |
ShouldBeApplicationOf Type QName | Expected a type to be an application of a particular datatype. |
ConstructorPatternInWrongDatatype QName QName | constructor, datatype |
CantResolveOverloadedConstructorsTargetingSameDatatype QName (List1 QName) | Datatype, constructors. |
DoesNotConstructAnElementOf QName Type | constructor, type |
WrongHidingInLHS | The left hand side of a function definition has a hidden argument where a non-hidden was expected. |
WrongHidingInLambda Type | Expected a non-hidden function and found a hidden lambda. |
WrongHidingInApplication Type | A function is applied to a hidden argument where a non-hidden was expected. |
WrongHidingInProjection QName | |
IllegalHidingInPostfixProjection (NamedArg Expr) | |
WrongNamedArgument (NamedArg Expr) [NamedName] | A function is applied to a hidden named argument it does not have. The list contains names of possible hidden arguments at this point. |
WrongIrrelevanceInLambda | Wrong user-given relevance annotation in lambda. |
WrongQuantityInLambda | Wrong user-given quantity annotation in lambda. |
WrongCohesionInLambda | Wrong user-given cohesion annotation in lambda. |
QuantityMismatch Quantity Quantity | The given quantity does not correspond to the expected quantity. |
HidingMismatch Hiding Hiding | The given hiding does not correspond to the expected hiding. |
RelevanceMismatch Relevance Relevance | The given relevance does not correspond to the expected relevane. |
UninstantiatedDotPattern Expr | |
ForcedConstructorNotInstantiated Pattern | |
IllformedProjectionPatternAbstract Pattern | |
IllformedProjectionPatternConcrete Pattern | |
CannotEliminateWithPattern (Maybe Blocker) (NamedArg Pattern) Type | |
CannotEliminateWithProjection (Arg Type) Bool QName | |
WrongNumberOfConstructorArguments QName Nat Nat | |
ShouldBeEmpty Type [DeBruijnPattern] | |
ShouldBeASort Type | The given type should have been a sort. |
ShouldBePi Type | The given type should have been a pi. |
ShouldBePath Type | |
ShouldBeRecordType Type | |
ShouldBeRecordPattern DeBruijnPattern | |
NotAProjectionPattern (NamedArg Pattern) | |
NotAProperTerm | |
InvalidTypeSort Sort | This sort is not a type expression. |
InvalidType Term | This term is not a type expression. |
SplitOnCoinductive | |
SplitOnIrrelevant (Dom Type) | |
SplitOnUnusableCohesion (Dom Type) | |
SplitOnNonVariable Term Type | |
SplitOnNonEtaRecord QName | |
SplitOnAbstract QName | |
SplitOnUnchecked QName | |
SplitOnPartial (Dom Type) | |
SplitInProp DataOrRecordE | |
DefinitionIsIrrelevant QName | |
DefinitionIsErased QName | |
ProjectionIsIrrelevant QName | |
VariableIsIrrelevant Name | |
VariableIsErased Name | |
VariableIsOfUnusableCohesion Name Cohesion | |
UnequalLevel Comparison Level Level | |
UnequalTerms Comparison Term Term CompareAs | |
UnequalTypes Comparison Type Type | |
UnequalRelevance Comparison Term Term | The two function types have different relevance. |
UnequalQuantity Comparison Term Term | The two function types have different relevance. |
UnequalCohesion Comparison Term Term | The two function types have different cohesion. |
UnequalFiniteness Comparison Term Term | One of the function types has a finite domain (i.e. is a |
UnequalHiding Term Term | The two function types have different hiding. |
UnequalSorts Sort Sort | |
UnequalBecauseOfUniverseConflict Comparison Term Term | |
NotLeqSort Sort Sort | |
MetaCannotDependOn MetaId Nat | The arguments are the meta variable and the parameter that it wants to depend on. |
MetaOccursInItself MetaId | |
MetaIrrelevantSolution MetaId Term | |
MetaErasedSolution MetaId Term | |
GenericError String | |
GenericDocError Doc | |
SortOfSplitVarError (Maybe Blocker) Doc | the meta is what we might be blocked on. |
BuiltinMustBeConstructor BuiltinId Expr | |
NoSuchBuiltinName String | |
DuplicateBuiltinBinding BuiltinId Term Term | |
NoBindingForBuiltin BuiltinId | |
NoBindingForPrimitive PrimitiveId | |
NoSuchPrimitiveFunction String | |
DuplicatePrimitiveBinding PrimitiveId QName QName | |
WrongArgInfoForPrimitive PrimitiveId ArgInfo ArgInfo | |
ShadowedModule Name [ModuleName] | |
BuiltinInParameterisedModule BuiltinId | |
IllegalDeclarationInDataDefinition [Declaration] | The declaration list comes from a single |
IllegalLetInTelescope TypedBinding | |
IllegalPatternInTelescope Binder | |
NoRHSRequiresAbsurdPattern [NamedArg Pattern] | |
TooManyFields QName [Name] [Name] | Record type, fields not supplied by user, non-fields but supplied. |
DuplicateFields [Name] | |
DuplicateConstructors [Name] | |
DuplicateOverlapPragma QName OverlapMode OverlapMode | |
WithOnFreeVariable Expr Term | |
UnexpectedWithPatterns [Pattern] | |
WithClausePatternMismatch Pattern (NamedArg DeBruijnPattern) | |
IllTypedPatternAfterWithAbstraction Pattern | |
FieldOutsideRecord | |
ModuleArityMismatch ModuleName Telescope [NamedArg Expr] | |
GeneralizeCyclicDependency | |
GeneralizeUnsolvedMeta | |
ReferencesFutureVariables Term (NonEmpty Int) (Arg Term) Int | The first term references the given list of variables, which are in "the future" with respect to the given lock (and its leftmost variable) |
DoesNotMentionTicks Term Type (Arg Term) | Arguments: later term, its type, lock term. The lock term does not mention any @lock variables. |
MismatchedProjectionsError QName QName | |
AttributeKindNotEnabled String String String | |
InvalidProjectionParameter (NamedArg Expr) | |
TacticAttributeNotAllowed | |
CannotRewriteByNonEquation Type | |
MacroResultTypeMismatch Type | |
NamedWhereModuleInRefinedContext [Term] [String] | |
CubicalPrimitiveNotFullyApplied QName | |
TooManyArgumentsToLeveledSort QName | |
TooManyArgumentsToUnivOmega QName | |
ComatchingDisabledForRecord QName | |
BuiltinMustBeIsOne Term | |
IncorrectTypeForRewriteRelation Term IncorrectTypeForRewriteRelationReason | |
UnexpectedParameter LamBinding | |
NoParameterOfName ArgName | |
UnexpectedModalityAnnotationInParameter LamBinding | |
ExpectedBindingForParameter (Dom Type) (Abs Type) | |
UnexpectedTypeSignatureForParameter (List1 (NamedArg Binder)) | |
SortDoesNotAdmitDataDefinitions QName Sort | |
SortCannotDependOnItsIndex QName Type | |
UnusableAtModality WhyCheckModality Modality Term | |
SplitError SplitError | |
ImpossibleConstructor QName NegativeUnification | |
TooManyPolarities QName Int | |
RecursiveRecordNeedsInductivity QName | A record type inferred as recursive needs a manual declaration whether it should be inductively or coinductively. Sized type errors |
CannotSolveSizeConstraints (List1 (ProblemConstraint, HypSizeConstraint)) Doc | The list of constraints is given redundantly as pairs of
|
ContradictorySizeConstraint (ProblemConstraint, HypSizeConstraint) | |
EmptyTypeOfSizes Term | This type, representing a type of sizes, might be empty. |
FunctionTypeInSizeUniv Term | This term, a function type constructor, lives in
|
LibraryError LibErrors | Collected errors when processing the |
LocalVsImportedModuleClash ModuleName | |
SolvedButOpenHoles | Some interaction points (holes) have not been filled by user.
There are not |
CyclicModuleDependency [TopLevelModuleName] | |
FileNotFound TopLevelModuleName [AbsolutePath] | |
OverlappingProjects AbsolutePath TopLevelModuleName TopLevelModuleName | |
AmbiguousTopLevelModuleName TopLevelModuleName [AbsolutePath] | |
ModuleNameUnexpected TopLevelModuleName TopLevelModuleName | Found module name, expected module name. |
ModuleNameDoesntMatchFileName TopLevelModuleName [AbsolutePath] | |
ClashingFileNamesFor ModuleName [AbsolutePath] | |
ModuleDefinedInOtherFile TopLevelModuleName AbsolutePath AbsolutePath | Module name, file from which it was loaded, file which the include path says contains the module. |
InvalidFileName AbsolutePath InvalidFileNameReason | The file name does not correspond to a module name. Scope errors |
BothWithAndRHS | |
AbstractConstructorNotInScope QName | |
NotInScope [QName] | |
NoSuchModule QName | |
AmbiguousName QName AmbiguousNameReason | |
AmbiguousModule QName (List1 ModuleName) | |
AmbiguousField Name [ModuleName] | |
AmbiguousConstructor QName [QName] | |
ClashingDefinition QName QName (Maybe NiceDeclaration) | |
ClashingModule ModuleName ModuleName | |
ClashingImport Name QName | |
ClashingModuleImport Name ModuleName | |
DefinitionInDifferentModule QName | The given data/record definition rests in a different module than its signature. |
DuplicateImports QName [ImportedName] | |
InvalidPattern Pattern | |
RepeatedVariablesInPattern [Name] | |
GeneralizeNotSupportedHere QName | |
GeneralizedVarInLetOpenedModule QName | |
MultipleFixityDecls [(Name, [Fixity'])] | |
MultiplePolarityPragmas [Name] | |
NotAModuleExpr Expr | The expr was used in the right hand side of an implicit module
definition, but it wasn't of the form |
NotAnExpression Expr | |
NotAValidLetBinding NiceDeclaration | |
NotValidBeforeField NiceDeclaration | |
NothingAppliedToHiddenArg Expr | |
NothingAppliedToInstanceArg Expr | |
AsPatternInPatternSynonym | |
DotPatternInPatternSynonym | |
BadArgumentsToPatternSynonym AmbiguousQName | |
TooFewArgumentsToPatternSynonym AmbiguousQName | |
CannotResolveAmbiguousPatternSynonym (List1 (QName, PatternSynDefn)) | |
IllegalInstanceVariableInPatternSynonym Name | This variable is bound in the lhs of the pattern synonym in instance position, but not on the rhs. This is forbidden because expansion of pattern synonyms would not be faithful to availability of instances in instance search. |
PatternSynonymArgumentShadowsConstructorOrPatternSynonym LHSOrPatSyn Name (List1 AbstractName) | A variable to be bound in the pattern synonym resolved on the rhs as name of a constructor or a pattern synonym. The resolvents are given in the list. |
UnusedVariableInPatternSynonym Name | This variable is only bound on the lhs of the pattern synonym, not on the rhs. |
UnboundVariablesInPatternSynonym [Name] | These variables are only bound on the rhs of the pattern synonym, not on the lhs. Operator errors |
NoParseForApplication (List2 Expr) | |
AmbiguousParseForApplication (List2 Expr) (List1 Expr) | |
NoParseForLHS LHSOrPatSyn [Pattern] Pattern | The list contains patterns that failed to be interpreted. If it is non-empty, the first entry could be printed as error hint. |
AmbiguousParseForLHS LHSOrPatSyn Pattern [Pattern] | Pattern and its possible interpretations. |
AmbiguousProjection QName [QName] | |
AmbiguousOverloadedProjection (List1 QName) Doc | |
OperatorInformation [NotationSection] TypeError | |
InstanceNoCandidate Type [(Term, TCErr)] | |
UnquoteFailed UnquoteError | |
DeBruijnIndexOutOfScope Nat Telescope [Name] | |
NeedOptionCopatterns | |
NeedOptionRewriting | |
NeedOptionProp | |
NeedOptionTwoLevel | |
NonFatalErrors [TCWarning] | |
InstanceSearchDepthExhausted Term Type Int | |
TriedToCopyConstrainedPrim QName | |
CustomBackendError String Doc | Used for backend-specific errors. The string is the backend name. |
Instances
PrettyTCM TypeError Source # | |||||
Defined in Agda.TypeChecking.Errors | |||||
NFData TypeError Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Generic TypeError Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
| |||||
Show TypeError Source # | |||||
type Rep TypeError Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep TypeError = D1 ('MetaData "TypeError" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (((((((C1 ('MetaCons "InternalError" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String)) :+: (C1 ('MetaCons "NotImplemented" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String)) :+: C1 ('MetaCons "NotSupported" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String)))) :+: (C1 ('MetaCons "CompilationError" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String)) :+: (C1 ('MetaCons "PropMustBeSingleton" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "DataMustEndInSort" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term))))) :+: ((C1 ('MetaCons "ShouldEndInApplicationOfTheDatatype" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)) :+: (C1 ('MetaCons "ShouldBeAppliedToTheDatatypeParameters" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)) :+: C1 ('MetaCons "ShouldBeApplicationOf" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)))) :+: (C1 ('MetaCons "ConstructorPatternInWrongDatatype" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: (C1 ('MetaCons "CantResolveOverloadedConstructorsTargetingSameDatatype" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 QName))) :+: C1 ('MetaCons "DoesNotConstructAnElementOf" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)))))) :+: (((C1 ('MetaCons "WrongHidingInLHS" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "WrongHidingInLambda" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)) :+: C1 ('MetaCons "WrongHidingInApplication" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)))) :+: (C1 ('MetaCons "WrongHidingInProjection" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: (C1 ('MetaCons "IllegalHidingInPostfixProjection" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (NamedArg Expr))) :+: C1 ('MetaCons "WrongNamedArgument" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (NamedArg Expr)) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [NamedName]))))) :+: ((C1 ('MetaCons "WrongIrrelevanceInLambda" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "WrongQuantityInLambda" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "WrongCohesionInLambda" 'PrefixI 'False) (U1 :: Type -> Type))) :+: (C1 ('MetaCons "QuantityMismatch" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Quantity) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Quantity)) :+: (C1 ('MetaCons "HidingMismatch" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Hiding) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Hiding)) :+: C1 ('MetaCons "RelevanceMismatch" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Relevance) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Relevance))))))) :+: ((((C1 ('MetaCons "UninstantiatedDotPattern" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr)) :+: (C1 ('MetaCons "ForcedConstructorNotInstantiated" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Pattern)) :+: C1 ('MetaCons "IllformedProjectionPatternAbstract" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Pattern)))) :+: (C1 ('MetaCons "IllformedProjectionPatternConcrete" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Pattern)) :+: (C1 ('MetaCons "CannotEliminateWithPattern" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe Blocker)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (NamedArg Pattern)) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type))) :+: C1 ('MetaCons "CannotEliminateWithProjection" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Arg Type)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)))))) :+: ((C1 ('MetaCons "WrongNumberOfConstructorArguments" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Nat) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Nat))) :+: (C1 ('MetaCons "ShouldBeEmpty" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [DeBruijnPattern])) :+: C1 ('MetaCons "ShouldBeASort" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)))) :+: (C1 ('MetaCons "ShouldBePi" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)) :+: (C1 ('MetaCons "ShouldBePath" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)) :+: C1 ('MetaCons "ShouldBeRecordType" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)))))) :+: (((C1 ('MetaCons "ShouldBeRecordPattern" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 DeBruijnPattern)) :+: (C1 ('MetaCons "NotAProjectionPattern" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (NamedArg Pattern))) :+: C1 ('MetaCons "NotAProperTerm" 'PrefixI 'False) (U1 :: Type -> Type))) :+: (C1 ('MetaCons "InvalidTypeSort" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Sort)) :+: (C1 ('MetaCons "InvalidType" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)) :+: C1 ('MetaCons "SplitOnCoinductive" 'PrefixI 'False) (U1 :: Type -> Type)))) :+: ((C1 ('MetaCons "SplitOnIrrelevant" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Dom Type))) :+: (C1 ('MetaCons "SplitOnUnusableCohesion" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Dom Type))) :+: C1 ('MetaCons "SplitOnNonVariable" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)))) :+: ((C1 ('MetaCons "SplitOnNonEtaRecord" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "SplitOnAbstract" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName))) :+: (C1 ('MetaCons "SplitOnUnchecked" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "SplitOnPartial" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Dom Type))))))))) :+: (((((C1 ('MetaCons "SplitInProp" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 DataOrRecordE)) :+: (C1 ('MetaCons "DefinitionIsIrrelevant" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "DefinitionIsErased" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)))) :+: (C1 ('MetaCons "ProjectionIsIrrelevant" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: (C1 ('MetaCons "VariableIsIrrelevant" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Name)) :+: C1 ('MetaCons "VariableIsErased" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Name))))) :+: ((C1 ('MetaCons "VariableIsOfUnusableCohesion" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Name) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Cohesion)) :+: (C1 ('MetaCons "UnequalLevel" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Level) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Level))) :+: C1 ('MetaCons "UnequalTerms" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 CompareAs))))) :+: (C1 ('MetaCons "UnequalTypes" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type))) :+: (C1 ('MetaCons "UnequalRelevance" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term))) :+: C1 ('MetaCons "UnequalQuantity" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term))))))) :+: (((C1 ('MetaCons "UnequalCohesion" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term))) :+: (C1 ('MetaCons "UnequalFiniteness" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term))) :+: C1 ('MetaCons "UnequalHiding" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)))) :+: (C1 ('MetaCons "UnequalSorts" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Sort) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Sort)) :+: (C1 ('MetaCons "UnequalBecauseOfUniverseConflict" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term))) :+: C1 ('MetaCons "NotLeqSort" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Sort) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Sort))))) :+: ((C1 ('MetaCons "MetaCannotDependOn" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 MetaId) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Nat)) :+: (C1 ('MetaCons "MetaOccursInItself" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 MetaId)) :+: C1 ('MetaCons "MetaIrrelevantSolution" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 MetaId) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)))) :+: (C1 ('MetaCons "MetaErasedSolution" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 MetaId) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)) :+: (C1 ('MetaCons "GenericError" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String)) :+: C1 ('MetaCons "GenericDocError" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Doc))))))) :+: ((((C1 ('MetaCons "SortOfSplitVarError" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe Blocker)) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Doc)) :+: (C1 ('MetaCons "BuiltinMustBeConstructor" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 BuiltinId) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr)) :+: C1 ('MetaCons "NoSuchBuiltinName" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String)))) :+: (C1 ('MetaCons "DuplicateBuiltinBinding" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 BuiltinId) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term))) :+: (C1 ('MetaCons "NoBindingForBuiltin" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 BuiltinId)) :+: C1 ('MetaCons "NoBindingForPrimitive" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 PrimitiveId))))) :+: ((C1 ('MetaCons "NoSuchPrimitiveFunction" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String)) :+: (C1 ('MetaCons "DuplicatePrimitiveBinding" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 PrimitiveId) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName))) :+: C1 ('MetaCons "WrongArgInfoForPrimitive" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 PrimitiveId) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ArgInfo) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ArgInfo))))) :+: (C1 ('MetaCons "ShadowedModule" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Name) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [ModuleName])) :+: (C1 ('MetaCons "BuiltinInParameterisedModule" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 BuiltinId)) :+: C1 ('MetaCons "IllegalDeclarationInDataDefinition" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Declaration])))))) :+: (((C1 ('MetaCons "IllegalLetInTelescope" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 TypedBinding)) :+: (C1 ('MetaCons "IllegalPatternInTelescope" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Binder)) :+: C1 ('MetaCons "NoRHSRequiresAbsurdPattern" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [NamedArg Pattern])))) :+: (C1 ('MetaCons "TooManyFields" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Name]) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Name]))) :+: (C1 ('MetaCons "DuplicateFields" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Name])) :+: C1 ('MetaCons "DuplicateConstructors" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Name]))))) :+: ((C1 ('MetaCons "DuplicateOverlapPragma" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 OverlapMode) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 OverlapMode))) :+: (C1 ('MetaCons "WithOnFreeVariable" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)) :+: C1 ('MetaCons "UnexpectedWithPatterns" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Pattern])))) :+: ((C1 ('MetaCons "WithClausePatternMismatch" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Pattern) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (NamedArg DeBruijnPattern))) :+: C1 ('MetaCons "IllTypedPatternAfterWithAbstraction" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Pattern))) :+: (C1 ('MetaCons "FieldOutsideRecord" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "ModuleArityMismatch" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ModuleName) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Telescope) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [NamedArg Expr])))))))))) :+: ((((((C1 ('MetaCons "GeneralizeCyclicDependency" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "GeneralizeUnsolvedMeta" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "ReferencesFutureVariables" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (NonEmpty Int))) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Arg Term)) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Int))))) :+: (C1 ('MetaCons "DoesNotMentionTicks" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Arg Term)))) :+: (C1 ('MetaCons "MismatchedProjectionsError" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "AttributeKindNotEnabled" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String)))))) :+: ((C1 ('MetaCons "InvalidProjectionParameter" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (NamedArg Expr))) :+: (C1 ('MetaCons "TacticAttributeNotAllowed" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "CannotRewriteByNonEquation" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)))) :+: (C1 ('MetaCons "MacroResultTypeMismatch" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)) :+: (C1 ('MetaCons "NamedWhereModuleInRefinedContext" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Term]) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [String])) :+: C1 ('MetaCons "CubicalPrimitiveNotFullyApplied" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)))))) :+: (((C1 ('MetaCons "TooManyArgumentsToLeveledSort" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: (C1 ('MetaCons "TooManyArgumentsToUnivOmega" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "ComatchingDisabledForRecord" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)))) :+: (C1 ('MetaCons "BuiltinMustBeIsOne" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)) :+: (C1 ('MetaCons "IncorrectTypeForRewriteRelation" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 IncorrectTypeForRewriteRelationReason)) :+: C1 ('MetaCons "UnexpectedParameter" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 LamBinding))))) :+: ((C1 ('MetaCons "NoParameterOfName" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ArgName)) :+: (C1 ('MetaCons "UnexpectedModalityAnnotationInParameter" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 LamBinding)) :+: C1 ('MetaCons "ExpectedBindingForParameter" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Dom Type)) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Abs Type))))) :+: (C1 ('MetaCons "UnexpectedTypeSignatureForParameter" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 (NamedArg Binder)))) :+: (C1 ('MetaCons "SortDoesNotAdmitDataDefinitions" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Sort)) :+: C1 ('MetaCons "SortCannotDependOnItsIndex" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type))))))) :+: ((((C1 ('MetaCons "UnusableAtModality" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 WhyCheckModality) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Modality) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term))) :+: (C1 ('MetaCons "SplitError" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 SplitError)) :+: C1 ('MetaCons "ImpossibleConstructor" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 NegativeUnification)))) :+: (C1 ('MetaCons "TooManyPolarities" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Int)) :+: (C1 ('MetaCons "RecursiveRecordNeedsInductivity" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "CannotSolveSizeConstraints" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 (ProblemConstraint, HypSizeConstraint))) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Doc))))) :+: ((C1 ('MetaCons "ContradictorySizeConstraint" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (ProblemConstraint, HypSizeConstraint))) :+: (C1 ('MetaCons "EmptyTypeOfSizes" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)) :+: C1 ('MetaCons "FunctionTypeInSizeUniv" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)))) :+: (C1 ('MetaCons "LibraryError" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 LibErrors)) :+: (C1 ('MetaCons "LocalVsImportedModuleClash" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ModuleName)) :+: C1 ('MetaCons "SolvedButOpenHoles" 'PrefixI 'False) (U1 :: Type -> Type))))) :+: (((C1 ('MetaCons "CyclicModuleDependency" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [TopLevelModuleName])) :+: (C1 ('MetaCons "FileNotFound" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 TopLevelModuleName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [AbsolutePath])) :+: C1 ('MetaCons "OverlappingProjects" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 AbsolutePath) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 TopLevelModuleName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 TopLevelModuleName))))) :+: (C1 ('MetaCons "AmbiguousTopLevelModuleName" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 TopLevelModuleName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [AbsolutePath])) :+: (C1 ('MetaCons "ModuleNameUnexpected" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 TopLevelModuleName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 TopLevelModuleName)) :+: C1 ('MetaCons "ModuleNameDoesntMatchFileName" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 TopLevelModuleName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [AbsolutePath]))))) :+: ((C1 ('MetaCons "ClashingFileNamesFor" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ModuleName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [AbsolutePath])) :+: (C1 ('MetaCons "ModuleDefinedInOtherFile" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 TopLevelModuleName) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 AbsolutePath) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 AbsolutePath))) :+: C1 ('MetaCons "InvalidFileName" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 AbsolutePath) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 InvalidFileNameReason)))) :+: ((C1 ('MetaCons "BothWithAndRHS" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "AbstractConstructorNotInScope" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName))) :+: (C1 ('MetaCons "NotInScope" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [QName])) :+: C1 ('MetaCons "NoSuchModule" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)))))))) :+: (((((C1 ('MetaCons "AmbiguousName" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 AmbiguousNameReason)) :+: (C1 ('MetaCons "AmbiguousModule" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 ModuleName))) :+: C1 ('MetaCons "AmbiguousField" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Name) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [ModuleName])))) :+: (C1 ('MetaCons "AmbiguousConstructor" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [QName])) :+: (C1 ('MetaCons "ClashingDefinition" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe NiceDeclaration)))) :+: C1 ('MetaCons "ClashingModule" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ModuleName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ModuleName))))) :+: ((C1 ('MetaCons "ClashingImport" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Name) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: (C1 ('MetaCons "ClashingModuleImport" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Name) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ModuleName)) :+: C1 ('MetaCons "DefinitionInDifferentModule" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)))) :+: (C1 ('MetaCons "DuplicateImports" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [ImportedName])) :+: (C1 ('MetaCons "InvalidPattern" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Pattern)) :+: C1 ('MetaCons "RepeatedVariablesInPattern" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Name])))))) :+: (((C1 ('MetaCons "GeneralizeNotSupportedHere" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: (C1 ('MetaCons "GeneralizedVarInLetOpenedModule" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "MultipleFixityDecls" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [(Name, [Fixity'])])))) :+: (C1 ('MetaCons "MultiplePolarityPragmas" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Name])) :+: (C1 ('MetaCons "NotAModuleExpr" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr)) :+: C1 ('MetaCons "NotAnExpression" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr))))) :+: ((C1 ('MetaCons "NotAValidLetBinding" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 NiceDeclaration)) :+: (C1 ('MetaCons "NotValidBeforeField" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 NiceDeclaration)) :+: C1 ('MetaCons "NothingAppliedToHiddenArg" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr)))) :+: (C1 ('MetaCons "NothingAppliedToInstanceArg" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr)) :+: (C1 ('MetaCons "AsPatternInPatternSynonym" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "DotPatternInPatternSynonym" 'PrefixI 'False) (U1 :: Type -> Type)))))) :+: ((((C1 ('MetaCons "BadArgumentsToPatternSynonym" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 AmbiguousQName)) :+: (C1 ('MetaCons "TooFewArgumentsToPatternSynonym" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 AmbiguousQName)) :+: C1 ('MetaCons "CannotResolveAmbiguousPatternSynonym" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 (QName, PatternSynDefn)))))) :+: (C1 ('MetaCons "IllegalInstanceVariableInPatternSynonym" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Name)) :+: (C1 ('MetaCons "PatternSynonymArgumentShadowsConstructorOrPatternSynonym" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 LHSOrPatSyn) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Name) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 AbstractName)))) :+: C1 ('MetaCons "UnusedVariableInPatternSynonym" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Name))))) :+: ((C1 ('MetaCons "UnboundVariablesInPatternSynonym" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Name])) :+: (C1 ('MetaCons "NoParseForApplication" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List2 Expr))) :+: C1 ('MetaCons "AmbiguousParseForApplication" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List2 Expr)) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 Expr))))) :+: (C1 ('MetaCons "NoParseForLHS" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 LHSOrPatSyn) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Pattern]) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Pattern))) :+: (C1 ('MetaCons "AmbiguousParseForLHS" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 LHSOrPatSyn) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Pattern) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Pattern]))) :+: C1 ('MetaCons "AmbiguousProjection" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [QName])))))) :+: (((C1 ('MetaCons "AmbiguousOverloadedProjection" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 QName)) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Doc)) :+: (C1 ('MetaCons "OperatorInformation" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [NotationSection]) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 TypeError)) :+: C1 ('MetaCons "InstanceNoCandidate" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [(Term, TCErr)])))) :+: (C1 ('MetaCons "UnquoteFailed" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 UnquoteError)) :+: (C1 ('MetaCons "DeBruijnIndexOutOfScope" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Nat) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Telescope) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Name]))) :+: C1 ('MetaCons "NeedOptionCopatterns" 'PrefixI 'False) (U1 :: Type -> Type)))) :+: ((C1 ('MetaCons "NeedOptionRewriting" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "NeedOptionProp" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "NeedOptionTwoLevel" 'PrefixI 'False) (U1 :: Type -> Type))) :+: ((C1 ('MetaCons "NonFatalErrors" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [TCWarning])) :+: C1 ('MetaCons "InstanceSearchDepthExhausted" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Int)))) :+: (C1 ('MetaCons "TriedToCopyConstrainedPrim" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "CustomBackendError" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Doc)))))))))) |
data LHSOrPatSyn Source #
Distinguish error message when parsing lhs or pattern synonym, resp.
Instances
EmbPrj LHSOrPatSyn Source # | |||||
Defined in Agda.TypeChecking.Serialise.Instances.Errors icode :: LHSOrPatSyn -> S Int32 Source # icod_ :: LHSOrPatSyn -> S Int32 Source # value :: Int32 -> R LHSOrPatSyn Source # | |||||
NFData LHSOrPatSyn Source # | |||||
Defined in Agda.TypeChecking.Monad.Base rnf :: LHSOrPatSyn -> () | |||||
Bounded LHSOrPatSyn Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Enum LHSOrPatSyn Source # | |||||
Defined in Agda.TypeChecking.Monad.Base succ :: LHSOrPatSyn -> LHSOrPatSyn pred :: LHSOrPatSyn -> LHSOrPatSyn toEnum :: Int -> LHSOrPatSyn fromEnum :: LHSOrPatSyn -> Int enumFrom :: LHSOrPatSyn -> [LHSOrPatSyn] enumFromThen :: LHSOrPatSyn -> LHSOrPatSyn -> [LHSOrPatSyn] enumFromTo :: LHSOrPatSyn -> LHSOrPatSyn -> [LHSOrPatSyn] enumFromThenTo :: LHSOrPatSyn -> LHSOrPatSyn -> LHSOrPatSyn -> [LHSOrPatSyn] | |||||
Generic LHSOrPatSyn Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
from :: LHSOrPatSyn -> Rep LHSOrPatSyn x to :: Rep LHSOrPatSyn x -> LHSOrPatSyn | |||||
Show LHSOrPatSyn Source # | |||||
Defined in Agda.TypeChecking.Monad.Base showsPrec :: Int -> LHSOrPatSyn -> ShowS show :: LHSOrPatSyn -> String showList :: [LHSOrPatSyn] -> ShowS | |||||
Eq LHSOrPatSyn Source # | |||||
Defined in Agda.TypeChecking.Monad.Base (==) :: LHSOrPatSyn -> LHSOrPatSyn -> Bool (/=) :: LHSOrPatSyn -> LHSOrPatSyn -> Bool | |||||
type Rep LHSOrPatSyn Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep LHSOrPatSyn = D1 ('MetaData "LHSOrPatSyn" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "IsLHS" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "IsPatSyn" 'PrefixI 'False) (U1 :: Type -> Type)) |
getScope :: ReadTCState m => m ScopeInfo Source #
Get the current scope.
getContext :: MonadTCEnv m => m Context Source #
Get the current context.
CheckClause Type SpineClause | |
CheckLHS SpineLHS | |
CheckPattern Pattern Telescope Type | |
CheckPatternLinearityType Name | |
CheckPatternLinearityValue Name | |
CheckLetBinding LetBinding | |
InferExpr Expr | |
CheckExprCall Comparison Expr Type | |
CheckDotPattern Expr Term | |
CheckProjection Range QName Type | |
IsTypeCall Comparison Expr Sort | |
IsType_ Expr | |
InferVar Name | |
InferDef QName | |
CheckArguments 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 | |
CheckConArgFitsIn QName Bool Type Sort | |
CheckFunDefCall Range QName [Clause] Bool | Highlight (interactively) if and only if the boolean is |
CheckPragma Range Pragma | |
CheckPrimitive Range QName Expr | |
CheckIsEmpty Range Type | |
CheckConfluence QName QName | |
CheckModuleParameters ModuleName Telescope | |
CheckWithFunctionType Type | |
CheckSectionApplication Range Erased ModuleName ModuleApplication | |
CheckNamedWhere ModuleName | |
CheckIApplyConfluence | Checking a clause for confluence with endpoint reductions. Always
|
ScopeCheckExpr Expr | |
ScopeCheckDeclaration NiceDeclaration | |
ScopeCheckLHS QName Pattern | |
NoHighlighting | |
ModuleContents | Interaction command: show module contents. |
SetRange Range | used by |
Instances
Pretty Call Source # | |||||
HasRange Call Source # | |||||
PrettyTCM Call Source # | |||||
Defined in Agda.TypeChecking.Pretty.Call | |||||
NFData Call Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Generic Call Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
| |||||
type Rep Call Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep Call = D1 ('MetaData "Call" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (((((C1 ('MetaCons "CheckClause" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 SpineClause)) :+: C1 ('MetaCons "CheckLHS" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 SpineLHS))) :+: (C1 ('MetaCons "CheckPattern" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Pattern) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Telescope) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type))) :+: C1 ('MetaCons "CheckPatternLinearityType" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Name)))) :+: ((C1 ('MetaCons "CheckPatternLinearityValue" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Name)) :+: C1 ('MetaCons "CheckLetBinding" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 LetBinding))) :+: (C1 ('MetaCons "InferExpr" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr)) :+: (C1 ('MetaCons "CheckExprCall" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type))) :+: C1 ('MetaCons "CheckDotPattern" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)))))) :+: (((C1 ('MetaCons "CheckProjection" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type))) :+: C1 ('MetaCons "IsTypeCall" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Sort)))) :+: (C1 ('MetaCons "IsType_" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr)) :+: C1 ('MetaCons "InferVar" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Name)))) :+: ((C1 ('MetaCons "InferDef" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "CheckArguments" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [NamedArg Expr])) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe Type))))) :+: (C1 ('MetaCons "CheckMetaSolution" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 MetaId)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term))) :+: (C1 ('MetaCons "CheckTargetType" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type))) :+: C1 ('MetaCons "CheckDataDef" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [LamBinding]) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Constructor])))))))) :+: ((((C1 ('MetaCons "CheckRecDef" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [LamBinding]) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Constructor]))) :+: C1 ('MetaCons "CheckConstructor" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Telescope)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Sort) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Constructor)))) :+: (C1 ('MetaCons "CheckConArgFitsIn" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Sort))) :+: C1 ('MetaCons "CheckFunDefCall" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Clause]) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool))))) :+: ((C1 ('MetaCons "CheckPragma" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Pragma)) :+: C1 ('MetaCons "CheckPrimitive" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr)))) :+: (C1 ('MetaCons "CheckIsEmpty" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)) :+: (C1 ('MetaCons "CheckConfluence" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "CheckModuleParameters" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ModuleName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Telescope)))))) :+: (((C1 ('MetaCons "CheckWithFunctionType" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)) :+: C1 ('MetaCons "CheckSectionApplication" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Erased)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ModuleName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ModuleApplication)))) :+: (C1 ('MetaCons "CheckNamedWhere" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ModuleName)) :+: (C1 ('MetaCons "CheckIApplyConfluence" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term))) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)))) :+: C1 ('MetaCons "ScopeCheckExpr" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr))))) :+: ((C1 ('MetaCons "ScopeCheckDeclaration" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 NiceDeclaration)) :+: C1 ('MetaCons "ScopeCheckLHS" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Pattern))) :+: (C1 ('MetaCons "NoHighlighting" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "ModuleContents" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "SetRange" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range)))))))) |
class (Functor m, Applicative m, MonadFail m) => HasBuiltins (m :: Type -> Type) where Source #
Nothing
getBuiltinThing :: SomeBuiltin -> m (Maybe (Builtin PrimFun)) Source #
default getBuiltinThing :: forall (t :: (Type -> Type) -> Type -> Type) (n :: Type -> Type). (MonadTrans t, HasBuiltins n, t n ~ m) => SomeBuiltin -> m (Maybe (Builtin PrimFun)) Source #
Instances
HasBuiltins AbsToCon Source # | |
Defined in Agda.Syntax.Translation.AbstractToConcrete getBuiltinThing :: SomeBuiltin -> AbsToCon (Maybe (Builtin PrimFun)) Source # | |
HasBuiltins TerM Source # | |
Defined in Agda.Termination.Monad getBuiltinThing :: SomeBuiltin -> TerM (Maybe (Builtin PrimFun)) Source # | |
HasBuiltins ReduceM Source # | |
Defined in Agda.TypeChecking.Reduce.Monad getBuiltinThing :: SomeBuiltin -> ReduceM (Maybe (Builtin PrimFun)) Source # | |
HasBuiltins BuiltinAccess Source # | |
Defined in Agda.TypeChecking.Monad.Builtin getBuiltinThing :: SomeBuiltin -> BuiltinAccess (Maybe (Builtin PrimFun)) Source # | |
HasBuiltins NLM Source # | |
Defined in Agda.TypeChecking.Rewriting.NonLinMatch getBuiltinThing :: SomeBuiltin -> NLM (Maybe (Builtin PrimFun)) Source # | |
HasBuiltins m => HasBuiltins (PureConversionT m) Source # | |
Defined in Agda.TypeChecking.Conversion.Pure getBuiltinThing :: SomeBuiltin -> PureConversionT m (Maybe (Builtin PrimFun)) Source # | |
HasBuiltins m => HasBuiltins (BlockT m) Source # | |
Defined in Agda.TypeChecking.Monad.Builtin getBuiltinThing :: SomeBuiltin -> BlockT m (Maybe (Builtin PrimFun)) Source # | |
MonadIO m => HasBuiltins (TCMT m) Source # | |
Defined in Agda.TypeChecking.Monad.Builtin getBuiltinThing :: SomeBuiltin -> TCMT m (Maybe (Builtin PrimFun)) Source # | |
HasBuiltins m => HasBuiltins (NamesT m) Source # | |
Defined in Agda.TypeChecking.Names getBuiltinThing :: SomeBuiltin -> NamesT m (Maybe (Builtin PrimFun)) Source # | |
HasBuiltins m => HasBuiltins (ListT m) Source # | |
Defined in Agda.TypeChecking.Monad.Builtin getBuiltinThing :: SomeBuiltin -> ListT m (Maybe (Builtin PrimFun)) Source # | |
HasBuiltins m => HasBuiltins (ChangeT m) Source # | |
Defined in Agda.TypeChecking.Monad.Builtin getBuiltinThing :: SomeBuiltin -> ChangeT m (Maybe (Builtin PrimFun)) Source # | |
HasBuiltins m => HasBuiltins (MaybeT m) Source # | |
Defined in Agda.TypeChecking.Monad.Builtin getBuiltinThing :: SomeBuiltin -> MaybeT m (Maybe (Builtin PrimFun)) Source # | |
HasBuiltins m => HasBuiltins (ExceptT e m) Source # | |
Defined in Agda.TypeChecking.Monad.Builtin getBuiltinThing :: SomeBuiltin -> ExceptT e m (Maybe (Builtin PrimFun)) Source # | |
HasBuiltins m => HasBuiltins (IdentityT m) Source # | |
Defined in Agda.TypeChecking.Monad.Builtin getBuiltinThing :: SomeBuiltin -> IdentityT m (Maybe (Builtin PrimFun)) Source # | |
HasBuiltins m => HasBuiltins (ReaderT e m) Source # | |
Defined in Agda.TypeChecking.Monad.Builtin getBuiltinThing :: SomeBuiltin -> ReaderT e m (Maybe (Builtin PrimFun)) Source # | |
HasBuiltins m => HasBuiltins (StateT s m) Source # | |
Defined in Agda.TypeChecking.Monad.Builtin getBuiltinThing :: SomeBuiltin -> StateT s m (Maybe (Builtin PrimFun)) Source # | |
(HasBuiltins m, Monoid w) => HasBuiltins (WriterT w m) Source # | |
Defined in Agda.TypeChecking.Monad.Builtin getBuiltinThing :: SomeBuiltin -> WriterT w m (Maybe (Builtin PrimFun)) Source # |
getBuiltinName' :: HasBuiltins m => BuiltinId -> m (Maybe QName) Source #
getAllPatternSyns :: ReadTCState m => m PatternSynDefns Source #
Get both local and imported pattern synonyms
type Context = [ContextEntry] Source #
The Context
is a stack of ContextEntry
s.
setCurrentRange :: (MonadTrace m, HasRange x) => x -> m a -> m a Source #
Sets the current range (for error messages etc.) to the range
of the given object, if it has a range (i.e., its range is not noRange
).
notUnderOpaque :: MonadTCEnv m => m a -> m a Source #
Outside of any opaque blocks.
registerInteractionPoint :: MonadInteractionPoints m => Bool -> Range -> Maybe Nat -> m InteractionId Source #
Register an interaction point during scope checking. If there is no interaction id yet, create one.
insideDotPattern :: TCM a -> TCM a Source #
isInsideDotPattern :: TCM Bool Source #
getCurrentPath :: MonadTCEnv m => m AbsolutePath Source #
Get the path of the currently checked file
data DisplayForm Source #
A DisplayForm
is in essence a rewrite rule q ts --> dt
for a defined symbol (could be a
constructor as well) q
. The right hand side is a DisplayTerm
which is used to reify
to a
more readable Syntax
.
The patterns ts
are just terms, but the first dfPatternVars
variables are pattern variables
that matches any term.
Display | |
|
Instances
Pretty DisplayForm Source # | |||||
Defined in Agda.TypeChecking.Monad.Base pretty :: DisplayForm -> Doc Source # prettyPrec :: Int -> DisplayForm -> Doc Source # prettyList :: [DisplayForm] -> Doc Source # | |||||
NamesIn DisplayForm Source # | |||||
Defined in Agda.Syntax.Internal.Names namesAndMetasIn' :: Monoid m => (Either QName MetaId -> m) -> DisplayForm -> m Source # | |||||
KillRange DisplayForm Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Free DisplayForm Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
InstantiateFull DisplayForm Source # | |||||
Defined in Agda.TypeChecking.Reduce | |||||
Normalise DisplayForm Source # | |||||
Defined in Agda.TypeChecking.Reduce | |||||
Simplify DisplayForm Source # | |||||
Defined in Agda.TypeChecking.Reduce | |||||
EmbPrj DisplayForm Source # | |||||
Defined in Agda.TypeChecking.Serialise.Instances.Internal icode :: DisplayForm -> S Int32 Source # icod_ :: DisplayForm -> S Int32 Source # value :: Int32 -> R DisplayForm Source # | |||||
Subst DisplayForm Source # | |||||
Defined in Agda.TypeChecking.Substitute
applySubst :: Substitution' (SubstArg DisplayForm) -> DisplayForm -> DisplayForm Source # | |||||
NFData DisplayForm Source # | |||||
Defined in Agda.TypeChecking.Monad.Base rnf :: DisplayForm -> () | |||||
Generic DisplayForm Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
from :: DisplayForm -> Rep DisplayForm x to :: Rep DisplayForm x -> DisplayForm | |||||
Show DisplayForm Source # | |||||
Defined in Agda.TypeChecking.Monad.Base showsPrec :: Int -> DisplayForm -> ShowS show :: DisplayForm -> String showList :: [DisplayForm] -> ShowS | |||||
type SubstArg DisplayForm Source # | |||||
Defined in Agda.TypeChecking.Substitute | |||||
type Rep DisplayForm Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep DisplayForm = D1 ('MetaData "DisplayForm" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "Display" 'PrefixI 'True) (S1 ('MetaSel ('Just "dfPatternVars") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Nat) :*: (S1 ('MetaSel ('Just "dfPats") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Elims) :*: S1 ('MetaSel ('Just "dfRHS") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 DisplayTerm)))) |
constructorForm :: HasBuiltins m => Term -> m Term Source #
Rewrite a literal to constructor form if possible.
class (MonadTCEnv m, ReadTCState m, MonadError TCErr m, MonadBlock m, HasOptions m, MonadDebug m) => MonadConstraint (m :: Type -> Type) where Source #
Monad service class containing methods for adding and solving constraints
addConstraint :: Blocker -> Constraint -> m () Source #
Unconditionally add the constraint.
addAwakeConstraint :: Blocker -> Constraint -> m () Source #
Add constraint as awake constraint.
solveConstraint :: Constraint -> m () Source #
solveSomeAwakeConstraints :: (ProblemConstraint -> Bool) -> Bool -> m () Source #
Solve awake constraints matching the predicate. If the second argument is
True solve constraints even if already isSolvingConstraints
.
wakeConstraints :: (ProblemConstraint -> WakeUp) -> m () Source #
stealConstraints :: ProblemId -> m () Source #
modifyAwakeConstraints :: (Constraints -> Constraints) -> m () Source #
modifySleepingConstraints :: (Constraints -> Constraints) -> m () Source #
Instances
MonadConstraint TCM Source # | |
Defined in Agda.TypeChecking.Constraints addConstraint :: Blocker -> Constraint -> TCM () Source # addAwakeConstraint :: Blocker -> Constraint -> TCM () Source # solveConstraint :: Constraint -> TCM () Source # solveSomeAwakeConstraints :: (ProblemConstraint -> Bool) -> Bool -> TCM () Source # wakeConstraints :: (ProblemConstraint -> WakeUp) -> TCM () Source # stealConstraints :: ProblemId -> TCM () Source # modifyAwakeConstraints :: (Constraints -> Constraints) -> TCM () Source # modifySleepingConstraints :: (Constraints -> Constraints) -> TCM () Source # | |
(PureTCM m, MonadBlock m) => MonadConstraint (PureConversionT m) Source # | |
Defined in Agda.TypeChecking.Conversion.Pure addConstraint :: Blocker -> Constraint -> PureConversionT m () Source # addAwakeConstraint :: Blocker -> Constraint -> PureConversionT m () Source # solveConstraint :: Constraint -> PureConversionT m () Source # solveSomeAwakeConstraints :: (ProblemConstraint -> Bool) -> Bool -> PureConversionT m () Source # wakeConstraints :: (ProblemConstraint -> WakeUp) -> PureConversionT m () Source # stealConstraints :: ProblemId -> PureConversionT m () Source # modifyAwakeConstraints :: (Constraints -> Constraints) -> PureConversionT m () Source # modifySleepingConstraints :: (Constraints -> Constraints) -> PureConversionT m () Source # | |
MonadConstraint m => MonadConstraint (ReaderT e m) Source # | |
Defined in Agda.TypeChecking.Monad.Constraints addConstraint :: Blocker -> Constraint -> ReaderT e m () Source # addAwakeConstraint :: Blocker -> Constraint -> ReaderT e m () Source # solveConstraint :: Constraint -> ReaderT e m () Source # solveSomeAwakeConstraints :: (ProblemConstraint -> Bool) -> Bool -> ReaderT e m () Source # wakeConstraints :: (ProblemConstraint -> WakeUp) -> ReaderT e m () Source # stealConstraints :: ProblemId -> ReaderT e m () Source # modifyAwakeConstraints :: (Constraints -> Constraints) -> ReaderT e m () Source # modifySleepingConstraints :: (Constraints -> Constraints) -> ReaderT e m () Source # |
typeOfConst :: (HasConstInfo m, ReadTCState m) => QName -> m Type Source #
Get type of a constant, instantiated to the current context.
Closure | |
|
Instances
Functor Closure Source # | |||||
Foldable Closure Source # | |||||
Defined in Agda.TypeChecking.Monad.Base fold :: Monoid m => Closure m -> m foldMap :: Monoid m => (a -> m) -> Closure a -> m foldMap' :: Monoid m => (a -> m) -> Closure a -> m foldr :: (a -> b -> b) -> b -> Closure a -> b foldr' :: (a -> b -> b) -> b -> Closure a -> b foldl :: (b -> a -> b) -> b -> Closure a -> b foldl' :: (b -> a -> b) -> b -> Closure a -> b foldr1 :: (a -> a -> a) -> Closure a -> a foldl1 :: (a -> a -> a) -> Closure a -> a elem :: Eq a => a -> Closure a -> Bool maximum :: Ord a => Closure a -> a minimum :: Ord a => Closure a -> a | |||||
LensIsAbstract (Closure a) Source # | |||||
Defined in Agda.TypeChecking.Monad.Base lensIsAbstract :: Lens' (Closure a) IsAbstract Source # | |||||
HasRange a => HasRange (Closure a) Source # | |||||
KillRange a => KillRange (Closure a) Source # | |||||
Defined in Agda.TypeChecking.Monad.Base killRange :: KillRangeT (Closure a) Source # | |||||
MentionsMeta a => MentionsMeta (Closure a) Source # | |||||
Defined in Agda.TypeChecking.MetaVars.Mention mentionsMetas :: HashSet MetaId -> Closure a -> Bool Source # | |||||
LensTCEnv (Closure a) Source # | |||||
PrettyTCM a => PrettyTCM (Closure a) Source # | |||||
Defined in Agda.TypeChecking.Pretty | |||||
Instantiate a => Instantiate (Closure a) Source # | |||||
Defined in Agda.TypeChecking.Reduce | |||||
InstantiateFull a => InstantiateFull (Closure a) Source # | |||||
Defined in Agda.TypeChecking.Reduce | |||||
Normalise a => Normalise (Closure a) Source # | |||||
Defined in Agda.TypeChecking.Reduce | |||||
Reduce a => Reduce (Closure a) Source # | |||||
Simplify a => Simplify (Closure a) Source # | |||||
Defined in Agda.TypeChecking.Reduce | |||||
NFData a => NFData (Closure a) Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Generic (Closure a) Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
| |||||
Show a => Show (Closure a) Source # | |||||
LensClosure (Closure a) a Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
type Rep (Closure a) Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep (Closure a) = D1 ('MetaData "Closure" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "Closure" 'PrefixI 'True) ((S1 ('MetaSel ('Just "clSignature") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Signature) :*: S1 ('MetaSel ('Just "clEnv") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 TCEnv)) :*: (S1 ('MetaSel ('Just "clScope") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ScopeInfo) :*: (S1 ('MetaSel ('Just "clModuleCheckpoints") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Map ModuleName CheckpointId)) :*: S1 ('MetaSel ('Just "clValue") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 a))))) |
sizeType :: (HasBuiltins m, MonadTCEnv m, ReadTCState m) => m Type Source #
The built-in type SIZE
.
isInstanceConstraint :: Constraint -> Bool Source #
shouldPostponeInstanceSearch :: (ReadTCState m, HasOptions m) => m Bool Source #
data RecordFieldWarning Source #
Instances
EmbPrj RecordFieldWarning Source # | |||||
Defined in Agda.TypeChecking.Serialise.Instances.Errors icode :: RecordFieldWarning -> S Int32 Source # icod_ :: RecordFieldWarning -> S Int32 Source # value :: Int32 -> R RecordFieldWarning Source # | |||||
NFData RecordFieldWarning | |||||
Defined in Agda.TypeChecking.Monad.Base rnf :: RecordFieldWarning -> () | |||||
Generic RecordFieldWarning Source # | |||||
Defined in Agda.TypeChecking.Monad.Base.Warning
from :: RecordFieldWarning -> Rep RecordFieldWarning x to :: Rep RecordFieldWarning x -> RecordFieldWarning | |||||
Show RecordFieldWarning Source # | |||||
Defined in Agda.TypeChecking.Monad.Base.Warning showsPrec :: Int -> RecordFieldWarning -> ShowS show :: RecordFieldWarning -> String showList :: [RecordFieldWarning] -> ShowS | |||||
type Rep RecordFieldWarning Source # | |||||
Defined in Agda.TypeChecking.Monad.Base.Warning type Rep RecordFieldWarning = D1 ('MetaData "RecordFieldWarning" "Agda.TypeChecking.Monad.Base.Warning" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "DuplicateFields" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [(Name, Range)])) :+: C1 ('MetaCons "TooManyFields" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Name]) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [(Name, Range)])))) |
A builtin name, defined by the BUILTIN
pragma.
BuiltinNat | |
BuiltinSuc | |
BuiltinZero | |
BuiltinNatPlus | |
BuiltinNatMinus | |
BuiltinNatTimes | |
BuiltinNatDivSucAux | |
BuiltinNatModSucAux | |
BuiltinNatEquals | |
BuiltinNatLess | |
BuiltinWord64 | |
BuiltinInteger | |
BuiltinIntegerPos | |
BuiltinIntegerNegSuc | |
BuiltinFloat | |
BuiltinChar | |
BuiltinString | |
BuiltinUnit | |
BuiltinUnitUnit | |
BuiltinSigma | |
BuiltinSigmaCon | |
BuiltinBool | |
BuiltinTrue | |
BuiltinFalse | |
BuiltinList | |
BuiltinNil | |
BuiltinCons | |
BuiltinMaybe | |
BuiltinNothing | |
BuiltinJust | |
BuiltinIO | |
BuiltinId | |
BuiltinReflId | |
BuiltinPath | |
BuiltinPathP | |
BuiltinIntervalUniv | |
BuiltinInterval | |
BuiltinIZero | |
BuiltinIOne | |
BuiltinPartial | |
BuiltinPartialP | |
BuiltinIsOne | |
BuiltinItIsOne | |
BuiltinEquiv | |
BuiltinEquivFun | |
BuiltinEquivProof | |
BuiltinTranspProof | |
BuiltinIsOne1 | |
BuiltinIsOne2 | |
BuiltinIsOneEmpty | |
BuiltinSub | |
BuiltinSubIn | |
BuiltinSizeUniv | |
BuiltinSize | |
BuiltinSizeLt | |
BuiltinSizeSuc | |
BuiltinSizeInf | |
BuiltinSizeMax | |
BuiltinInf | |
BuiltinSharp | |
BuiltinFlat | |
BuiltinEquality | |
BuiltinRefl | |
BuiltinRewrite | |
BuiltinLevelMax | |
BuiltinLevel | |
BuiltinLevelZero | |
BuiltinLevelSuc | |
BuiltinProp | |
BuiltinSet | |
BuiltinStrictSet | |
BuiltinPropOmega | |
BuiltinSetOmega | |
BuiltinSSetOmega | |
BuiltinLevelUniv | |
BuiltinFromNat | |
BuiltinFromNeg | |
BuiltinFromString | |
BuiltinQName | |
BuiltinAgdaSort | |
BuiltinAgdaSortSet | |
BuiltinAgdaSortLit | |
BuiltinAgdaSortProp | |
BuiltinAgdaSortPropLit | |
BuiltinAgdaSortInf | |
BuiltinAgdaSortUnsupported | |
BuiltinHiding | |
BuiltinHidden | |
BuiltinInstance | |
BuiltinVisible | |
BuiltinRelevance | |
BuiltinRelevant | |
BuiltinIrrelevant | |
BuiltinQuantity | |
BuiltinQuantity0 | |
BuiltinQuantityω | |
BuiltinModality | |
BuiltinModalityConstructor | |
BuiltinAssoc | |
BuiltinAssocLeft | |
BuiltinAssocRight | |
BuiltinAssocNon | |
BuiltinPrecedence | |
BuiltinPrecRelated | |
BuiltinPrecUnrelated | |
BuiltinFixity | |
BuiltinFixityFixity | |
BuiltinArg | |
BuiltinArgInfo | |
BuiltinArgArgInfo | |
BuiltinArgArg | |
BuiltinAbs | |
BuiltinAbsAbs | |
BuiltinAgdaTerm | |
BuiltinAgdaTermVar | |
BuiltinAgdaTermLam | |
BuiltinAgdaTermExtLam | |
BuiltinAgdaTermDef | |
BuiltinAgdaTermCon | |
BuiltinAgdaTermPi | |
BuiltinAgdaTermSort | |
BuiltinAgdaTermLit | |
BuiltinAgdaTermUnsupported | |
BuiltinAgdaTermMeta | |
BuiltinAgdaErrorPart | |
BuiltinAgdaErrorPartString | |
BuiltinAgdaErrorPartTerm | |
BuiltinAgdaErrorPartPatt | |
BuiltinAgdaErrorPartName | |
BuiltinAgdaLiteral | |
BuiltinAgdaLitNat | |
BuiltinAgdaLitWord64 | |
BuiltinAgdaLitFloat | |
BuiltinAgdaLitChar | |
BuiltinAgdaLitString | |
BuiltinAgdaLitQName | |
BuiltinAgdaLitMeta | |
BuiltinAgdaClause | |
BuiltinAgdaClauseClause | |
BuiltinAgdaClauseAbsurd | |
BuiltinAgdaPattern | |
BuiltinAgdaPatVar | |
BuiltinAgdaPatCon | |
BuiltinAgdaPatDot | |
BuiltinAgdaPatLit | |
BuiltinAgdaPatProj | |
BuiltinAgdaPatAbsurd | |
BuiltinAgdaDefinitionFunDef | |
BuiltinAgdaDefinitionDataDef | |
BuiltinAgdaDefinitionRecordDef | |
BuiltinAgdaDefinitionDataConstructor | |
BuiltinAgdaDefinitionPostulate | |
BuiltinAgdaDefinitionPrimitive | |
BuiltinAgdaDefinition | |
BuiltinAgdaMeta | |
BuiltinAgdaTCM | |
BuiltinAgdaTCMReturn | |
BuiltinAgdaTCMBind | |
BuiltinAgdaTCMUnify | |
BuiltinAgdaTCMTypeError | |
BuiltinAgdaTCMInferType | |
BuiltinAgdaTCMCheckType | |
BuiltinAgdaTCMNormalise | |
BuiltinAgdaTCMReduce | |
BuiltinAgdaTCMCatchError | |
BuiltinAgdaTCMGetContext | |
BuiltinAgdaTCMExtendContext | |
BuiltinAgdaTCMInContext | |
BuiltinAgdaTCMFreshName | |
BuiltinAgdaTCMDeclareDef | |
BuiltinAgdaTCMDeclarePostulate | |
BuiltinAgdaTCMDeclareData | |
BuiltinAgdaTCMDefineData | |
BuiltinAgdaTCMDefineFun | |
BuiltinAgdaTCMGetType | |
BuiltinAgdaTCMGetDefinition | |
BuiltinAgdaTCMBlock | |
BuiltinAgdaTCMCommit | |
BuiltinAgdaTCMQuoteTerm | |
BuiltinAgdaTCMUnquoteTerm | |
BuiltinAgdaTCMQuoteOmegaTerm | |
BuiltinAgdaTCMIsMacro | |
BuiltinAgdaTCMWithNormalisation | |
BuiltinAgdaTCMWithReconstructed | |
BuiltinAgdaTCMWithExpandLast | |
BuiltinAgdaTCMWithReduceDefs | |
BuiltinAgdaTCMAskNormalisation | |
BuiltinAgdaTCMAskReconstructed | |
BuiltinAgdaTCMAskExpandLast | |
BuiltinAgdaTCMAskReduceDefs | |
BuiltinAgdaTCMFormatErrorParts | |
BuiltinAgdaTCMDebugPrint | |
BuiltinAgdaTCMNoConstraints | |
BuiltinAgdaTCMWorkOnTypes | |
BuiltinAgdaTCMRunSpeculative | |
BuiltinAgdaTCMExec | |
BuiltinAgdaTCMGetInstances | |
BuiltinAgdaTCMSolveInstances | |
BuiltinAgdaTCMPragmaForeign | |
BuiltinAgdaTCMPragmaCompile | |
BuiltinAgdaBlocker | |
BuiltinAgdaBlockerAny | |
BuiltinAgdaBlockerAll | |
BuiltinAgdaBlockerMeta |
Instances
IsBuiltin BuiltinId Source # | |||||
Defined in Agda.Syntax.Builtin someBuiltin :: BuiltinId -> SomeBuiltin Source # getBuiltinId :: BuiltinId -> String Source # | |||||
Pretty BuiltinId Source # | |||||
KillRange BuiltinId Source # | |||||
Defined in Agda.Syntax.Builtin | |||||
EmbPrj BuiltinId Source # | |||||
NFData BuiltinId Source # | |||||
Defined in Agda.Syntax.Builtin | |||||
Bounded BuiltinId Source # | |||||
Defined in Agda.Syntax.Builtin | |||||
Enum BuiltinId Source # | |||||
Generic BuiltinId Source # | |||||
Defined in Agda.Syntax.Builtin
| |||||
Show BuiltinId Source # | |||||
Eq BuiltinId Source # | |||||
Ord BuiltinId Source # | |||||
Defined in Agda.Syntax.Builtin | |||||
Hashable BuiltinId Source # | |||||
Defined in Agda.Syntax.Builtin hashWithSalt :: Int -> BuiltinId -> Int | |||||
type Rep BuiltinId Source # | |||||
Defined in Agda.Syntax.Builtin type Rep BuiltinId = D1 ('MetaData "BuiltinId" "Agda.Syntax.Builtin" "Agda-2.6.20240714-inplace" 'False) (((((((C1 ('MetaCons "BuiltinNat" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "BuiltinSuc" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BuiltinZero" 'PrefixI 'False) (U1 :: Type -> Type))) :+: (C1 ('MetaCons "BuiltinNatPlus" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "BuiltinNatMinus" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BuiltinNatTimes" 'PrefixI 'False) (U1 :: Type -> Type)))) :+: ((C1 ('MetaCons "BuiltinNatDivSucAux" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "BuiltinNatModSucAux" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BuiltinNatEquals" 'PrefixI 'False) (U1 :: Type -> Type))) :+: (C1 ('MetaCons "BuiltinNatLess" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "BuiltinWord64" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BuiltinInteger" 'PrefixI 'False) (U1 :: Type -> Type))))) :+: (((C1 ('MetaCons "BuiltinIntegerPos" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "BuiltinIntegerNegSuc" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BuiltinFloat" 'PrefixI 'False) (U1 :: Type -> Type))) :+: (C1 ('MetaCons "BuiltinChar" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "BuiltinString" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BuiltinUnit" 'PrefixI 'False) (U1 :: Type -> Type)))) :+: ((C1 ('MetaCons "BuiltinUnitUnit" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "BuiltinSigma" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BuiltinSigmaCon" 'PrefixI 'False) (U1 :: Type -> Type))) :+: ((C1 ('MetaCons "BuiltinBool" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BuiltinTrue" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "BuiltinFalse" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BuiltinList" 'PrefixI 'False) (U1 :: Type -> Type)))))) :+: ((((C1 ('MetaCons "BuiltinNil" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "BuiltinCons" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BuiltinMaybe" 'PrefixI 'False) (U1 :: Type -> Type))) :+: (C1 ('MetaCons "BuiltinNothing" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "BuiltinJust" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BuiltinIO" 'PrefixI 'False) (U1 :: Type -> Type)))) :+: ((C1 ('MetaCons "BuiltinId" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "BuiltinReflId" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BuiltinPath" 'PrefixI 'False) (U1 :: Type -> Type))) :+: ((C1 ('MetaCons "BuiltinPathP" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BuiltinIntervalUniv" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "BuiltinInterval" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BuiltinIZero" 'PrefixI 'False) (U1 :: Type -> Type))))) :+: (((C1 ('MetaCons "BuiltinIOne" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "BuiltinPartial" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BuiltinPartialP" 'PrefixI 'False) (U1 :: Type -> Type))) :+: (C1 ('MetaCons "BuiltinIsOne" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "BuiltinItIsOne" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BuiltinEquiv" 'PrefixI 'False) (U1 :: Type -> Type)))) :+: ((C1 ('MetaCons "BuiltinEquivFun" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "BuiltinEquivProof" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BuiltinTranspProof" 'PrefixI 'False) (U1 :: Type -> Type))) :+: ((C1 ('MetaCons "BuiltinIsOne1" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BuiltinIsOne2" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "BuiltinIsOneEmpty" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BuiltinSub" 'PrefixI 'False) (U1 :: Type -> Type))))))) :+: (((((C1 ('MetaCons "BuiltinSubIn" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "BuiltinSizeUniv" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BuiltinSize" 'PrefixI 'False) (U1 :: Type -> Type))) :+: (C1 ('MetaCons "BuiltinSizeLt" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "BuiltinSizeSuc" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BuiltinSizeInf" 'PrefixI 'False) (U1 :: Type -> Type)))) :+: ((C1 ('MetaCons "BuiltinSizeMax" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "BuiltinInf" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BuiltinSharp" 'PrefixI 'False) (U1 :: Type -> Type))) :+: (C1 ('MetaCons "BuiltinFlat" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "BuiltinEquality" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BuiltinRefl" 'PrefixI 'False) (U1 :: Type -> Type))))) :+: (((C1 ('MetaCons "BuiltinRewrite" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "BuiltinLevelMax" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BuiltinLevel" 'PrefixI 'False) (U1 :: Type -> Type))) :+: (C1 ('MetaCons "BuiltinLevelZero" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "BuiltinLevelSuc" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BuiltinProp" 'PrefixI 'False) (U1 :: Type -> Type)))) :+: ((C1 ('MetaCons "BuiltinSet" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "BuiltinStrictSet" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BuiltinPropOmega" 'PrefixI 'False) (U1 :: Type -> Type))) :+: ((C1 ('MetaCons "BuiltinSetOmega" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BuiltinSSetOmega" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "BuiltinLevelUniv" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BuiltinFromNat" 'PrefixI 'False) (U1 :: Type -> Type)))))) :+: ((((C1 ('MetaCons "BuiltinFromNeg" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "BuiltinFromString" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BuiltinQName" 'PrefixI 'False) (U1 :: Type -> Type))) :+: (C1 ('MetaCons "BuiltinAgdaSort" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "BuiltinAgdaSortSet" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BuiltinAgdaSortLit" 'PrefixI 'False) (U1 :: Type -> Type)))) :+: ((C1 ('MetaCons "BuiltinAgdaSortProp" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "BuiltinAgdaSortPropLit" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BuiltinAgdaSortInf" 'PrefixI 'False) (U1 :: Type -> Type))) :+: ((C1 ('MetaCons "BuiltinAgdaSortUnsupported" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BuiltinHiding" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "BuiltinHidden" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BuiltinInstance" 'PrefixI 'False) (U1 :: Type -> Type))))) :+: (((C1 ('MetaCons "BuiltinVisible" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "BuiltinRelevance" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BuiltinRelevant" 'PrefixI 'False) (U1 :: Type -> Type))) :+: (C1 ('MetaCons "BuiltinIrrelevant" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "BuiltinQuantity" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BuiltinQuantity0" 'PrefixI 'False) (U1 :: Type -> Type)))) :+: ((C1 ('MetaCons "BuiltinQuantity\969" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "BuiltinModality" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BuiltinModalityConstructor" 'PrefixI 'False) (U1 :: Type -> Type))) :+: ((C1 ('MetaCons "BuiltinAssoc" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BuiltinAssocLeft" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "BuiltinAssocRight" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BuiltinAssocNon" 'PrefixI 'False) (U1 :: Type -> Type)))))))) :+: ((((((C1 ('MetaCons "BuiltinPrecedence" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "BuiltinPrecRelated" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BuiltinPrecUnrelated" 'PrefixI 'False) (U1 :: Type -> Type))) :+: (C1 ('MetaCons "BuiltinFixity" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "BuiltinFixityFixity" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BuiltinArg" 'PrefixI 'False) (U1 :: Type -> Type)))) :+: ((C1 ('MetaCons "BuiltinArgInfo" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "BuiltinArgArgInfo" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BuiltinArgArg" 'PrefixI 'False) (U1 :: Type -> Type))) :+: (C1 ('MetaCons "BuiltinAbs" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "BuiltinAbsAbs" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BuiltinAgdaTerm" 'PrefixI 'False) (U1 :: Type -> Type))))) :+: (((C1 ('MetaCons "BuiltinAgdaTermVar" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "BuiltinAgdaTermLam" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BuiltinAgdaTermExtLam" 'PrefixI 'False) (U1 :: Type -> Type))) :+: (C1 ('MetaCons "BuiltinAgdaTermDef" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "BuiltinAgdaTermCon" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BuiltinAgdaTermPi" 'PrefixI 'False) (U1 :: Type -> Type)))) :+: ((C1 ('MetaCons "BuiltinAgdaTermSort" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "BuiltinAgdaTermLit" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BuiltinAgdaTermUnsupported" 'PrefixI 'False) (U1 :: Type -> Type))) :+: ((C1 ('MetaCons "BuiltinAgdaTermMeta" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BuiltinAgdaErrorPart" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "BuiltinAgdaErrorPartString" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BuiltinAgdaErrorPartTerm" 'PrefixI 'False) (U1 :: Type -> Type)))))) :+: ((((C1 ('MetaCons "BuiltinAgdaErrorPartPatt" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "BuiltinAgdaErrorPartName" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BuiltinAgdaLiteral" 'PrefixI 'False) (U1 :: Type -> Type))) :+: (C1 ('MetaCons "BuiltinAgdaLitNat" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "BuiltinAgdaLitWord64" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BuiltinAgdaLitFloat" 'PrefixI 'False) (U1 :: Type -> Type)))) :+: ((C1 ('MetaCons "BuiltinAgdaLitChar" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "BuiltinAgdaLitString" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BuiltinAgdaLitQName" 'PrefixI 'False) (U1 :: Type -> Type))) :+: ((C1 ('MetaCons "BuiltinAgdaLitMeta" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BuiltinAgdaClause" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "BuiltinAgdaClauseClause" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BuiltinAgdaClauseAbsurd" 'PrefixI 'False) (U1 :: Type -> Type))))) :+: (((C1 ('MetaCons "BuiltinAgdaPattern" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "BuiltinAgdaPatVar" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BuiltinAgdaPatCon" 'PrefixI 'False) (U1 :: Type -> Type))) :+: (C1 ('MetaCons "BuiltinAgdaPatDot" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "BuiltinAgdaPatLit" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BuiltinAgdaPatProj" 'PrefixI 'False) (U1 :: Type -> Type)))) :+: ((C1 ('MetaCons "BuiltinAgdaPatAbsurd" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "BuiltinAgdaDefinitionFunDef" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BuiltinAgdaDefinitionDataDef" 'PrefixI 'False) (U1 :: Type -> Type))) :+: ((C1 ('MetaCons "BuiltinAgdaDefinitionRecordDef" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BuiltinAgdaDefinitionDataConstructor" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "BuiltinAgdaDefinitionPostulate" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BuiltinAgdaDefinitionPrimitive" 'PrefixI 'False) (U1 :: Type -> Type))))))) :+: (((((C1 ('MetaCons "BuiltinAgdaDefinition" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "BuiltinAgdaMeta" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BuiltinAgdaTCM" 'PrefixI 'False) (U1 :: Type -> Type))) :+: (C1 ('MetaCons "BuiltinAgdaTCMReturn" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "BuiltinAgdaTCMBind" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BuiltinAgdaTCMUnify" 'PrefixI 'False) (U1 :: Type -> Type)))) :+: ((C1 ('MetaCons "BuiltinAgdaTCMTypeError" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "BuiltinAgdaTCMInferType" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BuiltinAgdaTCMCheckType" 'PrefixI 'False) (U1 :: Type -> Type))) :+: (C1 ('MetaCons "BuiltinAgdaTCMNormalise" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "BuiltinAgdaTCMReduce" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BuiltinAgdaTCMCatchError" 'PrefixI 'False) (U1 :: Type -> Type))))) :+: (((C1 ('MetaCons "BuiltinAgdaTCMGetContext" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "BuiltinAgdaTCMExtendContext" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BuiltinAgdaTCMInContext" 'PrefixI 'False) (U1 :: Type -> Type))) :+: (C1 ('MetaCons "BuiltinAgdaTCMFreshName" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "BuiltinAgdaTCMDeclareDef" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BuiltinAgdaTCMDeclarePostulate" 'PrefixI 'False) (U1 :: Type -> Type)))) :+: ((C1 ('MetaCons "BuiltinAgdaTCMDeclareData" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "BuiltinAgdaTCMDefineData" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BuiltinAgdaTCMDefineFun" 'PrefixI 'False) (U1 :: Type -> Type))) :+: ((C1 ('MetaCons "BuiltinAgdaTCMGetType" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BuiltinAgdaTCMGetDefinition" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "BuiltinAgdaTCMBlock" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BuiltinAgdaTCMCommit" 'PrefixI 'False) (U1 :: Type -> Type)))))) :+: ((((C1 ('MetaCons "BuiltinAgdaTCMQuoteTerm" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "BuiltinAgdaTCMUnquoteTerm" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BuiltinAgdaTCMQuoteOmegaTerm" 'PrefixI 'False) (U1 :: Type -> Type))) :+: (C1 ('MetaCons "BuiltinAgdaTCMIsMacro" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "BuiltinAgdaTCMWithNormalisation" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BuiltinAgdaTCMWithReconstructed" 'PrefixI 'False) (U1 :: Type -> Type)))) :+: ((C1 ('MetaCons "BuiltinAgdaTCMWithExpandLast" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "BuiltinAgdaTCMWithReduceDefs" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BuiltinAgdaTCMAskNormalisation" 'PrefixI 'False) (U1 :: Type -> Type))) :+: ((C1 ('MetaCons "BuiltinAgdaTCMAskReconstructed" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BuiltinAgdaTCMAskExpandLast" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "BuiltinAgdaTCMAskReduceDefs" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BuiltinAgdaTCMFormatErrorParts" 'PrefixI 'False) (U1 :: Type -> Type))))) :+: (((C1 ('MetaCons "BuiltinAgdaTCMDebugPrint" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "BuiltinAgdaTCMNoConstraints" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BuiltinAgdaTCMWorkOnTypes" 'PrefixI 'False) (U1 :: Type -> Type))) :+: (C1 ('MetaCons "BuiltinAgdaTCMRunSpeculative" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "BuiltinAgdaTCMExec" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BuiltinAgdaTCMGetInstances" 'PrefixI 'False) (U1 :: Type -> Type)))) :+: ((C1 ('MetaCons "BuiltinAgdaTCMSolveInstances" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "BuiltinAgdaTCMPragmaForeign" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BuiltinAgdaTCMPragmaCompile" 'PrefixI 'False) (U1 :: Type -> Type))) :+: ((C1 ('MetaCons "BuiltinAgdaBlocker" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BuiltinAgdaBlockerAny" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "BuiltinAgdaBlockerAll" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BuiltinAgdaBlockerMeta" 'PrefixI 'False) (U1 :: Type -> Type))))))))) |
data PrimitiveId Source #
A primitive name, defined by the primitive
block.
PrimConId | |
PrimIdElim | |
PrimIMin | |
PrimIMax | |
PrimINeg | |
PrimPartial | |
PrimPartialP | |
PrimSubOut | |
PrimGlue | |
Prim_glue | |
Prim_unglue | |
Prim_glueU | |
Prim_unglueU | |
PrimFaceForall | |
PrimComp | |
PrimPOr | |
PrimTrans | |
PrimDepIMin | |
PrimIdFace | |
PrimIdPath | |
PrimHComp | |
PrimShowInteger | |
PrimNatPlus | |
PrimNatMinus | |
PrimNatTimes | |
PrimNatDivSucAux | |
PrimNatModSucAux | |
PrimNatEquality | |
PrimNatLess | |
PrimShowNat | |
PrimWord64FromNat | |
PrimWord64ToNat | |
PrimWord64ToNatInjective | |
PrimLevelZero | |
PrimLevelSuc | |
PrimLevelMax | |
PrimFloatEquality | |
PrimFloatInequality | |
PrimFloatLess | |
PrimFloatIsInfinite | |
PrimFloatIsNaN | |
PrimFloatIsNegativeZero | |
PrimFloatIsSafeInteger | |
PrimFloatToWord64 | |
PrimFloatToWord64Injective | |
PrimNatToFloat | |
PrimIntToFloat | |
PrimFloatRound | |
PrimFloatFloor | |
PrimFloatCeiling | |
PrimFloatToRatio | |
PrimRatioToFloat | |
PrimFloatDecode | |
PrimFloatEncode | |
PrimShowFloat | |
PrimFloatPlus | |
PrimFloatMinus | |
PrimFloatTimes | |
PrimFloatNegate | |
PrimFloatDiv | |
PrimFloatPow | |
PrimFloatSqrt | |
PrimFloatExp | |
PrimFloatLog | |
PrimFloatSin | |
PrimFloatCos | |
PrimFloatTan | |
PrimFloatASin | |
PrimFloatACos | |
PrimFloatATan | |
PrimFloatATan2 | |
PrimFloatSinh | |
PrimFloatCosh | |
PrimFloatTanh | |
PrimFloatASinh | |
PrimFloatACosh | |
PrimFloatATanh | |
PrimCharEquality | |
PrimIsLower | |
PrimIsDigit | |
PrimIsAlpha | |
PrimIsSpace | |
PrimIsAscii | |
PrimIsLatin1 | |
PrimIsPrint | |
PrimIsHexDigit | |
PrimToUpper | |
PrimToLower | |
PrimCharToNat | |
PrimCharToNatInjective | |
PrimNatToChar | |
PrimShowChar | |
PrimStringToList | |
PrimStringToListInjective | |
PrimStringFromList | |
PrimStringFromListInjective | |
PrimStringAppend | |
PrimStringEquality | |
PrimShowString | |
PrimStringUncons | |
PrimErase | |
PrimEraseEquality | |
PrimForce | |
PrimForceLemma | |
PrimQNameEquality | |
PrimQNameLess | |
PrimShowQName | |
PrimQNameFixity | |
PrimQNameToWord64s | |
PrimQNameToWord64sInjective | |
PrimMetaEquality | |
PrimMetaLess | |
PrimShowMeta | |
PrimMetaToNat | |
PrimMetaToNatInjective | |
PrimLockUniv |
Instances
IsBuiltin PrimitiveId Source # | |||||
Defined in Agda.Syntax.Builtin someBuiltin :: PrimitiveId -> SomeBuiltin Source # getBuiltinId :: PrimitiveId -> String Source # | |||||
Pretty PrimitiveId Source # | |||||
Defined in Agda.Syntax.Builtin pretty :: PrimitiveId -> Doc Source # prettyPrec :: Int -> PrimitiveId -> Doc Source # prettyList :: [PrimitiveId] -> Doc Source # | |||||
KillRange PrimitiveId Source # | |||||
Defined in Agda.Syntax.Builtin | |||||
InstantiateFull PrimitiveId Source # | |||||
Defined in Agda.TypeChecking.Reduce | |||||
EmbPrj PrimitiveId Source # | |||||
Defined in Agda.TypeChecking.Serialise.Instances.Common icode :: PrimitiveId -> S Int32 Source # icod_ :: PrimitiveId -> S Int32 Source # value :: Int32 -> R PrimitiveId Source # | |||||
NFData PrimitiveId Source # | |||||
Defined in Agda.Syntax.Builtin rnf :: PrimitiveId -> () | |||||
Bounded PrimitiveId Source # | |||||
Defined in Agda.Syntax.Builtin | |||||
Enum PrimitiveId Source # | |||||
Defined in Agda.Syntax.Builtin succ :: PrimitiveId -> PrimitiveId pred :: PrimitiveId -> PrimitiveId toEnum :: Int -> PrimitiveId fromEnum :: PrimitiveId -> Int enumFrom :: PrimitiveId -> [PrimitiveId] enumFromThen :: PrimitiveId -> PrimitiveId -> [PrimitiveId] enumFromTo :: PrimitiveId -> PrimitiveId -> [PrimitiveId] enumFromThenTo :: PrimitiveId -> PrimitiveId -> PrimitiveId -> [PrimitiveId] | |||||
Generic PrimitiveId Source # | |||||
Defined in Agda.Syntax.Builtin
from :: PrimitiveId -> Rep PrimitiveId x to :: Rep PrimitiveId x -> PrimitiveId | |||||
Show PrimitiveId Source # | |||||
Defined in Agda.Syntax.Builtin showsPrec :: Int -> PrimitiveId -> ShowS show :: PrimitiveId -> String showList :: [PrimitiveId] -> ShowS | |||||
Eq PrimitiveId Source # | |||||
Defined in Agda.Syntax.Builtin (==) :: PrimitiveId -> PrimitiveId -> Bool (/=) :: PrimitiveId -> PrimitiveId -> Bool | |||||
Ord PrimitiveId Source # | |||||
Defined in Agda.Syntax.Builtin compare :: PrimitiveId -> PrimitiveId -> Ordering (<) :: PrimitiveId -> PrimitiveId -> Bool (<=) :: PrimitiveId -> PrimitiveId -> Bool (>) :: PrimitiveId -> PrimitiveId -> Bool (>=) :: PrimitiveId -> PrimitiveId -> Bool max :: PrimitiveId -> PrimitiveId -> PrimitiveId min :: PrimitiveId -> PrimitiveId -> PrimitiveId | |||||
Hashable PrimitiveId Source # | |||||
Defined in Agda.Syntax.Builtin hashWithSalt :: Int -> PrimitiveId -> Int hash :: PrimitiveId -> Int | |||||
type Rep PrimitiveId Source # | |||||
Defined in Agda.Syntax.Builtin type Rep PrimitiveId = D1 ('MetaData "PrimitiveId" "Agda.Syntax.Builtin" "Agda-2.6.20240714-inplace" 'False) ((((((C1 ('MetaCons "PrimConId" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "PrimIdElim" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "PrimIMin" 'PrefixI 'False) (U1 :: Type -> Type))) :+: ((C1 ('MetaCons "PrimIMax" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "PrimINeg" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "PrimPartial" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "PrimPartialP" 'PrefixI 'False) (U1 :: Type -> Type)))) :+: ((C1 ('MetaCons "PrimSubOut" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "PrimGlue" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "Prim_glue" 'PrefixI 'False) (U1 :: Type -> Type))) :+: ((C1 ('MetaCons "Prim_unglue" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "Prim_glueU" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "Prim_unglueU" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "PrimFaceForall" 'PrefixI 'False) (U1 :: Type -> Type))))) :+: (((C1 ('MetaCons "PrimComp" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "PrimPOr" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "PrimTrans" 'PrefixI 'False) (U1 :: Type -> Type))) :+: ((C1 ('MetaCons "PrimDepIMin" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "PrimIdFace" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "PrimIdPath" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "PrimHComp" 'PrefixI 'False) (U1 :: Type -> Type)))) :+: (((C1 ('MetaCons "PrimShowInteger" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "PrimNatPlus" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "PrimNatMinus" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "PrimNatTimes" 'PrefixI 'False) (U1 :: Type -> Type))) :+: ((C1 ('MetaCons "PrimNatDivSucAux" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "PrimNatModSucAux" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "PrimNatEquality" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "PrimNatLess" 'PrefixI 'False) (U1 :: Type -> Type)))))) :+: ((((C1 ('MetaCons "PrimShowNat" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "PrimWord64FromNat" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "PrimWord64ToNat" 'PrefixI 'False) (U1 :: Type -> Type))) :+: ((C1 ('MetaCons "PrimWord64ToNatInjective" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "PrimLevelZero" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "PrimLevelSuc" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "PrimLevelMax" 'PrefixI 'False) (U1 :: Type -> Type)))) :+: ((C1 ('MetaCons "PrimFloatEquality" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "PrimFloatInequality" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "PrimFloatLess" 'PrefixI 'False) (U1 :: Type -> Type))) :+: ((C1 ('MetaCons "PrimFloatIsInfinite" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "PrimFloatIsNaN" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "PrimFloatIsNegativeZero" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "PrimFloatIsSafeInteger" 'PrefixI 'False) (U1 :: Type -> Type))))) :+: (((C1 ('MetaCons "PrimFloatToWord64" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "PrimFloatToWord64Injective" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "PrimNatToFloat" 'PrefixI 'False) (U1 :: Type -> Type))) :+: ((C1 ('MetaCons "PrimIntToFloat" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "PrimFloatRound" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "PrimFloatFloor" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "PrimFloatCeiling" 'PrefixI 'False) (U1 :: Type -> Type)))) :+: (((C1 ('MetaCons "PrimFloatToRatio" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "PrimRatioToFloat" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "PrimFloatDecode" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "PrimFloatEncode" 'PrefixI 'False) (U1 :: Type -> Type))) :+: ((C1 ('MetaCons "PrimShowFloat" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "PrimFloatPlus" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "PrimFloatMinus" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "PrimFloatTimes" 'PrefixI 'False) (U1 :: Type -> Type))))))) :+: (((((C1 ('MetaCons "PrimFloatNegate" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "PrimFloatDiv" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "PrimFloatPow" 'PrefixI 'False) (U1 :: Type -> Type))) :+: ((C1 ('MetaCons "PrimFloatSqrt" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "PrimFloatExp" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "PrimFloatLog" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "PrimFloatSin" 'PrefixI 'False) (U1 :: Type -> Type)))) :+: ((C1 ('MetaCons "PrimFloatCos" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "PrimFloatTan" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "PrimFloatASin" 'PrefixI 'False) (U1 :: Type -> Type))) :+: ((C1 ('MetaCons "PrimFloatACos" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "PrimFloatATan" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "PrimFloatATan2" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "PrimFloatSinh" 'PrefixI 'False) (U1 :: Type -> Type))))) :+: (((C1 ('MetaCons "PrimFloatCosh" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "PrimFloatTanh" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "PrimFloatASinh" 'PrefixI 'False) (U1 :: Type -> Type))) :+: ((C1 ('MetaCons "PrimFloatACosh" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "PrimFloatATanh" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "PrimCharEquality" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "PrimIsLower" 'PrefixI 'False) (U1 :: Type -> Type)))) :+: (((C1 ('MetaCons "PrimIsDigit" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "PrimIsAlpha" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "PrimIsSpace" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "PrimIsAscii" 'PrefixI 'False) (U1 :: Type -> Type))) :+: ((C1 ('MetaCons "PrimIsLatin1" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "PrimIsPrint" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "PrimIsHexDigit" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "PrimToUpper" 'PrefixI 'False) (U1 :: Type -> Type)))))) :+: ((((C1 ('MetaCons "PrimToLower" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "PrimCharToNat" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "PrimCharToNatInjective" 'PrefixI 'False) (U1 :: Type -> Type))) :+: ((C1 ('MetaCons "PrimNatToChar" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "PrimShowChar" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "PrimStringToList" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "PrimStringToListInjective" 'PrefixI 'False) (U1 :: Type -> Type)))) :+: ((C1 ('MetaCons "PrimStringFromList" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "PrimStringFromListInjective" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "PrimStringAppend" 'PrefixI 'False) (U1 :: Type -> Type))) :+: ((C1 ('MetaCons "PrimStringEquality" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "PrimShowString" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "PrimStringUncons" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "PrimErase" 'PrefixI 'False) (U1 :: Type -> Type))))) :+: (((C1 ('MetaCons "PrimEraseEquality" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "PrimForce" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "PrimForceLemma" 'PrefixI 'False) (U1 :: Type -> Type))) :+: ((C1 ('MetaCons "PrimQNameEquality" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "PrimQNameLess" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "PrimShowQName" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "PrimQNameFixity" 'PrefixI 'False) (U1 :: Type -> Type)))) :+: (((C1 ('MetaCons "PrimQNameToWord64s" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "PrimQNameToWord64sInjective" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "PrimMetaEquality" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "PrimMetaLess" 'PrefixI 'False) (U1 :: Type -> Type))) :+: ((C1 ('MetaCons "PrimShowMeta" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "PrimMetaToNat" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "PrimMetaToNatInjective" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "PrimLockUniv" 'PrefixI 'False) (U1 :: Type -> Type)))))))) |
writeToCurrentLog :: (MonadDebug m, MonadTCState m, ReadTCState m) => TypeCheckAction -> m () Source #
Writes a TypeCheckAction
to the current log, using the current
PostScopeState
readFromCachedLog :: (MonadDebug m, MonadTCState m, ReadTCState m) => m (Maybe (TypeCheckAction, PostScopeState)) Source #
Reads the next entry in the cached type check log, if present.
cleanCachedLog :: (MonadDebug m, MonadTCState m) => m () Source #
Empties the "to read" CachedState. To be used when it gets invalid.
cacheCurrentLog :: (MonadDebug m, MonadTCState m) => m () Source #
Caches the current type check log. Discardes the old cache. Does nothing if caching is inactive.
activateLoadedFileCache :: (HasOptions m, MonadDebug m, MonadTCState m) => m () Source #
Makes sure that the stLoadedFileCache
is Just
, with a clean
current log. Crashes is stLoadedFileCache
is already active with a
dirty log. Should be called when we start typechecking the current
file.
cachingStarts :: (MonadDebug m, MonadTCState m, ReadTCState m) => m () Source #
To be called before any write or restore calls.
areWeCaching :: ReadTCState m => m Bool Source #
localCache :: (MonadTCState m, ReadTCState m) => m a -> m a Source #
Runs the action and restores the current cache at the end of it.
withoutCache :: (MonadTCState m, ReadTCState m) => m a -> m a Source #
Runs the action without cache and restores the current cache at the end of it.
restorePostScopeState :: (MonadDebug m, MonadTCState m) => PostScopeState -> m () Source #
enterClosure :: (MonadTCEnv m, ReadTCState m, LensClosure c a) => c -> (a -> m b) -> m b Source #
solvingProblem :: MonadConstraint m => ProblemId -> m a -> m a Source #
type VerboseKey = String Source #
type VerboseLevel = Int Source #
currentModule :: MonadTCEnv m => m ModuleName Source #
Get the name of the current module, if any.
addImport :: TopLevelModuleName -> TCM () Source #
addImportCycleCheck :: TopLevelModuleName -> TCM a -> TCM a Source #
checkForImportCycle :: TCM () Source #
Assumes that the first module in the import path is the module we are worried about.
dropDecodedModule :: TopLevelModuleName -> TCM () Source #
getPrettyVisitedModules :: ReadTCState m => m Doc Source #
getVisitedModule :: ReadTCState m => TopLevelModuleName -> m (Maybe ModuleInfo) Source #
getVisitedModules :: ReadTCState m => m VisitedModules Source #
setDecodedModules :: DecodedModules -> TCM () Source #
setVisitedModules :: VisitedModules -> TCM () Source #
storeDecodedModule :: ModuleInfo -> TCM () Source #
visitModule :: ModuleInfo -> TCM () Source #
withImportPath :: [TopLevelModuleName] -> TCM a -> TCM a Source #
reportSLn :: MonadDebug m => VerboseKey -> VerboseLevel -> String -> m () Source #
hideAndRelParams :: (LensHiding a, LensRelevance a) => a -> a Source #
Prepare parts of a parameter telescope for abstraction in constructors and projections.
noMutualBlock :: TCM a -> TCM a Source #
makeOpen :: (ReadTCState m, MonadTCEnv m) => a -> m (Open a) Source #
Create an open term in the current context.
getOpen :: (TermSubst a, MonadTCEnv m) => Open a -> m a Source #
Extract the value from an open term. The checkpoint at which it was created must be in scope.
tryGetOpen :: (TermSubst a, ReadTCState m, MonadTCEnv m) => (Substitution -> a -> Maybe a) -> Open a -> m (Maybe a) Source #
Extract the value from an open term. If the checkpoint is no longer in scope use the provided function to pull the object to the most recent common checkpoint. The function is given the substitution from the common ancestor to the checkpoint of the thing.
currentModuleNameHash :: ReadTCState m => m ModuleNameHash Source #
reportSDoc :: MonadDebug m => VerboseKey -> VerboseLevel -> TCM Doc -> m () Source #
Conditionally render debug Doc
and print it.
setPragmaOptions :: PragmaOptions -> TCM () Source #
Sets the pragma options. Checks for unsafe combinations.
class MonadTCEnv m => MonadAddContext (m :: Type -> Type) where Source #
Nothing
addCtx :: Name -> Dom Type -> m a -> m a Source #
addCtx x arg cont
add a variable to the context.
Chooses an unused Name
.
Warning: Does not update module parameter substitution!
default addCtx :: forall (n :: Type -> Type) (t :: (Type -> Type) -> Type -> Type) a. (MonadAddContext n, MonadTransControl t, t n ~ m) => Name -> Dom Type -> m a -> m a Source #
addLetBinding' :: Origin -> Name -> Term -> Dom Type -> m a -> m a Source #
Add a let bound variable to the context
default addLetBinding' :: forall (n :: Type -> Type) (t :: (Type -> Type) -> Type -> Type) a. (MonadAddContext n, MonadTransControl t, t n ~ m) => Origin -> Name -> Term -> Dom Type -> m a -> m a Source #
updateContext :: Substitution -> (Context -> Context) -> m a -> m a Source #
Update the context. Requires a substitution that transports things living in the old context to the new.
default updateContext :: forall (n :: Type -> Type) (t :: (Type -> Type) -> Type -> Type) a. (MonadAddContext n, MonadTransControl t, t n ~ m) => Substitution -> (Context -> Context) -> m a -> m a Source #
withFreshName :: Range -> ArgName -> (Name -> m a) -> m a Source #
default withFreshName :: forall (n :: Type -> Type) (t :: (Type -> Type) -> Type -> Type) a. (MonadAddContext n, MonadTransControl t, t n ~ m) => Range -> ArgName -> (Name -> m a) -> m a Source #
Instances
MonadAddContext AbsToCon Source # | |
Defined in Agda.Syntax.Translation.AbstractToConcrete addCtx :: Name -> Dom Type -> AbsToCon a -> AbsToCon a Source # addLetBinding' :: Origin -> Name -> Term -> Dom Type -> AbsToCon a -> AbsToCon a Source # updateContext :: Substitution -> (Context -> Context) -> AbsToCon a -> AbsToCon a Source # withFreshName :: Range -> ArgName -> (Name -> AbsToCon a) -> AbsToCon a Source # | |
MonadAddContext TerM Source # | |
Defined in Agda.Termination.Monad | |
MonadAddContext ReduceM Source # | |
Defined in Agda.TypeChecking.Reduce.Monad addCtx :: Name -> Dom Type -> ReduceM a -> ReduceM a Source # addLetBinding' :: Origin -> Name -> Term -> Dom Type -> ReduceM a -> ReduceM a Source # updateContext :: Substitution -> (Context -> Context) -> ReduceM a -> ReduceM a Source # withFreshName :: Range -> ArgName -> (Name -> ReduceM a) -> ReduceM a Source # | |
MonadAddContext TCM Source # | |
Defined in Agda.TypeChecking.Monad.Context | |
MonadAddContext NLM Source # | |
Defined in Agda.TypeChecking.Rewriting.NonLinMatch | |
MonadAddContext m => MonadAddContext (PureConversionT m) Source # | |
Defined in Agda.TypeChecking.Conversion.Pure addCtx :: Name -> Dom Type -> PureConversionT m a -> PureConversionT m a Source # addLetBinding' :: Origin -> Name -> Term -> Dom Type -> PureConversionT m a -> PureConversionT m a Source # updateContext :: Substitution -> (Context -> Context) -> PureConversionT m a -> PureConversionT m a Source # withFreshName :: Range -> ArgName -> (Name -> PureConversionT m a) -> PureConversionT m a Source # | |
MonadAddContext m => MonadAddContext (BlockT m) Source # | |
Defined in Agda.TypeChecking.Monad.Context addCtx :: Name -> Dom Type -> BlockT m a -> BlockT m a Source # addLetBinding' :: Origin -> Name -> Term -> Dom Type -> BlockT m a -> BlockT m a Source # updateContext :: Substitution -> (Context -> Context) -> BlockT m a -> BlockT m a Source # withFreshName :: Range -> ArgName -> (Name -> BlockT m a) -> BlockT m a Source # | |
MonadAddContext m => MonadAddContext (NamesT m) Source # | |
Defined in Agda.TypeChecking.Names addCtx :: Name -> Dom Type -> NamesT m a -> NamesT m a Source # addLetBinding' :: Origin -> Name -> Term -> Dom Type -> NamesT m a -> NamesT m a Source # updateContext :: Substitution -> (Context -> Context) -> NamesT m a -> NamesT m a Source # withFreshName :: Range -> ArgName -> (Name -> NamesT m a) -> NamesT m a Source # | |
MonadAddContext m => MonadAddContext (ListT m) Source # | |
Defined in Agda.TypeChecking.Monad.Context addCtx :: Name -> Dom Type -> ListT m a -> ListT m a Source # addLetBinding' :: Origin -> Name -> Term -> Dom Type -> ListT m a -> ListT m a Source # updateContext :: Substitution -> (Context -> Context) -> ListT m a -> ListT m a Source # withFreshName :: Range -> ArgName -> (Name -> ListT m a) -> ListT m a Source # | |
MonadAddContext m => MonadAddContext (ChangeT m) Source # | |
Defined in Agda.TypeChecking.Monad.Context addCtx :: Name -> Dom Type -> ChangeT m a -> ChangeT m a Source # addLetBinding' :: Origin -> Name -> Term -> Dom Type -> ChangeT m a -> ChangeT m a Source # updateContext :: Substitution -> (Context -> Context) -> ChangeT m a -> ChangeT m a Source # withFreshName :: Range -> ArgName -> (Name -> ChangeT m a) -> ChangeT m a Source # | |
MonadAddContext m => MonadAddContext (MaybeT m) Source # | |
Defined in Agda.TypeChecking.Monad.Context addCtx :: Name -> Dom Type -> MaybeT m a -> MaybeT m a Source # addLetBinding' :: Origin -> Name -> Term -> Dom Type -> MaybeT m a -> MaybeT m a Source # updateContext :: Substitution -> (Context -> Context) -> MaybeT m a -> MaybeT m a Source # withFreshName :: Range -> ArgName -> (Name -> MaybeT m a) -> MaybeT m a Source # | |
MonadAddContext m => MonadAddContext (ExceptT e m) Source # | |
Defined in Agda.TypeChecking.Monad.Context addCtx :: Name -> Dom Type -> ExceptT e m a -> ExceptT e m a Source # addLetBinding' :: Origin -> Name -> Term -> Dom Type -> ExceptT e m a -> ExceptT e m a Source # updateContext :: Substitution -> (Context -> Context) -> ExceptT e m a -> ExceptT e m a Source # withFreshName :: Range -> ArgName -> (Name -> ExceptT e m a) -> ExceptT e m a Source # | |
MonadAddContext m => MonadAddContext (IdentityT m) Source # | |
Defined in Agda.TypeChecking.Monad.Context addCtx :: Name -> Dom Type -> IdentityT m a -> IdentityT m a Source # addLetBinding' :: Origin -> Name -> Term -> Dom Type -> IdentityT m a -> IdentityT m a Source # updateContext :: Substitution -> (Context -> Context) -> IdentityT m a -> IdentityT m a Source # withFreshName :: Range -> ArgName -> (Name -> IdentityT m a) -> IdentityT m a Source # | |
MonadAddContext m => MonadAddContext (ReaderT r m) Source # | |
Defined in Agda.TypeChecking.Monad.Context addCtx :: Name -> Dom Type -> ReaderT r m a -> ReaderT r m a Source # addLetBinding' :: Origin -> Name -> Term -> Dom Type -> ReaderT r m a -> ReaderT r m a Source # updateContext :: Substitution -> (Context -> Context) -> ReaderT r m a -> ReaderT r m a Source # withFreshName :: Range -> ArgName -> (Name -> ReaderT r m a) -> ReaderT r m a Source # | |
MonadAddContext m => MonadAddContext (StateT r m) Source # | |
Defined in Agda.TypeChecking.Monad.Context addCtx :: Name -> Dom Type -> StateT r m a -> StateT r m a Source # addLetBinding' :: Origin -> Name -> Term -> Dom Type -> StateT r m a -> StateT r m a Source # updateContext :: Substitution -> (Context -> Context) -> StateT r m a -> StateT r m a Source # withFreshName :: Range -> ArgName -> (Name -> StateT r m a) -> StateT r m a Source # | |
(Monoid w, MonadAddContext m) => MonadAddContext (WriterT w m) Source # | |
Defined in Agda.TypeChecking.Monad.Context addCtx :: Name -> Dom Type -> WriterT w m a -> WriterT w m a Source # addLetBinding' :: Origin -> Name -> Term -> Dom Type -> WriterT w m a -> WriterT w m a Source # updateContext :: Substitution -> (Context -> Context) -> WriterT w m a -> WriterT w m a Source # withFreshName :: Range -> ArgName -> (Name -> WriterT w m a) -> WriterT w m a Source # |
class (Applicative m, MonadTCEnv m, ReadTCState m, HasOptions m) => MonadReduce (m :: Type -> Type) where Source #
Nothing
liftReduce :: ReduceM a -> m a Source #
default liftReduce :: forall (t :: (Type -> Type) -> Type -> Type) (n :: Type -> Type) a. (MonadTrans t, MonadReduce n, t n ~ m) => ReduceM a -> m a Source #
Instances
MonadReduce AbsToCon Source # | |
Defined in Agda.Syntax.Translation.AbstractToConcrete liftReduce :: ReduceM a -> AbsToCon a Source # | |
MonadReduce TerM Source # | |
Defined in Agda.Termination.Monad liftReduce :: ReduceM a -> TerM a Source # | |
MonadReduce ReduceM Source # | |
Defined in Agda.TypeChecking.Monad.Base liftReduce :: ReduceM a -> ReduceM a Source # | |
MonadReduce NLM Source # | |
Defined in Agda.TypeChecking.Rewriting.NonLinMatch liftReduce :: ReduceM a -> NLM a Source # | |
MonadReduce m => MonadReduce (PureConversionT m) Source # | |
Defined in Agda.TypeChecking.Conversion.Pure liftReduce :: ReduceM a -> PureConversionT m a Source # | |
MonadReduce m => MonadReduce (BlockT m) Source # | |
Defined in Agda.TypeChecking.Monad.Base liftReduce :: ReduceM a -> BlockT m a Source # | |
MonadIO m => MonadReduce (TCMT m) Source # | |
Defined in Agda.TypeChecking.Monad.Base liftReduce :: ReduceM a -> TCMT m a Source # | |
MonadReduce m => MonadReduce (NamesT m) Source # | |
Defined in Agda.TypeChecking.Names liftReduce :: ReduceM a -> NamesT m a Source # | |
MonadReduce m => MonadReduce (ListT m) Source # | |
Defined in Agda.TypeChecking.Monad.Base liftReduce :: ReduceM a -> ListT m a Source # | |
MonadReduce m => MonadReduce (ChangeT m) Source # | |
Defined in Agda.TypeChecking.Monad.Base liftReduce :: ReduceM a -> ChangeT m a Source # | |
MonadReduce m => MonadReduce (MaybeT m) Source # | |
Defined in Agda.TypeChecking.Monad.Base liftReduce :: ReduceM a -> MaybeT m a Source # | |
MonadReduce m => MonadReduce (ExceptT err m) Source # | |
Defined in Agda.TypeChecking.Monad.Base liftReduce :: ReduceM a -> ExceptT err m a Source # | |
MonadReduce m => MonadReduce (IdentityT m) Source # | |
Defined in Agda.TypeChecking.Monad.Base liftReduce :: ReduceM a -> IdentityT m a Source # | |
MonadReduce m => MonadReduce (ReaderT r m) Source # | |
Defined in Agda.TypeChecking.Monad.Base liftReduce :: ReduceM a -> ReaderT r m a Source # | |
MonadReduce m => MonadReduce (StateT w m) Source # | |
Defined in Agda.TypeChecking.Monad.Base liftReduce :: ReduceM a -> StateT w m a Source # | |
(Monoid w, MonadReduce m) => MonadReduce (WriterT w m) Source # | |
Defined in Agda.TypeChecking.Monad.Base liftReduce :: ReduceM a -> WriterT w m a Source # |
setHardCompileTimeModeIfErased Source #
If the first argument is
, then hard
compile-time mode is enabled when the continuation is run.Erased
something
verboseS :: MonadDebug m => VerboseKey -> VerboseLevel -> m () -> m () Source #
Run a computation if a certain verbosity level is activated.
Precondition: The level must be non-negative.
resetState :: TCM () Source #
Resets the non-persistent part of the type checking state.
class ReadTCState m => MonadStatistics (m :: Type -> Type) where Source #
Nothing
modifyCounter :: String -> (Integer -> Integer) -> m () Source #
default modifyCounter :: forall (n :: Type -> Type) (t :: (Type -> Type) -> Type -> Type). (MonadStatistics n, MonadTrans t, t n ~ m) => String -> (Integer -> Integer) -> m () Source #
Instances
MonadStatistics TerM Source # | |
Defined in Agda.Termination.Monad modifyCounter :: String -> (Integer -> Integer) -> TerM () Source # | |
MonadStatistics TCM Source # | |
Defined in Agda.TypeChecking.Monad.Statistics modifyCounter :: String -> (Integer -> Integer) -> TCM () Source # | |
ReadTCState m => MonadStatistics (PureConversionT m) Source # | |
Defined in Agda.TypeChecking.Conversion.Pure modifyCounter :: String -> (Integer -> Integer) -> PureConversionT m () Source # | |
MonadStatistics m => MonadStatistics (MaybeT m) Source # | |
Defined in Agda.TypeChecking.Monad.Statistics modifyCounter :: String -> (Integer -> Integer) -> MaybeT m () Source # | |
MonadStatistics m => MonadStatistics (ExceptT e m) Source # | |
Defined in Agda.TypeChecking.Monad.Statistics modifyCounter :: String -> (Integer -> Integer) -> ExceptT e m () Source # | |
MonadStatistics m => MonadStatistics (ReaderT r m) Source # | |
Defined in Agda.TypeChecking.Monad.Statistics modifyCounter :: String -> (Integer -> Integer) -> ReaderT r m () Source # | |
MonadStatistics m => MonadStatistics (StateT s m) Source # | |
Defined in Agda.TypeChecking.Monad.Statistics modifyCounter :: String -> (Integer -> Integer) -> StateT s m () Source # | |
(MonadStatistics m, Monoid w) => MonadStatistics (WriterT w m) Source # | |
Defined in Agda.TypeChecking.Monad.Statistics modifyCounter :: String -> (Integer -> Integer) -> WriterT w m () Source # |
type Statistics = Map String Integer Source #
tick :: MonadStatistics m => String -> m () Source #
Increase specified counter by 1
.
tickN :: MonadStatistics m => String -> Integer -> m () Source #
Increase specified counter by n
.
tickMax :: MonadStatistics m => String -> Integer -> m () Source #
Set the specified counter to the maximum of its current value and n
.
getStatistics :: ReadTCState m => m Statistics Source #
Get the statistics.
modifyStatistics :: (Statistics -> Statistics) -> TCM () Source #
Modify the statistics via given function.
printStatistics :: (MonadDebug m, MonadTCEnv m, HasOptions m) => Maybe TopLevelModuleName -> Statistics -> m () Source #
Print the given statistics.
interestingCall :: Call -> Bool Source #
localTCState :: TCM a -> TCM a Source #
Restore TCState
after performing subcomputation.
In contrast to localState
, the Benchmark
info from the subcomputation is saved.
applyWhenVerboseS :: MonadDebug m => VerboseKey -> VerboseLevel -> (m a -> m a) -> m a -> m a Source #
Apply a function if a certain verbosity level is activated.
Precondition: The level must be non-negative.
defAbstract :: Definition -> IsAbstract Source #
data CompilerPragma Source #
The backends are responsible for parsing their own pragmas.
CompilerPragma Range String |
Instances
HasRange CompilerPragma Source # | |||||
Defined in Agda.TypeChecking.Monad.Base getRange :: CompilerPragma -> Range Source # | |||||
KillRange CompiledRepresentation Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
EmbPrj CompilerPragma Source # | |||||
Defined in Agda.TypeChecking.Serialise.Instances.Compilers icode :: CompilerPragma -> S Int32 Source # icod_ :: CompilerPragma -> S Int32 Source # value :: Int32 -> R CompilerPragma Source # | |||||
NFData CompilerPragma Source # | |||||
Defined in Agda.TypeChecking.Monad.Base rnf :: CompilerPragma -> () | |||||
Generic CompilerPragma Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
from :: CompilerPragma -> Rep CompilerPragma x to :: Rep CompilerPragma x -> CompilerPragma | |||||
Show CompilerPragma Source # | |||||
Defined in Agda.TypeChecking.Monad.Base showsPrec :: Int -> CompilerPragma -> ShowS show :: CompilerPragma -> String showList :: [CompilerPragma] -> ShowS | |||||
Eq CompilerPragma Source # | |||||
Defined in Agda.TypeChecking.Monad.Base (==) :: CompilerPragma -> CompilerPragma -> Bool (/=) :: CompilerPragma -> CompilerPragma -> Bool | |||||
type Rep CompilerPragma Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep CompilerPragma = D1 ('MetaData "CompilerPragma" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "CompilerPragma" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String))) |
getBuiltin' :: HasBuiltins m => BuiltinId -> m (Maybe Term) Source #
unsafeModifyContext :: MonadTCEnv tcm => (Context -> Context) -> tcm a -> tcm a Source #
Modify a Context
in a computation. Warning: does not update
the checkpoints. Use updateContext
instead.
applyDef :: HasConstInfo m => ProjOrigin -> QName -> Arg Term -> m Term Source #
Apply a function f
to its first argument, producing the proper
postfix projection if f
is a projection which is not irrelevant.
TCEnv | |
|
Instances
LensIsAbstract TCEnv Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
LensIsOpaque TCEnv Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
LensTCEnv TCEnv Source # | |||||
NFData TCEnv Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Generic TCEnv Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
| |||||
type Rep TCEnv Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep TCEnv = D1 ('MetaData "TCEnv" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "TCEnv" 'PrefixI 'True) (((((S1 ('MetaSel ('Just "envContext") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Context) :*: (S1 ('MetaSel ('Just "envLetBindings") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 LetBindings) :*: S1 ('MetaSel ('Just "envCurrentModule") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ModuleName))) :*: (S1 ('MetaSel ('Just "envCurrentPath") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe AbsolutePath)) :*: (S1 ('MetaSel ('Just "envAnonymousModules") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [(ModuleName, Nat)]) :*: S1 ('MetaSel ('Just "envImportPath") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [TopLevelModuleName])))) :*: ((S1 ('MetaSel ('Just "envMutualBlock") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe MutualId)) :*: (S1 ('MetaSel ('Just "envTerminationCheck") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (TerminationCheck ())) :*: S1 ('MetaSel ('Just "envCoverageCheck") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 CoverageCheck))) :*: ((S1 ('MetaSel ('Just "envMakeCase") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool) :*: S1 ('MetaSel ('Just "envSolvingConstraints") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool)) :*: (S1 ('MetaSel ('Just "envCheckingWhere") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool) :*: S1 ('MetaSel ('Just "envWorkingOnTypes") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool))))) :*: (((S1 ('MetaSel ('Just "envAssignMetas") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool) :*: (S1 ('MetaSel ('Just "envActiveProblems") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Set ProblemId)) :*: S1 ('MetaSel ('Just "envUnquoteProblem") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe ProblemId)))) :*: ((S1 ('MetaSel ('Just "envAbstractMode") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 AbstractMode) :*: S1 ('MetaSel ('Just "envRelevance") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Relevance)) :*: (S1 ('MetaSel ('Just "envQuantity") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Quantity) :*: S1 ('MetaSel ('Just "envHardCompileTimeMode") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool)))) :*: ((S1 ('MetaSel ('Just "envSplitOnStrict") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool) :*: (S1 ('MetaSel ('Just "envDisplayFormsEnabled") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool) :*: S1 ('MetaSel ('Just "envFoldLetBindings") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool))) :*: ((S1 ('MetaSel ('Just "envRange") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range) :*: S1 ('MetaSel ('Just "envHighlightingRange") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range)) :*: (S1 ('MetaSel ('Just "envClause") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 IPClause) :*: S1 ('MetaSel ('Just "envCall") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe (Closure Call)))))))) :*: ((((S1 ('MetaSel ('Just "envHighlightingLevel") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 HighlightingLevel) :*: (S1 ('MetaSel ('Just "envHighlightingMethod") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 HighlightingMethod) :*: S1 ('MetaSel ('Just "envExpandLast") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ExpandHidden))) :*: (S1 ('MetaSel ('Just "envAppDef") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe QName)) :*: (S1 ('MetaSel ('Just "envSimplification") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Simplification) :*: S1 ('MetaSel ('Just "envAllowedReductions") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 AllowedReductions)))) :*: ((S1 ('MetaSel ('Just "envReduceDefs") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ReduceDefs) :*: (S1 ('MetaSel ('Just "envReconstructed") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool) :*: S1 ('MetaSel ('Just "envInjectivityDepth") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Int))) :*: ((S1 ('MetaSel ('Just "envCompareBlocked") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool) :*: S1 ('MetaSel ('Just "envPrintDomainFreePi") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool)) :*: (S1 ('MetaSel ('Just "envPrintMetasBare") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool) :*: S1 ('MetaSel ('Just "envInsideDotPattern") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool))))) :*: (((S1 ('MetaSel ('Just "envUnquoteFlags") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 UnquoteFlags) :*: (S1 ('MetaSel ('Just "envInstanceDepth") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Int) :*: S1 ('MetaSel ('Just "envIsDebugPrinting") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool))) :*: ((S1 ('MetaSel ('Just "envPrintingPatternLambdas") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [QName]) :*: S1 ('MetaSel ('Just "envCallByNeed") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool)) :*: (S1 ('MetaSel ('Just "envCurrentCheckpoint") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 CheckpointId) :*: S1 ('MetaSel ('Just "envCheckpoints") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Map CheckpointId Substitution))))) :*: ((S1 ('MetaSel ('Just "envGeneralizeMetas") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 DoGeneralize) :*: (S1 ('MetaSel ('Just "envGeneralizedVars") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Map QName GeneralizedValue)) :*: S1 ('MetaSel ('Just "envActiveBackendName") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe BackendName)))) :*: ((S1 ('MetaSel ('Just "envConflComputingOverlap") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool) :*: S1 ('MetaSel ('Just "envCurrentlyElaborating") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool)) :*: (S1 ('MetaSel ('Just "envSyntacticEqualityFuel") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (Maybe Int)) :*: S1 ('MetaSel ('Just "envCurrentOpaqueId") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (Maybe OpaqueId))))))))) |
Interface | |
|
Instances
Pretty Interface Source # | |||||
InstantiateFull Interface Source # | |||||
Defined in Agda.TypeChecking.Reduce | |||||
EmbPrj Interface Source # | |||||
NFData Interface Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Generic Interface Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
| |||||
Show Interface Source # | |||||
type Rep Interface Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep Interface = D1 ('MetaData "Interface" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "Interface" 'PrefixI 'True) ((((S1 ('MetaSel ('Just "iSourceHash") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Hash) :*: (S1 ('MetaSel ('Just "iSource") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Text) :*: S1 ('MetaSel ('Just "iFileType") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 FileType))) :*: (S1 ('MetaSel ('Just "iImportedModules") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [(TopLevelModuleName, Hash)]) :*: (S1 ('MetaSel ('Just "iModuleName") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ModuleName) :*: S1 ('MetaSel ('Just "iTopLevelModuleName") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 TopLevelModuleName)))) :*: ((S1 ('MetaSel ('Just "iScope") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Map ModuleName Scope)) :*: (S1 ('MetaSel ('Just "iInsideScope") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ScopeInfo) :*: S1 ('MetaSel ('Just "iSignature") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Signature))) :*: (S1 ('MetaSel ('Just "iMetaBindings") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 RemoteMetaStore) :*: (S1 ('MetaSel ('Just "iDisplayForms") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 DisplayForms) :*: S1 ('MetaSel ('Just "iUserWarnings") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Map QName Text)))))) :*: (((S1 ('MetaSel ('Just "iImportWarning") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe Text)) :*: (S1 ('MetaSel ('Just "iBuiltin") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (BuiltinThings (PrimitiveId, QName))) :*: S1 ('MetaSel ('Just "iForeignCode") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Map BackendName ForeignCodeStack)))) :*: (S1 ('MetaSel ('Just "iHighlighting") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 HighlightingInfo) :*: (S1 ('MetaSel ('Just "iDefaultPragmaOptions") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [OptionsPragma]) :*: S1 ('MetaSel ('Just "iFilePragmaOptions") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [OptionsPragma])))) :*: ((S1 ('MetaSel ('Just "iOptionsUsed") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 PragmaOptions) :*: (S1 ('MetaSel ('Just "iPatternSyns") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 PatternSynDefns) :*: S1 ('MetaSel ('Just "iWarnings") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [TCWarning]))) :*: (S1 ('MetaSel ('Just "iPartialDefs") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Set QName)) :*: (S1 ('MetaSel ('Just "iOpaqueBlocks") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Map OpaqueId OpaqueBlock)) :*: S1 ('MetaSel ('Just "iOpaqueNames") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Map QName OpaqueId)))))))) |
type MonadTCError (m :: Type -> Type) = (MonadTCEnv m, ReadTCState m, MonadError TCErr m) Source #
The constraints needed for typeError
and similar.
genericError :: (HasCallStack, MonadTCError m) => String -> m a Source #
:: AbsolutePath | The file name. |
-> TopLevelModuleName | The top-level module name. |
-> TCM [AgdaLibFile] |
Returns the library files for a given file.
Nothing is returned if optUseLibs
is False
.
An error is raised if optUseLibs
is True
and a library file is
located too far down the directory hierarchy (see
checkLibraryFileNotTooFarDown
).
class ReportS a where Source #
Debug print some lines if the verbosity level for the given
VerboseKey
is at least VerboseLevel
.
Note: In the presence of OverloadedStrings
, just
@
reportS key level "Literate string"
gives an
Ambiguous type variable error in
GHC@.
Use the legacy functions reportSLn
and reportSDoc
instead then.
reportS :: MonadDebug m => VerboseKey -> VerboseLevel -> a -> m () Source #
Instances
ReportS Doc Source # | |
Defined in Agda.TypeChecking.Monad.Debug reportS :: MonadDebug m => VerboseKey -> VerboseLevel -> Doc -> m () Source # | |
ReportS String Source # | |
Defined in Agda.TypeChecking.Monad.Debug reportS :: MonadDebug m => VerboseKey -> VerboseLevel -> String -> m () Source # | |
ReportS (TCM Doc) Source # | |
Defined in Agda.TypeChecking.Monad.Debug reportS :: MonadDebug m => VerboseKey -> VerboseLevel -> TCM Doc -> m () Source # | |
ReportS [Doc] Source # | |
Defined in Agda.TypeChecking.Monad.Debug reportS :: MonadDebug m => VerboseKey -> VerboseLevel -> [Doc] -> m () Source # | |
ReportS [TCM Doc] Source # | |
Defined in Agda.TypeChecking.Monad.Debug reportS :: MonadDebug m => VerboseKey -> VerboseLevel -> [TCM Doc] -> m () Source # | |
ReportS [String] Source # | |
Defined in Agda.TypeChecking.Monad.Debug reportS :: MonadDebug m => VerboseKey -> VerboseLevel -> [String] -> m () Source # |
type BackendName = String Source #
TCSt | |
|
Instances
LensCommandLineOptions TCState Source # | |||||
Defined in Agda.Interaction.Options.Lenses | |||||
LensIncludePaths TCState Source # | |||||
Defined in Agda.Interaction.Options.Lenses getIncludePaths :: TCState -> [FilePath] Source # setIncludePaths :: [FilePath] -> TCState -> TCState Source # mapIncludePaths :: ([FilePath] -> [FilePath]) -> TCState -> TCState Source # getAbsoluteIncludePaths :: TCState -> [AbsolutePath] Source # setAbsoluteIncludePaths :: [AbsolutePath] -> TCState -> TCState Source # mapAbsoluteIncludePaths :: ([AbsolutePath] -> [AbsolutePath]) -> TCState -> TCState Source # | |||||
LensPersistentVerbosity TCState Source # | |||||
LensPragmaOptions TCState Source # | |||||
Defined in Agda.Interaction.Options.Lenses getPragmaOptions :: TCState -> PragmaOptions Source # setPragmaOptions :: PragmaOptions -> TCState -> TCState Source # mapPragmaOptions :: (PragmaOptions -> PragmaOptions) -> TCState -> TCState Source # | |||||
LensSafeMode TCState Source # | |||||
Defined in Agda.Interaction.Options.Lenses | |||||
LensVerbosity TCState Source # | |||||
Defined in Agda.Interaction.Options.Lenses | |||||
NFData TCState Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Generic TCState Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
| |||||
Show TCState Source # | |||||
type Rep TCState Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep TCState = D1 ('MetaData "TCState" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "TCSt" 'PrefixI 'True) (S1 ('MetaSel ('Just "stPreScopeState") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 PreScopeState) :*: (S1 ('MetaSel ('Just "stPostScopeState") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 PostScopeState) :*: S1 ('MetaSel ('Just "stPersistentState") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 PersistentTCState)))) |
class (MonadTCEnv m, ReadTCState m) => MonadInteractionPoints (m :: Type -> Type) where Source #
Nothing
freshInteractionId :: m InteractionId Source #
default freshInteractionId :: forall (t :: (Type -> Type) -> Type -> Type) (n :: Type -> Type). (MonadTrans t, MonadInteractionPoints n, t n ~ m) => m InteractionId Source #
modifyInteractionPoints :: (InteractionPoints -> InteractionPoints) -> m () Source #
default modifyInteractionPoints :: forall (t :: (Type -> Type) -> Type -> Type) (n :: Type -> Type). (MonadTrans t, MonadInteractionPoints n, t n ~ m) => (InteractionPoints -> InteractionPoints) -> m () Source #
Instances
MonadInteractionPoints AbsToCon Source # | |
MonadInteractionPoints TCM Source # | |
Defined in Agda.TypeChecking.Monad.MetaVars | |
(PureTCM m, MonadBlock m) => MonadInteractionPoints (PureConversionT m) Source # | |
Defined in Agda.TypeChecking.Conversion.Pure | |
MonadInteractionPoints m => MonadInteractionPoints (IdentityT m) Source # | |
Defined in Agda.TypeChecking.Monad.MetaVars freshInteractionId :: IdentityT m InteractionId Source # modifyInteractionPoints :: (InteractionPoints -> InteractionPoints) -> IdentityT m () Source # | |
MonadInteractionPoints m => MonadInteractionPoints (ReaderT r m) Source # | |
Defined in Agda.TypeChecking.Monad.MetaVars freshInteractionId :: ReaderT r m InteractionId Source # modifyInteractionPoints :: (InteractionPoints -> InteractionPoints) -> ReaderT r m () Source # | |
MonadInteractionPoints m => MonadInteractionPoints (StateT s m) Source # | |
Defined in Agda.TypeChecking.Monad.MetaVars freshInteractionId :: StateT s m InteractionId Source # modifyInteractionPoints :: (InteractionPoints -> InteractionPoints) -> StateT s m () Source # |
class (HasBuiltins m, HasConstInfo m, MonadAddContext m, MonadDebug m, MonadReduce m, MonadTCEnv m, ReadTCState m) => PureTCM (m :: Type -> Type) Source #
Instances
PureTCM AbsToCon Source # | |
Defined in Agda.Syntax.Translation.AbstractToConcrete | |
PureTCM TerM Source # | |
Defined in Agda.Termination.Monad | |
PureTCM ReduceM Source # | |
Defined in Agda.TypeChecking.Reduce.Monad | |
PureTCM TCM Source # | |
Defined in Agda.TypeChecking.Monad.Pure | |
PureTCM NLM Source # | |
Defined in Agda.TypeChecking.Rewriting.NonLinMatch | |
(HasBuiltins m, HasConstInfo m, MonadAddContext m, MonadReduce m) => PureTCM (PureConversionT m) Source # | |
Defined in Agda.TypeChecking.Conversion.Pure | |
PureTCM m => PureTCM (BlockT m) Source # | |
Defined in Agda.TypeChecking.Monad.Pure | |
(HasBuiltins m, HasConstInfo m, MonadAddContext m, MonadReduce m) => PureTCM (NamesT m) Source # | |
Defined in Agda.TypeChecking.Names | |
PureTCM m => PureTCM (ListT m) Source # | |
Defined in Agda.TypeChecking.Monad.Pure | |
PureTCM m => PureTCM (ChangeT m) Source # | |
Defined in Agda.TypeChecking.Monad.Pure | |
PureTCM m => PureTCM (MaybeT m) Source # | |
Defined in Agda.TypeChecking.Monad.Pure | |
PureTCM m => PureTCM (ExceptT e m) Source # | |
Defined in Agda.TypeChecking.Monad.Pure | |
PureTCM m => PureTCM (IdentityT m) Source # | |
Defined in Agda.TypeChecking.Monad.Pure | |
PureTCM m => PureTCM (ReaderT r m) Source # | |
Defined in Agda.TypeChecking.Monad.Pure | |
PureTCM m => PureTCM (StateT s m) Source # | |
Defined in Agda.TypeChecking.Monad.Pure | |
(PureTCM m, Monoid w) => PureTCM (WriterT w m) Source # | |
Defined in Agda.TypeChecking.Monad.Pure |
type RewriteRules = [RewriteRule] Source #
Signature lookup errors.
SigUnknown String | The name is not in the signature; default error message. |
SigAbstract | The name is not available, since it is abstract. |
SigCubicalNotErasure | The name is not available because it was defined in Cubical
Agda, but the current language is Erased Cubical Agda, and
|
checkpointSubstitution :: MonadTCEnv tcm => CheckpointId -> tcm Substitution Source #
Get the substitution from the context at a given checkpoint to the current context.
PrimFun | |
|
Instances
NamesIn PrimFun Source # | Note that the | ||||
Defined in Agda.Syntax.Internal.Names | |||||
Abstract PrimFun Source # | |||||
Apply PrimFun Source # | |||||
NFData PrimFun Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Generic PrimFun Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
| |||||
type Rep PrimFun Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep PrimFun = D1 ('MetaData "PrimFun" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "PrimFun" 'PrefixI 'True) ((S1 ('MetaSel ('Just "primFunName") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Just "primFunArity") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Arity)) :*: (S1 ('MetaSel ('Just "primFunArgOccurrences") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Occurrence]) :*: S1 ('MetaSel ('Just "primFunImplementation") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ([Arg Term] -> Int -> ReduceM (Reduced MaybeReducedArgs Term)))))) |
data RunMetaOccursCheck Source #
Instances
NFData RunMetaOccursCheck Source # | |||||
Defined in Agda.TypeChecking.Monad.Base rnf :: RunMetaOccursCheck -> () | |||||
Generic RunMetaOccursCheck Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
from :: RunMetaOccursCheck -> Rep RunMetaOccursCheck x to :: Rep RunMetaOccursCheck x -> RunMetaOccursCheck | |||||
Show RunMetaOccursCheck Source # | |||||
Defined in Agda.TypeChecking.Monad.Base showsPrec :: Int -> RunMetaOccursCheck -> ShowS show :: RunMetaOccursCheck -> String showList :: [RunMetaOccursCheck] -> ShowS | |||||
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 | |||||
type Rep RunMetaOccursCheck Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep RunMetaOccursCheck = D1 ('MetaData "RunMetaOccursCheck" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "RunMetaOccursCheck" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "DontRunMetaOccursCheck" 'PrefixI 'False) (U1 :: Type -> Type)) |
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
Pretty CompareAs Source # | |||||
TermLike CompareAs Source # | |||||
AllMetas CompareAs Source # | |||||
Free CompareAs Source # | |||||
MentionsMeta CompareAs Source # | |||||
Defined in Agda.TypeChecking.MetaVars.Mention mentionsMetas :: HashSet MetaId -> CompareAs -> Bool Source # | |||||
IsSizeType CompareAs Source # | |||||
Defined in Agda.TypeChecking.Monad.SizedTypes isSizeType :: (HasOptions m, HasBuiltins m) => CompareAs -> m (Maybe BoundedSize) Source # | |||||
PrettyTCM CompareAs Source # | |||||
Defined in Agda.TypeChecking.Pretty | |||||
Instantiate CompareAs Source # | |||||
Defined in Agda.TypeChecking.Reduce | |||||
InstantiateFull CompareAs Source # | |||||
Defined in Agda.TypeChecking.Reduce | |||||
IsMeta CompareAs Source # | |||||
Normalise CompareAs Source # | |||||
Defined in Agda.TypeChecking.Reduce | |||||
Reduce CompareAs Source # | |||||
Simplify CompareAs Source # | |||||
Defined in Agda.TypeChecking.Reduce | |||||
Subst CompareAs Source # | |||||
Defined in Agda.TypeChecking.Substitute applySubst :: Substitution' (SubstArg CompareAs) -> CompareAs -> CompareAs Source # | |||||
NFData CompareAs Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Generic CompareAs Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
| |||||
Show CompareAs Source # | |||||
type SubstArg CompareAs Source # | |||||
Defined in Agda.TypeChecking.Substitute | |||||
type Rep CompareAs Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep CompareAs = D1 ('MetaData "CompareAs" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "AsTermsOf" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)) :+: (C1 ('MetaCons "AsSizes" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "AsTypes" 'PrefixI 'False) (U1 :: Type -> Type))) |
data CompareDirection Source #
An extension of Comparison
to >=
.
Instances
Pretty CompareDirection Source # | |
Defined in Agda.TypeChecking.Monad.Base pretty :: CompareDirection -> Doc Source # prettyPrec :: Int -> CompareDirection -> Doc Source # prettyList :: [CompareDirection] -> Doc Source # | |
Show CompareDirection Source # | |
Defined in Agda.TypeChecking.Monad.Base showsPrec :: Int -> CompareDirection -> ShowS show :: CompareDirection -> String showList :: [CompareDirection] -> ShowS | |
Eq CompareDirection Source # | |
Defined in Agda.TypeChecking.Monad.Base (==) :: CompareDirection -> CompareDirection -> Bool (/=) :: CompareDirection -> CompareDirection -> Bool |
data MetaVariable Source #
Information about local meta-variables.
MetaVar | |
|
Instances
LensModality MetaVariable Source # | |||||
Defined in Agda.TypeChecking.Monad.Base getModality :: MetaVariable -> Modality Source # setModality :: Modality -> MetaVariable -> MetaVariable Source # mapModality :: (Modality -> Modality) -> MetaVariable -> MetaVariable Source # | |||||
LensQuantity MetaVariable Source # | |||||
Defined in Agda.TypeChecking.Monad.Base getQuantity :: MetaVariable -> Quantity Source # setQuantity :: Quantity -> MetaVariable -> MetaVariable Source # mapQuantity :: (Quantity -> Quantity) -> MetaVariable -> MetaVariable Source # | |||||
LensRelevance MetaVariable Source # | |||||
Defined in Agda.TypeChecking.Monad.Base getRelevance :: MetaVariable -> Relevance Source # setRelevance :: Relevance -> MetaVariable -> MetaVariable Source # mapRelevance :: (Relevance -> Relevance) -> MetaVariable -> MetaVariable Source # | |||||
HasRange MetaVariable Source # | |||||
Defined in Agda.TypeChecking.Monad.Base getRange :: MetaVariable -> Range Source # | |||||
SetRange MetaVariable Source # | |||||
Defined in Agda.TypeChecking.Monad.Base setRange :: Range -> MetaVariable -> MetaVariable Source # | |||||
NFData MetaVariable Source # | |||||
Defined in Agda.TypeChecking.Monad.Base rnf :: MetaVariable -> () | |||||
Generic MetaVariable Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
from :: MetaVariable -> Rep MetaVariable x to :: Rep MetaVariable x -> MetaVariable | |||||
LensClosure MetaVariable Range Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
type Rep MetaVariable Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep MetaVariable = D1 ('MetaData "MetaVariable" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "MetaVar" 'PrefixI 'True) (((S1 ('MetaSel ('Just "mvInfo") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 MetaInfo) :*: S1 ('MetaSel ('Just "mvPriority") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 MetaPriority)) :*: (S1 ('MetaSel ('Just "mvPermutation") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Permutation) :*: S1 ('MetaSel ('Just "mvJudgement") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Judgement MetaId)))) :*: ((S1 ('MetaSel ('Just "mvInstantiation") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 MetaInstantiation) :*: S1 ('MetaSel ('Just "mvListeners") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Set Listener))) :*: (S1 ('MetaSel ('Just "mvFrozen") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Frozen) :*: S1 ('MetaSel ('Just "mvTwin") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe MetaId)))))) |
class (MonadConstraint m, MonadReduce m, MonadAddContext m, MonadTCEnv m, ReadTCState m, HasBuiltins m, HasConstInfo m, MonadDebug m) => MonadMetaSolver (m :: Type -> Type) where Source #
Monad service class for creating, solving and eta-expanding of metavariables.
newMeta' :: MetaInstantiation -> Frozen -> MetaInfo -> MetaPriority -> Permutation -> Judgement a -> m MetaId Source #
Generate a new meta variable with some instantiation given.
For instance, the instantiation could be a PostponedTypeCheckingProblem
.
assignV :: CompareDirection -> MetaId -> Args -> Term -> CompareAs -> m () Source #
Assign to an open metavar which may not be frozen. First check that metavar args are in pattern fragment. Then do extended occurs check on given thing.
Assignment is aborted by throwing a PatternErr
via a call to
patternViolation
. This error is caught by catchConstraint
during equality checking (compareAtom
) and leads to
restoration of the original constraints.
assignTerm' :: MetaId -> [Arg ArgName] -> Term -> m () Source #
Directly instantiate the metavariable. Skip pattern check, occurs check and frozen check. Used for eta expanding frozen metas.
etaExpandMeta :: [MetaClass] -> MetaId -> m () Source #
Eta-expand a local meta-variable, if it is of the specified class. Don't do anything if the meta-variable is a blocked term.
updateMetaVar :: MetaId -> (MetaVariable -> MetaVariable) -> m () Source #
Update the status of the metavariable
speculateMetas :: m () -> m KeepMetas -> m () Source #
'speculateMetas fallback m' speculatively runs m
, but if the
result is RollBackMetas
any changes to metavariables are
rolled back and fallback
is run instead.
Instances
MonadMetaSolver TCM Source # | |
Defined in Agda.TypeChecking.MetaVars newMeta' :: MetaInstantiation -> Frozen -> MetaInfo -> MetaPriority -> Permutation -> Judgement a -> TCM MetaId Source # assignV :: CompareDirection -> MetaId -> Args -> Term -> CompareAs -> TCM () Source # assignTerm' :: MetaId -> [Arg ArgName] -> Term -> TCM () Source # etaExpandMeta :: [MetaClass] -> MetaId -> TCM () Source # updateMetaVar :: MetaId -> (MetaVariable -> MetaVariable) -> TCM () Source # speculateMetas :: TCM () -> TCM KeepMetas -> TCM () Source # | |
(PureTCM m, MonadBlock m) => MonadMetaSolver (PureConversionT m) Source # | |
Defined in Agda.TypeChecking.Conversion.Pure newMeta' :: MetaInstantiation -> Frozen -> MetaInfo -> MetaPriority -> Permutation -> Judgement a -> PureConversionT m MetaId Source # assignV :: CompareDirection -> MetaId -> Args -> Term -> CompareAs -> PureConversionT m () Source # assignTerm' :: MetaId -> [Arg ArgName] -> Term -> PureConversionT m () Source # etaExpandMeta :: [MetaClass] -> MetaId -> PureConversionT m () Source # updateMetaVar :: MetaId -> (MetaVariable -> MetaVariable) -> PureConversionT m () Source # speculateMetas :: PureConversionT m () -> PureConversionT m KeepMetas -> PureConversionT m () Source # | |
MonadMetaSolver m => MonadMetaSolver (ReaderT r m) Source # | |
Defined in Agda.TypeChecking.Monad.MetaVars newMeta' :: MetaInstantiation -> Frozen -> MetaInfo -> MetaPriority -> Permutation -> Judgement a -> ReaderT r m MetaId Source # assignV :: CompareDirection -> MetaId -> Args -> Term -> CompareAs -> ReaderT r m () Source # assignTerm' :: MetaId -> [Arg ArgName] -> Term -> ReaderT r m () Source # etaExpandMeta :: [MetaClass] -> MetaId -> ReaderT r m () Source # updateMetaVar :: MetaId -> (MetaVariable -> MetaVariable) -> ReaderT r m () Source # speculateMetas :: ReaderT r m () -> ReaderT r m KeepMetas -> ReaderT r m () Source # |
data InstanceInfo Source #
Information about an instance
definition.
InstanceInfo | |
|
Instances
KillRange InstanceInfo Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
EmbPrj InstanceInfo Source # | |||||
Defined in Agda.TypeChecking.Serialise.Instances.Internal icode :: InstanceInfo -> S Int32 Source # icod_ :: InstanceInfo -> S Int32 Source # value :: Int32 -> R InstanceInfo Source # | |||||
NFData InstanceInfo Source # | |||||
Defined in Agda.TypeChecking.Monad.Base rnf :: InstanceInfo -> () | |||||
Generic InstanceInfo Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
from :: InstanceInfo -> Rep InstanceInfo x to :: Rep InstanceInfo x -> InstanceInfo | |||||
Show InstanceInfo Source # | |||||
Defined in Agda.TypeChecking.Monad.Base showsPrec :: Int -> InstanceInfo -> ShowS show :: InstanceInfo -> String showList :: [InstanceInfo] -> ShowS | |||||
type Rep InstanceInfo Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep InstanceInfo = D1 ('MetaData "InstanceInfo" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "InstanceInfo" 'PrefixI 'True) (S1 ('MetaSel ('Just "instanceClass") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Just "instanceOverlap") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 OverlapMode))) |
class Monad m => MonadBlock (m :: Type -> Type) where Source #
patternViolation :: Blocker -> m a Source #
`patternViolation b` aborts the current computation
default patternViolation :: forall (t :: (Type -> Type) -> Type -> Type) (n :: Type -> Type) a. (MonadTrans t, MonadBlock n, m ~ t n) => Blocker -> m a Source #
catchPatternErr :: (Blocker -> m a) -> m a -> m a Source #
`catchPatternErr handle m` runs m, handling pattern violations
with handle
(doesn't roll back the state)
Instances
MonadBlock TCM Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
MonadBlock NLM Source # | |
Defined in Agda.TypeChecking.Rewriting.NonLinMatch | |
Monad m => MonadBlock (PureConversionT m) Source # | |
Defined in Agda.TypeChecking.Conversion.Pure patternViolation :: Blocker -> PureConversionT m a Source # catchPatternErr :: (Blocker -> PureConversionT m a) -> PureConversionT m a -> PureConversionT m a Source # | |
Monad m => MonadBlock (BlockT m) Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
MonadBlock m => MonadBlock (MaybeT m) Source # | |
Defined in Agda.TypeChecking.Monad.Base patternViolation :: Blocker -> MaybeT m a Source # catchPatternErr :: (Blocker -> MaybeT m a) -> MaybeT m a -> MaybeT m a Source # | |
Monad m => MonadBlock (ExceptT TCErr m) Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
MonadBlock m => MonadBlock (ReaderT e m) Source # | |
Defined in Agda.TypeChecking.Monad.Base patternViolation :: Blocker -> ReaderT e m a Source # catchPatternErr :: (Blocker -> ReaderT e m a) -> ReaderT e m a -> ReaderT e m a Source # |
Three cases: 1. not reduced, 2. reduced, but blocked, 3. reduced, not blocked.
NotReduced | |
Reduced (Blocked ()) |
data ProblemConstraint Source #
Instances
HasRange ProblemConstraint Source # | |||||
Defined in Agda.TypeChecking.Monad.Base getRange :: ProblemConstraint -> Range Source # | |||||
Reify ProblemConstraint Source # | |||||
Defined in Agda.Interaction.BasicOps
reify :: MonadReify m => ProblemConstraint -> m (ReifiesTo ProblemConstraint) Source # reifyWhen :: MonadReify m => Bool -> ProblemConstraint -> m (ReifiesTo ProblemConstraint) Source # | |||||
MentionsMeta ProblemConstraint Source # | |||||
Defined in Agda.TypeChecking.MetaVars.Mention mentionsMetas :: HashSet MetaId -> ProblemConstraint -> Bool Source # | |||||
PrettyTCM ProblemConstraint Source # | |||||
Defined in Agda.TypeChecking.Pretty.Constraint prettyTCM :: MonadPretty m => ProblemConstraint -> m Doc Source # | |||||
InstantiateFull ProblemConstraint Source # | |||||
Defined in Agda.TypeChecking.Reduce | |||||
Normalise ProblemConstraint Source # | |||||
Defined in Agda.TypeChecking.Reduce | |||||
Simplify ProblemConstraint Source # | |||||
Defined in Agda.TypeChecking.Reduce | |||||
NFData ProblemConstraint Source # | |||||
Defined in Agda.TypeChecking.Monad.Base rnf :: ProblemConstraint -> () | |||||
Generic ProblemConstraint Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
from :: ProblemConstraint -> Rep ProblemConstraint x to :: Rep ProblemConstraint x -> ProblemConstraint | |||||
Show ProblemConstraint Source # | |||||
Defined in Agda.TypeChecking.Monad.Base showsPrec :: Int -> ProblemConstraint -> ShowS show :: ProblemConstraint -> String showList :: [ProblemConstraint] -> ShowS | |||||
type ReifiesTo ProblemConstraint Source # | |||||
Defined in Agda.Interaction.BasicOps | |||||
type Rep ProblemConstraint Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep ProblemConstraint = D1 ('MetaData "ProblemConstraint" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "PConstr" 'PrefixI 'True) (S1 ('MetaSel ('Just "constraintProblems") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Set ProblemId)) :*: (S1 ('MetaSel ('Just "constraintUnblocker") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Blocker) :*: S1 ('MetaSel ('Just "theConstraint") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Closure Constraint))))) |
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
TermLike NLPat Source # | |||||
AllMetas NLPat Source # | |||||
NamesIn NLPat Source # | |||||
Defined in Agda.Syntax.Internal.Names | |||||
KillRange NLPat Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Free NLPat Source # | Only computes free variables that are not bound (see | ||||
PrettyTCM NLPat Source # | |||||
Defined in Agda.TypeChecking.Pretty | |||||
InstantiateFull NLPat Source # | |||||
Defined in Agda.TypeChecking.Reduce | |||||
GetMatchables NLPat Source # | |||||
Defined in Agda.TypeChecking.Rewriting.NonLinPattern getMatchables :: NLPat -> [QName] Source # | |||||
NLPatVars NLPat Source # | |||||
Defined in Agda.TypeChecking.Rewriting.NonLinPattern | |||||
EmbPrj NLPat Source # | |||||
Subst NLPat Source # | |||||
Defined in Agda.TypeChecking.Substitute applySubst :: Substitution' (SubstArg NLPat) -> NLPat -> NLPat Source # | |||||
DeBruijn NLPat Source # | |||||
Defined in Agda.TypeChecking.Substitute deBruijnVar :: Int -> NLPat Source # debruijnNamedVar :: String -> Int -> NLPat Source # deBruijnView :: NLPat -> Maybe Int Source # | |||||
NFData NLPat Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Generic NLPat Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
| |||||
Show NLPat Source # | |||||
Match NLPat Level Source # | |||||
Match NLPat Term Source # | |||||
NLPatToTerm NLPat Level Source # | |||||
Defined in Agda.TypeChecking.Rewriting.NonLinPattern | |||||
NLPatToTerm NLPat Term Source # | |||||
Defined in Agda.TypeChecking.Rewriting.NonLinPattern | |||||
PatternFrom Level NLPat Source # | |||||
Defined in Agda.TypeChecking.Rewriting.NonLinPattern | |||||
PatternFrom Term NLPat Source # | |||||
Defined in Agda.TypeChecking.Rewriting.NonLinPattern | |||||
PatternFrom Elims [Elim' NLPat] Source # | |||||
PrettyTCM (Type' NLPat) Source # | |||||
Defined in Agda.TypeChecking.Pretty | |||||
PrettyTCM (Elim' NLPat) Source # | |||||
Defined in Agda.TypeChecking.Pretty | |||||
Match [Elim' NLPat] Elims Source # | |||||
ToNLPat (Arg DeBruijnPattern) (Elim' NLPat) Source # | |||||
Defined in Agda.TypeChecking.Rewriting.Clause | |||||
ToNLPat (NamedArg DeBruijnPattern) (Elim' NLPat) Source # | |||||
Defined in Agda.TypeChecking.Rewriting.Clause | |||||
type TypeOf NLPat Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
type SubstArg NLPat Source # | |||||
Defined in Agda.TypeChecking.Substitute | |||||
type Rep NLPat Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep NLPat = D1 ('MetaData "NLPat" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) ((C1 ('MetaCons "PVar" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Int) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Arg Int])) :+: (C1 ('MetaCons "PDef" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 PElims)) :+: C1 ('MetaCons "PLam" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ArgInfo) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Abs NLPat))))) :+: ((C1 ('MetaCons "PPi" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Dom NLPType)) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Abs NLPType))) :+: C1 ('MetaCons "PSort" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 NLPSort))) :+: (C1 ('MetaCons "PBoundVar" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'SourceUnpack 'SourceStrict 'DecidedStrict) (Rec0 Int) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 PElims)) :+: C1 ('MetaCons "PTerm" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term))))) | |||||
type TypeOf [Elim' NLPat] Source # | |||||
class Monad m => MonadFresh i (m :: Type -> Type) where Source #
Nothing
default fresh :: forall (t :: (Type -> Type) -> Type -> Type) (n :: Type -> Type). (MonadTrans t, MonadFresh i n, t n ~ m) => m i Source #
Instances
MonadFresh NameId AbsToCon Source # | |
HasFresh i => MonadFresh i TCM Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
Monad m => MonadFresh NameId (PureConversionT m) Source # | |
Defined in Agda.TypeChecking.Conversion.Pure fresh :: PureConversionT m NameId Source # | |
Monad m => MonadFresh ProblemId (PureConversionT m) Source # | |
Defined in Agda.TypeChecking.Conversion.Pure fresh :: PureConversionT m ProblemId Source # | |
Monad m => MonadFresh Int (PureConversionT m) Source # | |
Defined in Agda.TypeChecking.Conversion.Pure fresh :: PureConversionT m Int Source # | |
MonadFresh i m => MonadFresh i (ListT m) Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
MonadFresh i m => MonadFresh i (IdentityT m) Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
MonadFresh i m => MonadFresh i (ReaderT r m) Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
MonadFresh i m => MonadFresh i (StateT s m) Source # | |
Defined in Agda.TypeChecking.Monad.Base |
class IsBuiltin a where Source #
The class of types which can be converted to SomeBuiltin
.
someBuiltin :: a -> SomeBuiltin Source #
Convert this value to a builtin.
getBuiltinId :: a -> String Source #
Get the identifier for this builtin, generally used for error messages.
Instances
IsBuiltin BuiltinId Source # | |
Defined in Agda.Syntax.Builtin someBuiltin :: BuiltinId -> SomeBuiltin Source # getBuiltinId :: BuiltinId -> String Source # | |
IsBuiltin PrimitiveId Source # | |
Defined in Agda.Syntax.Builtin someBuiltin :: PrimitiveId -> SomeBuiltin Source # getBuiltinId :: PrimitiveId -> String Source # | |
IsBuiltin SomeBuiltin Source # | |
Defined in Agda.Syntax.Builtin someBuiltin :: SomeBuiltin -> SomeBuiltin Source # getBuiltinId :: SomeBuiltin -> String Source # |
builtinsNoDef :: [BuiltinId] Source #
sizeBuiltins :: [BuiltinId] Source #
primitiveById :: String -> Maybe PrimitiveId Source #
Lookup a primitive by its identifier.
primConId :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primIdElim :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primIMin :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primIMax :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primINeg :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primPartial :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primPartialP :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primSubOut :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primGlue :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primFaceForall :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primHComp :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primNatPlus :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primNatMinus :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primNatTimes :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primNatDivSucAux :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primNatModSucAux :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primNatEquality :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primNatLess :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primLevelZero :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primLevelSuc :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primLevelMax :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primLockUniv :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
getPolarity :: HasConstInfo m => QName -> m [Polarity] Source #
Look up the polarity of a definition.
lookupMeta :: ReadTCState m => MetaId -> m (Maybe (Either RemoteMetaVariable MetaVariable)) Source #
Find information about the (local or remote) meta-variable, if any.
If no meta-variable is found, then the reason could be that the
dead-code elimination
(eliminateDeadCode
) failed to find the
meta-variable, perhaps because some NamesIn
instance is
incorrectly defined.
isLocal :: ReadTCState m => QName -> m Bool Source #
data ForeignCode Source #
ForeignCode Range String |
Instances
EmbPrj ForeignCode Source # | |||||
Defined in Agda.TypeChecking.Serialise.Instances.Compilers icode :: ForeignCode -> S Int32 Source # icod_ :: ForeignCode -> S Int32 Source # value :: Int32 -> R ForeignCode Source # | |||||
NFData ForeignCode Source # | |||||
Defined in Agda.TypeChecking.Monad.Base rnf :: ForeignCode -> () | |||||
Generic ForeignCode Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
from :: ForeignCode -> Rep ForeignCode x to :: Rep ForeignCode x -> ForeignCode | |||||
Show ForeignCode Source # | |||||
Defined in Agda.TypeChecking.Monad.Base showsPrec :: Int -> ForeignCode -> ShowS show :: ForeignCode -> String showList :: [ForeignCode] -> ShowS | |||||
type Rep ForeignCode Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep ForeignCode = D1 ('MetaData "ForeignCode" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "ForeignCode" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String))) |
defOpaque :: Definition -> IsOpaque Source #
data LetBinding Source #
Instances
InstantiateFull LetBinding Source # | |||||
Defined in Agda.TypeChecking.Reduce | |||||
Subst LetBinding Source # | |||||
Defined in Agda.TypeChecking.Substitute
applySubst :: Substitution' (SubstArg LetBinding) -> LetBinding -> LetBinding Source # | |||||
NFData LetBinding Source # | |||||
Defined in Agda.TypeChecking.Monad.Base rnf :: LetBinding -> () | |||||
Generic LetBinding Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
from :: LetBinding -> Rep LetBinding x to :: Rep LetBinding x -> LetBinding | |||||
Show LetBinding Source # | |||||
Defined in Agda.TypeChecking.Monad.Base showsPrec :: Int -> LetBinding -> ShowS show :: LetBinding -> String showList :: [LetBinding] -> ShowS | |||||
type SubstArg LetBinding Source # | |||||
Defined in Agda.TypeChecking.Substitute | |||||
type Rep LetBinding Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep LetBinding = D1 ('MetaData "LetBinding" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "LetBinding" 'PrefixI 'True) (S1 ('MetaSel ('Just "letOrigin") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Origin) :*: (S1 ('MetaSel ('Just "letTerm") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Just "letType") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Dom Type))))) |
Instances
Pretty Section Source # | |
NamesIn Section Source # | |
Defined in Agda.Syntax.Internal.Names | |
KillRange Section Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
KillRange Sections Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
InstantiateFull Section Source # | |
Defined in Agda.TypeChecking.Reduce | |
EmbPrj Section Source # | |
NFData Section Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
Show Section Source # | |
Eq Section | |
data PreScopeState Source #
PreScopeState | |
|
Instances
NFData PreScopeState Source # | This instance could be optimised, some things are guaranteed to be forced. | ||||
Defined in Agda.TypeChecking.Monad.Base rnf :: PreScopeState -> () | |||||
Generic PreScopeState Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
from :: PreScopeState -> Rep PreScopeState x to :: Rep PreScopeState x -> PreScopeState | |||||
type Rep PreScopeState Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep PreScopeState = D1 ('MetaData "PreScopeState" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "PreScopeState" 'PrefixI 'True) ((((S1 ('MetaSel ('Just "stPreTokens") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 HighlightingInfo) :*: S1 ('MetaSel ('Just "stPreImports") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Signature)) :*: (S1 ('MetaSel ('Just "stPreImportedModules") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (Set TopLevelModuleName)) :*: (S1 ('MetaSel ('Just "stPreModuleToSource") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 ModuleToSource) :*: S1 ('MetaSel ('Just "stPreVisitedModules") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 VisitedModules)))) :*: ((S1 ('MetaSel ('Just "stPreScope") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 ScopeInfo) :*: (S1 ('MetaSel ('Just "stPrePatternSyns") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 PatternSynDefns) :*: S1 ('MetaSel ('Just "stPrePatternSynImports") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 PatternSynDefns))) :*: (S1 ('MetaSel ('Just "stPreGeneralizedVars") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (Maybe (Set QName))) :*: (S1 ('MetaSel ('Just "stPrePragmaOptions") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 PragmaOptions) :*: S1 ('MetaSel ('Just "stPreImportedBuiltins") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (BuiltinThings PrimFun)))))) :*: (((S1 ('MetaSel ('Just "stPreImportedDisplayForms") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 DisplayForms) :*: S1 ('MetaSel ('Just "stPreFreshInteractionId") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 InteractionId)) :*: (S1 ('MetaSel ('Just "stPreImportedUserWarnings") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (Map QName Text)) :*: (S1 ('MetaSel ('Just "stPreLocalUserWarnings") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (Map QName Text)) :*: S1 ('MetaSel ('Just "stPreWarningOnImport") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (Maybe Text))))) :*: ((S1 ('MetaSel ('Just "stPreImportedPartialDefs") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (Set QName)) :*: (S1 ('MetaSel ('Just "stPreProjectConfigs") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (Map FilePath ProjectConfig)) :*: S1 ('MetaSel ('Just "stPreAgdaLibFiles") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (Map FilePath AgdaLibFile)))) :*: (S1 ('MetaSel ('Just "stPreImportedMetaStore") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 RemoteMetaStore) :*: (S1 ('MetaSel ('Just "stPreCopiedNames") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (HashMap QName QName)) :*: S1 ('MetaSel ('Just "stPreNameCopies") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (HashMap QName (HashSet QName))))))))) |
data PostScopeState Source #
PostScopeState | |
|
Instances
NFData PostScopeState Source # | This instance could be optimised, some things are guaranteed to be forced. | ||||
Defined in Agda.TypeChecking.Monad.Base rnf :: PostScopeState -> () | |||||
Generic PostScopeState Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
from :: PostScopeState -> Rep PostScopeState x to :: Rep PostScopeState x -> PostScopeState | |||||
type Rep PostScopeState Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep PostScopeState = D1 ('MetaData "PostScopeState" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "PostScopeState" 'PrefixI 'True) (((((S1 ('MetaSel ('Just "stPostSyntaxInfo") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 HighlightingInfo) :*: S1 ('MetaSel ('Just "stPostDisambiguatedNames") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 DisambiguatedNames)) :*: (S1 ('MetaSel ('Just "stPostOpenMetaStore") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 LocalMetaStore) :*: S1 ('MetaSel ('Just "stPostSolvedMetaStore") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 LocalMetaStore))) :*: ((S1 ('MetaSel ('Just "stPostInteractionPoints") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 InteractionPoints) :*: S1 ('MetaSel ('Just "stPostAwakeConstraints") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Constraints)) :*: (S1 ('MetaSel ('Just "stPostSleepingConstraints") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Constraints) :*: (S1 ('MetaSel ('Just "stPostDirty") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Bool) :*: S1 ('MetaSel ('Just "stPostOccursCheckDefs") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (Set QName)))))) :*: (((S1 ('MetaSel ('Just "stPostSignature") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Signature) :*: S1 ('MetaSel ('Just "stPostModuleCheckpoints") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (Map ModuleName CheckpointId))) :*: (S1 ('MetaSel ('Just "stPostImportsDisplayForms") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 DisplayForms) :*: S1 ('MetaSel ('Just "stPostForeignCode") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (Map BackendName ForeignCodeStack)))) :*: ((S1 ('MetaSel ('Just "stPostCurrentModule") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (Maybe (ModuleName, TopLevelModuleName))) :*: S1 ('MetaSel ('Just "stPostPendingInstances") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (Set QName))) :*: (S1 ('MetaSel ('Just "stPostTemporaryInstances") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (Set QName)) :*: (S1 ('MetaSel ('Just "stPostConcreteNames") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 ConcreteNames) :*: S1 ('MetaSel ('Just "stPostUsedNames") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (Map RawName (DList RawName)))))))) :*: ((((S1 ('MetaSel ('Just "stPostShadowingNames") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (Map Name (DList RawName))) :*: S1 ('MetaSel ('Just "stPostStatistics") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Statistics)) :*: (S1 ('MetaSel ('Just "stPostTCWarnings") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 [TCWarning]) :*: S1 ('MetaSel ('Just "stPostMutualBlocks") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (Map MutualId MutualBlock)))) :*: ((S1 ('MetaSel ('Just "stPostLocalBuiltins") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (BuiltinThings PrimFun)) :*: S1 ('MetaSel ('Just "stPostFreshMetaId") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 MetaId)) :*: (S1 ('MetaSel ('Just "stPostFreshMutualId") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 MutualId) :*: (S1 ('MetaSel ('Just "stPostFreshProblemId") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 ProblemId) :*: S1 ('MetaSel ('Just "stPostFreshCheckpointId") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 CheckpointId))))) :*: (((S1 ('MetaSel ('Just "stPostFreshInt") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Int) :*: S1 ('MetaSel ('Just "stPostFreshNameId") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 NameId)) :*: (S1 ('MetaSel ('Just "stPostFreshOpaqueId") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 OpaqueId) :*: (S1 ('MetaSel ('Just "stPostAreWeCaching") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Bool) :*: S1 ('MetaSel ('Just "stPostPostponeInstanceSearch") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Bool)))) :*: ((S1 ('MetaSel ('Just "stPostConsideringInstance") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Bool) :*: S1 ('MetaSel ('Just "stPostInstantiateBlocking") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Bool)) :*: (S1 ('MetaSel ('Just "stPostLocalPartialDefs") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (Set QName)) :*: (S1 ('MetaSel ('Just "stPostOpaqueBlocks") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Map OpaqueId OpaqueBlock)) :*: S1 ('MetaSel ('Just "stPostOpaqueIds") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Map QName OpaqueId))))))))) |
data PersistentTCState Source #
A part of the state which is not reverted when an error is thrown or the state is reset.
PersistentTCSt | |
|
Instances
LensCommandLineOptions PersistentTCState Source # | |||||
Defined in Agda.Interaction.Options.Lenses | |||||
LensIncludePaths PersistentTCState Source # | |||||
Defined in Agda.Interaction.Options.Lenses getIncludePaths :: PersistentTCState -> [FilePath] Source # setIncludePaths :: [FilePath] -> PersistentTCState -> PersistentTCState Source # mapIncludePaths :: ([FilePath] -> [FilePath]) -> PersistentTCState -> PersistentTCState Source # getAbsoluteIncludePaths :: PersistentTCState -> [AbsolutePath] Source # setAbsoluteIncludePaths :: [AbsolutePath] -> PersistentTCState -> PersistentTCState Source # mapAbsoluteIncludePaths :: ([AbsolutePath] -> [AbsolutePath]) -> PersistentTCState -> PersistentTCState Source # | |||||
LensPersistentVerbosity PersistentTCState Source # | |||||
Defined in Agda.Interaction.Options.Lenses | |||||
LensSafeMode PersistentTCState Source # | |||||
Defined in Agda.Interaction.Options.Lenses getSafeMode :: PersistentTCState -> SafeMode Source # setSafeMode :: SafeMode -> PersistentTCState -> PersistentTCState Source # mapSafeMode :: (SafeMode -> SafeMode) -> PersistentTCState -> PersistentTCState Source # | |||||
NFData PersistentTCState Source # | |||||
Defined in Agda.TypeChecking.Monad.Base rnf :: PersistentTCState -> () | |||||
Generic PersistentTCState Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
from :: PersistentTCState -> Rep PersistentTCState x to :: Rep PersistentTCState x -> PersistentTCState | |||||
type Rep PersistentTCState Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep PersistentTCState = D1 ('MetaData "PersistentTCState" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "PersistentTCSt" 'PrefixI 'True) (((S1 ('MetaSel ('Just "stDecodedModules") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 DecodedModules) :*: S1 ('MetaSel ('Just "stPersistentTopLevelModuleNames") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (BiMap RawTopLevelModuleName ModuleNameHash))) :*: (S1 ('MetaSel ('Just "stPersistentOptions") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 CommandLineOptions) :*: S1 ('MetaSel ('Just "stInteractionOutputCallback") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 InteractionOutputCallback))) :*: ((S1 ('MetaSel ('Just "stBenchmark") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Benchmark) :*: S1 ('MetaSel ('Just "stAccumStatistics") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Statistics)) :*: (S1 ('MetaSel ('Just "stPersistLoadedFileCache") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (Maybe LoadedFileCache)) :*: S1 ('MetaSel ('Just "stPersistBackends") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Backend_boot TCM]))))) |
type VisitedModules = Map TopLevelModuleName ModuleInfo Source #
type BuiltinThings pf = Map SomeBuiltin (Builtin pf) Source #
type DisplayForms = HashMap QName [LocalDisplayForm] Source #
type RemoteMetaStore = HashMap MetaId RemoteMetaVariable Source #
Used for meta-variables from other modules (and in Interface
s).
data DisambiguatedName Source #
Name disambiguation for the sake of highlighting.
Instances
NFData DisambiguatedName Source # | |||||
Defined in Agda.TypeChecking.Monad.Base rnf :: DisambiguatedName -> () | |||||
Generic DisambiguatedName Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
from :: DisambiguatedName -> Rep DisambiguatedName x to :: Rep DisambiguatedName x -> DisambiguatedName | |||||
type Rep DisambiguatedName Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep DisambiguatedName = D1 ('MetaData "DisambiguatedName" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "DisambiguatedName" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 NameKind) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName))) |
type DisambiguatedNames = IntMap DisambiguatedName Source #
type ConcreteNames = Map Name [Name] Source #
type LocalMetaStore = Map MetaId MetaVariable Source #
Used for meta-variables from the current module.
type InteractionPoints = BiMap InteractionId InteractionPoint Source #
Data structure managing the interaction points.
We never remove interaction points from this map, only set their
ipSolved
to True
. (Issue #2368)
newtype CheckpointId Source #
CheckpointId Int |
Instances
Pretty CheckpointId Source # | |
Defined in Agda.TypeChecking.Monad.Base pretty :: CheckpointId -> Doc Source # prettyPrec :: Int -> CheckpointId -> Doc Source # prettyList :: [CheckpointId] -> Doc Source # | |
HasFresh CheckpointId Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
PrettyTCM CheckpointId Source # | |
Defined in Agda.TypeChecking.Pretty prettyTCM :: MonadPretty m => CheckpointId -> m Doc Source # | |
EmbPrj CheckpointId Source # | |
Defined in Agda.TypeChecking.Serialise.Instances.Internal icode :: CheckpointId -> S Int32 Source # icod_ :: CheckpointId -> S Int32 Source # value :: Int32 -> R CheckpointId Source # | |
NFData CheckpointId Source # | |
Defined in Agda.TypeChecking.Monad.Base rnf :: CheckpointId -> () | |
Enum CheckpointId Source # | |
Defined in Agda.TypeChecking.Monad.Base succ :: CheckpointId -> CheckpointId pred :: CheckpointId -> CheckpointId toEnum :: Int -> CheckpointId fromEnum :: CheckpointId -> Int enumFrom :: CheckpointId -> [CheckpointId] enumFromThen :: CheckpointId -> CheckpointId -> [CheckpointId] enumFromTo :: CheckpointId -> CheckpointId -> [CheckpointId] enumFromThenTo :: CheckpointId -> CheckpointId -> CheckpointId -> [CheckpointId] | |
Num CheckpointId Source # | |
Defined in Agda.TypeChecking.Monad.Base (+) :: CheckpointId -> CheckpointId -> CheckpointId (-) :: CheckpointId -> CheckpointId -> CheckpointId (*) :: CheckpointId -> CheckpointId -> CheckpointId negate :: CheckpointId -> CheckpointId abs :: CheckpointId -> CheckpointId signum :: CheckpointId -> CheckpointId fromInteger :: Integer -> CheckpointId | |
Integral CheckpointId Source # | |
Defined in Agda.TypeChecking.Monad.Base quot :: CheckpointId -> CheckpointId -> CheckpointId rem :: CheckpointId -> CheckpointId -> CheckpointId div :: CheckpointId -> CheckpointId -> CheckpointId mod :: CheckpointId -> CheckpointId -> CheckpointId quotRem :: CheckpointId -> CheckpointId -> (CheckpointId, CheckpointId) divMod :: CheckpointId -> CheckpointId -> (CheckpointId, CheckpointId) toInteger :: CheckpointId -> Integer | |
Real CheckpointId Source # | |
Defined in Agda.TypeChecking.Monad.Base toRational :: CheckpointId -> Rational | |
Show CheckpointId Source # | |
Defined in Agda.TypeChecking.Monad.Base showsPrec :: Int -> CheckpointId -> ShowS show :: CheckpointId -> String showList :: [CheckpointId] -> ShowS | |
Eq CheckpointId Source # | |
Defined in Agda.TypeChecking.Monad.Base (==) :: CheckpointId -> CheckpointId -> Bool (/=) :: CheckpointId -> CheckpointId -> Bool | |
Ord CheckpointId Source # | |
Defined in Agda.TypeChecking.Monad.Base compare :: CheckpointId -> CheckpointId -> Ordering (<) :: CheckpointId -> CheckpointId -> Bool (<=) :: CheckpointId -> CheckpointId -> Bool (>) :: CheckpointId -> CheckpointId -> Bool (>=) :: CheckpointId -> CheckpointId -> Bool max :: CheckpointId -> CheckpointId -> CheckpointId min :: CheckpointId -> CheckpointId -> CheckpointId |
newtype ForeignCodeStack Source #
Foreign code fragments are stored in reversed order to support efficient appending: head points to the latest pragma in module.
Instances
EmbPrj ForeignCodeStack Source # | |||||
Defined in Agda.TypeChecking.Serialise.Instances.Compilers icode :: ForeignCodeStack -> S Int32 Source # icod_ :: ForeignCodeStack -> S Int32 Source # value :: Int32 -> R ForeignCodeStack Source # | |||||
NFData ForeignCodeStack Source # | |||||
Defined in Agda.TypeChecking.Monad.Base rnf :: ForeignCodeStack -> () | |||||
Generic ForeignCodeStack Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
from :: ForeignCodeStack -> Rep ForeignCodeStack x to :: Rep ForeignCodeStack x -> ForeignCodeStack | |||||
Show ForeignCodeStack Source # | |||||
Defined in Agda.TypeChecking.Monad.Base showsPrec :: Int -> ForeignCodeStack -> ShowS show :: ForeignCodeStack -> String showList :: [ForeignCodeStack] -> ShowS | |||||
type Rep ForeignCodeStack Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep ForeignCodeStack = D1 ('MetaData "ForeignCodeStack" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'True) (C1 ('MetaCons "ForeignCodeStack" 'PrefixI 'True) (S1 ('MetaSel ('Just "getForeignCodeStack") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [ForeignCode]))) |
MutId Int32 |
Instances
KillRange MutualId Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
HasFresh MutualId Source # | |
EmbPrj MutualId Source # | |
NFData MutualId Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
Enum MutualId Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
Num MutualId Source # | |
Show MutualId Source # | |
Eq MutualId Source # | |
Ord MutualId Source # | |
Defined in Agda.TypeChecking.Monad.Base |
data MutualBlock Source #
A mutual block of names in the signature.
MutualBlock | |
|
Instances
Null MutualBlock Source # | |||||
Defined in Agda.TypeChecking.Monad.Base empty :: MutualBlock Source # null :: MutualBlock -> Bool Source # | |||||
NFData MutualBlock Source # | |||||
Defined in Agda.TypeChecking.Monad.Base rnf :: MutualBlock -> () | |||||
Generic MutualBlock Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
from :: MutualBlock -> Rep MutualBlock x to :: Rep MutualBlock x -> MutualBlock | |||||
Show MutualBlock Source # | |||||
Defined in Agda.TypeChecking.Monad.Base showsPrec :: Int -> MutualBlock -> ShowS show :: MutualBlock -> String showList :: [MutualBlock] -> ShowS | |||||
Eq MutualBlock Source # | |||||
Defined in Agda.TypeChecking.Monad.Base (==) :: MutualBlock -> MutualBlock -> Bool (/=) :: MutualBlock -> MutualBlock -> Bool | |||||
type Rep MutualBlock Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep MutualBlock = D1 ('MetaData "MutualBlock" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "MutualBlock" 'PrefixI 'True) (S1 ('MetaSel ('Just "mutualInfo") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 MutualInfo) :*: S1 ('MetaSel ('Just "mutualNames") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Set QName)))) |
data OpaqueBlock Source #
A block of opaque definitions.
OpaqueBlock | |
|
Instances
Pretty OpaqueBlock Source # | |||||
Defined in Agda.TypeChecking.Monad.Base pretty :: OpaqueBlock -> Doc Source # prettyPrec :: Int -> OpaqueBlock -> Doc Source # prettyList :: [OpaqueBlock] -> Doc Source # | |||||
EmbPrj OpaqueBlock Source # | |||||
Defined in Agda.TypeChecking.Serialise.Instances.Internal icode :: OpaqueBlock -> S Int32 Source # icod_ :: OpaqueBlock -> S Int32 Source # value :: Int32 -> R OpaqueBlock Source # | |||||
NFData OpaqueBlock Source # | |||||
Defined in Agda.TypeChecking.Monad.Base rnf :: OpaqueBlock -> () | |||||
Generic OpaqueBlock Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
from :: OpaqueBlock -> Rep OpaqueBlock x to :: Rep OpaqueBlock x -> OpaqueBlock | |||||
Show OpaqueBlock Source # | |||||
Defined in Agda.TypeChecking.Monad.Base showsPrec :: Int -> OpaqueBlock -> ShowS show :: OpaqueBlock -> String showList :: [OpaqueBlock] -> ShowS | |||||
Eq OpaqueBlock Source # | |||||
Defined in Agda.TypeChecking.Monad.Base (==) :: OpaqueBlock -> OpaqueBlock -> Bool (/=) :: OpaqueBlock -> OpaqueBlock -> Bool | |||||
Hashable OpaqueBlock Source # | |||||
Defined in Agda.TypeChecking.Monad.Base hashWithSalt :: Int -> OpaqueBlock -> Int hash :: OpaqueBlock -> Int | |||||
type Rep OpaqueBlock Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep OpaqueBlock = D1 ('MetaData "OpaqueBlock" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "OpaqueBlock" 'PrefixI 'True) ((S1 ('MetaSel ('Just "opaqueId") 'SourceUnpack 'SourceStrict 'DecidedStrict) (Rec0 OpaqueId) :*: S1 ('MetaSel ('Just "opaqueUnfolding") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (HashSet QName))) :*: (S1 ('MetaSel ('Just "opaqueDecls") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (HashSet QName)) :*: (S1 ('MetaSel ('Just "opaqueParent") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe OpaqueId)) :*: S1 ('MetaSel ('Just "opaqueRange") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range))))) |
type DecodedModules = Map TopLevelModuleName ModuleInfo Source #
data LoadedFileCache Source #
Instances
NFData LoadedFileCache Source # | |||||
Defined in Agda.TypeChecking.Monad.Base rnf :: LoadedFileCache -> () | |||||
Generic LoadedFileCache Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
from :: LoadedFileCache -> Rep LoadedFileCache x to :: Rep LoadedFileCache x -> LoadedFileCache | |||||
type Rep LoadedFileCache Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep LoadedFileCache = D1 ('MetaData "LoadedFileCache" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "LoadedFileCache" 'PrefixI 'True) (S1 ('MetaSel ('Just "lfcCached") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 CachedTypeCheckLog) :*: S1 ('MetaSel ('Just "lfcCurrent") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 CurrentTypeCheckLog))) |
type 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).
initPersistentState :: PersistentTCState Source #
Empty persistent state.
initialMetaId :: MetaId Source #
An initial MetaId
.
initPreScopeState :: PreScopeState Source #
Empty state of type checker.
stImportedModules :: Lens' TCState (Set TopLevelModuleName) Source #
stForeignCode :: Lens' TCState (Map BackendName ForeignCodeStack) Source #
getUserWarnings :: ReadTCState m => m (Map QName Text) Source #
getPartialDefs :: ReadTCState m => m (Set QName) Source #
stBackends :: Lens' TCState [Backend_boot TCM] Source #
stProjectConfigs :: Lens' TCState (Map FilePath ProjectConfig) Source #
stAgdaLibFiles :: Lens' TCState (Map FilePath AgdaLibFile) Source #
stOpaqueBlocks :: Lens' TCState (Map OpaqueId OpaqueBlock) Source #
stModuleCheckpoints :: Lens' TCState (Map ModuleName CheckpointId) Source #
stCurrentModule :: Lens' TCState (Maybe (ModuleName, TopLevelModuleName)) Source #
Note that the lens is "strict".
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.
stMutualBlocks :: Lens' TCState (Map MutualId MutualBlock) Source #
stFreshInt :: Lens' TCState Int Source #
stAreWeCaching :: Lens' TCState Bool Source #
stPostponeInstanceSearch :: Lens' TCState Bool Source #
stConsideringInstance :: Lens' TCState Bool Source #
stInstantiateBlocking :: Lens' TCState Bool Source #
unionBuiltin :: Builtin a -> Builtin a -> Builtin a Source #
Union two Builtin
s. Only defined for BuiltinRewriteRelations
.
class Enum i => HasFresh i where Source #
Instances
HasFresh InteractionId Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
HasFresh MetaId Source # | |
HasFresh NameId Source # | |
HasFresh OpaqueId Source # | |
HasFresh ProblemId Source # | |
HasFresh CheckpointId Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
HasFresh MutualId Source # | |
HasFresh Int Source # | |
Defined in Agda.TypeChecking.Monad.Base |
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 Range Source # | |
Defined in Agda.TypeChecking.Monad.Base freshName_ :: MonadFresh NameId m => Range -> m Name Source # | |
FreshName String Source # | |
Defined in Agda.TypeChecking.Monad.Base freshName_ :: MonadFresh NameId m => String -> m Name Source # | |
FreshName () Source # | |
Defined in Agda.TypeChecking.Monad.Base freshName_ :: MonadFresh NameId m => () -> m Name Source # | |
FreshName (Range, String) Source # | |
Defined in Agda.TypeChecking.Monad.Base freshName_ :: MonadFresh NameId m => (Range, String) -> m Name Source # |
class Monad m => MonadStConcreteNames (m :: Type -> Type) where Source #
A monad that has read and write access to the stConcreteNames part of the TCState. Basically, this is a synonym for `MonadState ConcreteNames m` (which cannot be used directly because of the limitations of Haskell's typeclass system).
runStConcreteNames :: StateT ConcreteNames m a -> m a Source #
useConcreteNames :: m ConcreteNames Source #
modifyConcreteNames :: (ConcreteNames -> ConcreteNames) -> m () Source #
Instances
MonadStConcreteNames AbsToCon Source # | |
Defined in Agda.Syntax.Translation.AbstractToConcrete runStConcreteNames :: StateT ConcreteNames AbsToCon a -> AbsToCon a Source # useConcreteNames :: AbsToCon ConcreteNames Source # modifyConcreteNames :: (ConcreteNames -> ConcreteNames) -> AbsToCon () Source # | |
MonadStConcreteNames TCM Source # | |
Defined in Agda.TypeChecking.Monad.Base runStConcreteNames :: StateT ConcreteNames TCM a -> TCM a Source # useConcreteNames :: TCM ConcreteNames Source # modifyConcreteNames :: (ConcreteNames -> ConcreteNames) -> TCM () Source # | |
ReadTCState m => MonadStConcreteNames (PureConversionT m) Source # | |
Defined in Agda.TypeChecking.Conversion.Pure runStConcreteNames :: StateT ConcreteNames (PureConversionT m) a -> PureConversionT m a Source # useConcreteNames :: PureConversionT m ConcreteNames Source # modifyConcreteNames :: (ConcreteNames -> ConcreteNames) -> PureConversionT m () Source # | |
MonadStConcreteNames m => MonadStConcreteNames (IdentityT m) Source # | |
Defined in Agda.TypeChecking.Monad.Base runStConcreteNames :: StateT ConcreteNames (IdentityT m) a -> IdentityT m a Source # useConcreteNames :: IdentityT m ConcreteNames Source # modifyConcreteNames :: (ConcreteNames -> ConcreteNames) -> IdentityT m () Source # | |
MonadStConcreteNames m => MonadStConcreteNames (ReaderT r m) Source # | |
Defined in Agda.TypeChecking.Monad.Base runStConcreteNames :: StateT ConcreteNames (ReaderT r m) a -> ReaderT r m a Source # useConcreteNames :: ReaderT r m ConcreteNames Source # modifyConcreteNames :: (ConcreteNames -> ConcreteNames) -> ReaderT r m () Source # | |
MonadStConcreteNames m => MonadStConcreteNames (StateT s m) Source # | |
Defined in Agda.TypeChecking.Monad.Base runStConcreteNames :: StateT ConcreteNames (StateT s m) a -> StateT s m a Source # useConcreteNames :: StateT s m ConcreteNames Source # modifyConcreteNames :: (ConcreteNames -> ConcreteNames) -> StateT s m () Source # |
stateTCLensM :: MonadTCState m => Lens' TCState a -> (a -> m (r, a)) -> m r Source #
Modify a part of the state monadically, and return some result.
data ModuleCheckMode Source #
Distinguishes between type-checked and scope-checked interfaces
when stored in the map of VisitedModules
.
Instances
NFData ModuleCheckMode Source # | |||||
Defined in Agda.TypeChecking.Monad.Base rnf :: ModuleCheckMode -> () | |||||
Bounded ModuleCheckMode Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Enum ModuleCheckMode Source # | |||||
Defined in Agda.TypeChecking.Monad.Base succ :: ModuleCheckMode -> ModuleCheckMode pred :: ModuleCheckMode -> ModuleCheckMode toEnum :: Int -> ModuleCheckMode fromEnum :: ModuleCheckMode -> Int enumFrom :: ModuleCheckMode -> [ModuleCheckMode] enumFromThen :: ModuleCheckMode -> ModuleCheckMode -> [ModuleCheckMode] enumFromTo :: ModuleCheckMode -> ModuleCheckMode -> [ModuleCheckMode] enumFromThenTo :: ModuleCheckMode -> ModuleCheckMode -> ModuleCheckMode -> [ModuleCheckMode] | |||||
Generic ModuleCheckMode Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
from :: ModuleCheckMode -> Rep ModuleCheckMode x to :: Rep ModuleCheckMode x -> ModuleCheckMode | |||||
Show ModuleCheckMode Source # | |||||
Defined in Agda.TypeChecking.Monad.Base showsPrec :: Int -> ModuleCheckMode -> ShowS show :: ModuleCheckMode -> String showList :: [ModuleCheckMode] -> ShowS | |||||
Eq ModuleCheckMode Source # | |||||
Defined in Agda.TypeChecking.Monad.Base (==) :: ModuleCheckMode -> ModuleCheckMode -> Bool (/=) :: ModuleCheckMode -> ModuleCheckMode -> Bool | |||||
Ord ModuleCheckMode Source # | |||||
Defined in Agda.TypeChecking.Monad.Base compare :: ModuleCheckMode -> ModuleCheckMode -> Ordering (<) :: ModuleCheckMode -> ModuleCheckMode -> Bool (<=) :: ModuleCheckMode -> ModuleCheckMode -> Bool (>) :: ModuleCheckMode -> ModuleCheckMode -> Bool (>=) :: ModuleCheckMode -> ModuleCheckMode -> Bool max :: ModuleCheckMode -> ModuleCheckMode -> ModuleCheckMode min :: ModuleCheckMode -> ModuleCheckMode -> ModuleCheckMode | |||||
type Rep ModuleCheckMode Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep ModuleCheckMode = D1 ('MetaData "ModuleCheckMode" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "ModuleScopeChecked" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "ModuleTypeChecked" 'PrefixI 'False) (U1 :: Type -> Type)) |
iFullHash :: Interface -> Hash Source #
Combines the source hash and the (full) hashes of the imported modules.
intSignature :: Lens' Interface Signature Source #
A lens for the iSignature
field of the Interface
type.
class LensClosure b a | b -> a where Source #
lensClosure :: Lens' b (Closure a) Source #
Instances
LensClosure MetaInfo Range Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
LensClosure MetaVariable Range Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
LensClosure (Closure a) a Source # | |
Defined in Agda.TypeChecking.Monad.Base |
buildClosure :: (MonadTCEnv m, ReadTCState m) => a -> m (Closure a) Source #
data WhyCheckModality Source #
Why are we performing a modality check?
ConstructorType | Because --without-K is enabled, so the types of data constructors must be usable at the context's modality. |
IndexedClause | Because --without-K is enabled, so the result type of clauses must be usable at the context's modality. |
IndexedClauseArg Name Name | Because --without-K is enabled, so any argument (second name) which mentions a dotted argument (first name) must have a type which is usable at the context's modality. |
GeneratedClause | Because we double-check the --cubical-compatible clauses. This is an internal error! |
Instances
NFData WhyCheckModality Source # | |||||
Defined in Agda.TypeChecking.Monad.Base rnf :: WhyCheckModality -> () | |||||
Generic WhyCheckModality Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
from :: WhyCheckModality -> Rep WhyCheckModality x to :: Rep WhyCheckModality x -> WhyCheckModality | |||||
Show WhyCheckModality Source # | |||||
Defined in Agda.TypeChecking.Monad.Base showsPrec :: Int -> WhyCheckModality -> ShowS show :: WhyCheckModality -> String showList :: [WhyCheckModality] -> ShowS | |||||
type Rep WhyCheckModality Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep WhyCheckModality = D1 ('MetaData "WhyCheckModality" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) ((C1 ('MetaCons "ConstructorType" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "IndexedClause" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "IndexedClauseArg" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Name) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Name)) :+: C1 ('MetaCons "GeneratedClause" 'PrefixI 'False) (U1 :: Type -> Type))) |
Information about whether an argument is forced by the type of a function.
Instances
KillRange IsForced Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
PrettyTCM IsForced Source # | |||||
Defined in Agda.TypeChecking.Pretty | |||||
ChooseFlex IsForced Source # | |||||
Defined in Agda.TypeChecking.Rules.LHS.Problem chooseFlex :: IsForced -> IsForced -> FlexChoice Source # | |||||
EmbPrj IsForced Source # | |||||
NFData IsForced Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Generic IsForced Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
| |||||
Show IsForced Source # | |||||
Eq IsForced Source # | |||||
type Rep IsForced Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep IsForced = D1 ('MetaData "IsForced" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "Forced" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "NotForced" 'PrefixI 'False) (U1 :: Type -> Type)) |
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
HasOverlapMode Candidate Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Free Candidate Source # | |||||
PrettyTCM Candidate Source # | |||||
Defined in Agda.TypeChecking.Pretty | |||||
Instantiate Candidate Source # | |||||
Defined in Agda.TypeChecking.Reduce | |||||
InstantiateFull Candidate Source # | |||||
Defined in Agda.TypeChecking.Reduce | |||||
Normalise Candidate Source # | |||||
Defined in Agda.TypeChecking.Reduce | |||||
Reduce Candidate Source # | |||||
Simplify Candidate Source # | |||||
Defined in Agda.TypeChecking.Reduce | |||||
Subst Candidate Source # | |||||
Defined in Agda.TypeChecking.Substitute applySubst :: Substitution' (SubstArg Candidate) -> Candidate -> Candidate Source # | |||||
NFData Candidate Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Generic Candidate Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
| |||||
Show Candidate Source # | |||||
Eq Candidate | |||||
Ord Candidate | |||||
Defined in Agda.TypeChecking.Substitute | |||||
type SubstArg Candidate Source # | |||||
Defined in Agda.TypeChecking.Substitute | |||||
type Rep Candidate Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep Candidate = D1 ('MetaData "Candidate" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "Candidate" 'PrefixI 'True) ((S1 ('MetaSel ('Just "candidateKind") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 CandidateKind) :*: S1 ('MetaSel ('Just "candidateTerm") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)) :*: (S1 ('MetaSel ('Just "candidateType") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type) :*: S1 ('MetaSel ('Just "candidateOverlap") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 OverlapMode)))) |
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
Parametrized since it is used without MetaId when creating a new meta.
HasType | |
| |
IsSort | |
Instances
Pretty a => Pretty (Judgement a) Source # | |||||
PrettyTCM a => PrettyTCM (Judgement a) Source # | |||||
Defined in Agda.TypeChecking.Pretty | |||||
InstantiateFull (Judgement MetaId) Source # | |||||
Defined in Agda.TypeChecking.Reduce | |||||
EmbPrj a => EmbPrj (Judgement a) Source # | |||||
NFData a => NFData (Judgement a) Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Generic (Judgement a) Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
| |||||
Show a => Show (Judgement a) Source # | |||||
type Rep (Judgement a) Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep (Judgement a) = D1 ('MetaData "Judgement" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "HasType" 'PrefixI 'True) (S1 ('MetaSel ('Just "jMetaId") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 a) :*: (S1 ('MetaSel ('Just "jComparison") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: S1 ('MetaSel ('Just "jMetaType") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type))) :+: C1 ('MetaCons "IsSort" 'PrefixI 'True) (S1 ('MetaSel ('Just "jMetaId") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 a) :*: S1 ('MetaSel ('Just "jMetaType") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type))) |
data DoGeneralize Source #
YesGeneralizeVar | Generalize because it is a generalizable variable. |
YesGeneralizeMeta | Generalize because it is a metavariable and we're currently checking the type of a generalizable variable (this should get the default modality). |
NoGeneralize | Don't generalize. |
Instances
KillRange DoGeneralize Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
EmbPrj DoGeneralize Source # | |||||
Defined in Agda.TypeChecking.Serialise.Instances.Internal icode :: DoGeneralize -> S Int32 Source # icod_ :: DoGeneralize -> S Int32 Source # value :: Int32 -> R DoGeneralize Source # | |||||
NFData DoGeneralize Source # | |||||
Defined in Agda.TypeChecking.Monad.Base rnf :: DoGeneralize -> () | |||||
Generic DoGeneralize Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
from :: DoGeneralize -> Rep DoGeneralize x to :: Rep DoGeneralize x -> DoGeneralize | |||||
Show DoGeneralize Source # | |||||
Defined in Agda.TypeChecking.Monad.Base showsPrec :: Int -> DoGeneralize -> ShowS show :: DoGeneralize -> String showList :: [DoGeneralize] -> ShowS | |||||
Eq DoGeneralize Source # | |||||
Defined in Agda.TypeChecking.Monad.Base (==) :: DoGeneralize -> DoGeneralize -> Bool (/=) :: DoGeneralize -> DoGeneralize -> Bool | |||||
Ord DoGeneralize Source # | |||||
Defined in Agda.TypeChecking.Monad.Base compare :: DoGeneralize -> DoGeneralize -> Ordering (<) :: DoGeneralize -> DoGeneralize -> Bool (<=) :: DoGeneralize -> DoGeneralize -> Bool (>) :: DoGeneralize -> DoGeneralize -> Bool (>=) :: DoGeneralize -> DoGeneralize -> Bool max :: DoGeneralize -> DoGeneralize -> DoGeneralize min :: DoGeneralize -> DoGeneralize -> DoGeneralize | |||||
type Rep DoGeneralize Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep DoGeneralize = D1 ('MetaData "DoGeneralize" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "YesGeneralizeVar" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "YesGeneralizeMeta" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "NoGeneralize" 'PrefixI 'False) (U1 :: Type -> Type))) |
data GeneralizedValue Source #
The value of a generalizable variable. This is created to be a generalizable meta before checking the type to be generalized.
Instances
NFData GeneralizedValue Source # | |||||
Defined in Agda.TypeChecking.Monad.Base rnf :: GeneralizedValue -> () | |||||
Generic GeneralizedValue Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
from :: GeneralizedValue -> Rep GeneralizedValue x to :: Rep GeneralizedValue x -> GeneralizedValue | |||||
Show GeneralizedValue Source # | |||||
Defined in Agda.TypeChecking.Monad.Base showsPrec :: Int -> GeneralizedValue -> ShowS show :: GeneralizedValue -> String showList :: [GeneralizedValue] -> ShowS | |||||
type Rep GeneralizedValue Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep GeneralizedValue = D1 ('MetaData "GeneralizedValue" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "GeneralizedValue" 'PrefixI 'True) (S1 ('MetaSel ('Just "genvalCheckpoint") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 CheckpointId) :*: (S1 ('MetaSel ('Just "genvalTerm") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Just "genvalType") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)))) |
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.
MetaPriority Int |
Instances
NFData MetaPriority Source # | |
Defined in Agda.TypeChecking.Monad.Base rnf :: MetaPriority -> () | |
Show MetaPriority Source # | |
Defined in Agda.TypeChecking.Monad.Base showsPrec :: Int -> MetaPriority -> ShowS show :: MetaPriority -> String showList :: [MetaPriority] -> ShowS | |
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 |
data MetaInstantiation Source #
Solution status of meta.
InstV Instantiation | Solved by |
OpenMeta MetaKind | Unsolved (open to solutions). |
BlockedConst Term | Solved, but solution blocked by unsolved constraints. |
PostponedTypeCheckingProblem (Closure TypeCheckingProblem) | Meta stands for value of the expression that is still to be type checked. |
Instances
Pretty MetaInstantiation Source # | |||||
Defined in Agda.TypeChecking.Monad.Base pretty :: MetaInstantiation -> Doc Source # prettyPrec :: Int -> MetaInstantiation -> Doc Source # prettyList :: [MetaInstantiation] -> Doc Source # | |||||
NFData MetaInstantiation Source # | |||||
Defined in Agda.TypeChecking.Monad.Base rnf :: MetaInstantiation -> () | |||||
Generic MetaInstantiation Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
from :: MetaInstantiation -> Rep MetaInstantiation x to :: Rep MetaInstantiation x -> MetaInstantiation | |||||
type Rep MetaInstantiation Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep MetaInstantiation = D1 ('MetaData "MetaInstantiation" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) ((C1 ('MetaCons "InstV" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Instantiation)) :+: C1 ('MetaCons "OpenMeta" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 MetaKind))) :+: (C1 ('MetaCons "BlockedConst" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)) :+: C1 ('MetaCons "PostponedTypeCheckingProblem" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Closure TypeCheckingProblem))))) |
Instances
NFData Listener Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Generic Listener Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
| |||||
Eq Listener Source # | |||||
Ord Listener Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
type Rep Listener Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep Listener = D1 ('MetaData "Listener" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "EtaExpand" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 MetaId)) :+: C1 ('MetaCons "CheckConstraint" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Nat) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ProblemConstraint))) |
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 |
Instances
NFData Frozen Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Generic Frozen Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
| |||||
Show Frozen Source # | |||||
Eq Frozen Source # | |||||
type Rep Frozen Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep Frozen = D1 ('MetaData "Frozen" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "Frozen" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "Instantiable" 'PrefixI 'False) (U1 :: Type -> Type)) |
data Instantiation Source #
Meta-variable instantiations.
Instances
InstantiateFull Instantiation Source # | |||||
Defined in Agda.TypeChecking.Reduce | |||||
EmbPrj Instantiation Source # | |||||
Defined in Agda.TypeChecking.Serialise.Instances.Internal icode :: Instantiation -> S Int32 Source # icod_ :: Instantiation -> S Int32 Source # value :: Int32 -> R Instantiation Source # | |||||
NFData Instantiation Source # | |||||
Defined in Agda.TypeChecking.Monad.Base rnf :: Instantiation -> () | |||||
Generic Instantiation Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
from :: Instantiation -> Rep Instantiation x to :: Rep Instantiation x -> Instantiation | |||||
Show Instantiation Source # | |||||
Defined in Agda.TypeChecking.Monad.Base showsPrec :: Int -> Instantiation -> ShowS show :: Instantiation -> String showList :: [Instantiation] -> ShowS | |||||
type Rep Instantiation Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep Instantiation = D1 ('MetaData "Instantiation" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "Instantiation" 'PrefixI 'True) (S1 ('MetaSel ('Just "instTel") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Arg String]) :*: S1 ('MetaSel ('Just "instBody") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term))) |
data TypeCheckingProblem Source #
CheckExpr Comparison Expr Type | |
CheckArgs Comparison ExpandHidden Range [NamedArg Expr] Type Type (ArgsCheckState CheckedTarget -> TCM Term) | |
CheckProjAppToKnownPrincipalArg Comparison Expr ProjOrigin (List1 QName) Args Type Int Term Type PrincipalArgTypeMetas | |
CheckLambda Comparison (Arg (List1 (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 # | |||||
NFData TypeCheckingProblem Source # | |||||
Defined in Agda.TypeChecking.Monad.Base rnf :: TypeCheckingProblem -> () | |||||
Generic TypeCheckingProblem Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
from :: TypeCheckingProblem -> Rep TypeCheckingProblem x to :: Rep TypeCheckingProblem x -> TypeCheckingProblem | |||||
type Rep TypeCheckingProblem Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep TypeCheckingProblem = D1 ('MetaData "TypeCheckingProblem" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) ((C1 ('MetaCons "CheckExpr" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type))) :+: C1 ('MetaCons "CheckArgs" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ExpandHidden) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range))) :*: ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [NamedArg Expr]) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (ArgsCheckState CheckedTarget -> TCM Term)))))) :+: (C1 ('MetaCons "CheckProjAppToKnownPrincipalArg" 'PrefixI 'False) (((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ProjOrigin) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 QName)) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Args)))) :*: ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Int)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 PrincipalArgTypeMetas))))) :+: (C1 ('MetaCons "CheckLambda" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Arg (List1 (WithHiding Name), Maybe Type)))) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type))) :+: C1 ('MetaCons "DoQuoteTerm" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)))))) |
data RemoteMetaVariable Source #
Information about remote meta-variables.
Remote meta-variables are meta-variables originating in other modules. These meta-variables are always instantiated. We do not retain all the information about a local meta-variable when creating an interface:
- The
mvPriority
field is not needed, because the meta-variable cannot be instantiated. - The
mvFrozen
field is not needed, because there is no point in freezing instantiated meta-variables. - The
mvListeners
field is not needed, because no meta-variable should be listening to this one. - The
mvTwin
field is not needed, because the meta-variable has already been instantiated. - The
mvPermutation
is currently removed, but could be retained if it turns out to be useful for something. - The only part of the
mvInfo
field that is kept is themiModality
field. ThemiMetaOccursCheck
andmiGeneralizable
fields are omitted, because the meta-variable has already been instantiated. TheRange
that is part of themiClosRange
field and themiNameSuggestion
field are omitted because instantiated meta-variables are typically not presented to users. Finally theClosure
part of themiClosRange
field is omitted because it can be large (at least if we ignore potential sharing).
Instances
LensModality RemoteMetaVariable Source # | |||||
Defined in Agda.TypeChecking.Monad.Base getModality :: RemoteMetaVariable -> Modality Source # setModality :: Modality -> RemoteMetaVariable -> RemoteMetaVariable Source # mapModality :: (Modality -> Modality) -> RemoteMetaVariable -> RemoteMetaVariable Source # | |||||
LensQuantity RemoteMetaVariable Source # | |||||
Defined in Agda.TypeChecking.Monad.Base getQuantity :: RemoteMetaVariable -> Quantity Source # setQuantity :: Quantity -> RemoteMetaVariable -> RemoteMetaVariable Source # mapQuantity :: (Quantity -> Quantity) -> RemoteMetaVariable -> RemoteMetaVariable Source # | |||||
LensRelevance RemoteMetaVariable Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
InstantiateFull RemoteMetaVariable Source # | |||||
Defined in Agda.TypeChecking.Reduce | |||||
EmbPrj RemoteMetaVariable Source # | |||||
Defined in Agda.TypeChecking.Serialise.Instances.Internal icode :: RemoteMetaVariable -> S Int32 Source # icod_ :: RemoteMetaVariable -> S Int32 Source # value :: Int32 -> R RemoteMetaVariable Source # | |||||
NFData RemoteMetaVariable Source # | |||||
Defined in Agda.TypeChecking.Monad.Base rnf :: RemoteMetaVariable -> () | |||||
Generic RemoteMetaVariable Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
from :: RemoteMetaVariable -> Rep RemoteMetaVariable x to :: Rep RemoteMetaVariable x -> RemoteMetaVariable | |||||
Show RemoteMetaVariable Source # | |||||
Defined in Agda.TypeChecking.Monad.Base showsPrec :: Int -> RemoteMetaVariable -> ShowS show :: RemoteMetaVariable -> String showList :: [RemoteMetaVariable] -> ShowS | |||||
type Rep RemoteMetaVariable Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep RemoteMetaVariable = D1 ('MetaData "RemoteMetaVariable" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "RemoteMetaVariable" 'PrefixI 'True) (S1 ('MetaSel ('Just "rmvInstantiation") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Instantiation) :*: (S1 ('MetaSel ('Just "rmvModality") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Modality) :*: S1 ('MetaSel ('Just "rmvJudgement") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Judgement MetaId))))) |
data 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 PrincipalArgTypeMetas Source #
PrincipalArgTypeMetas | |
|
Instances
NFData PrincipalArgTypeMetas Source # | |||||
Defined in Agda.TypeChecking.Monad.Base rnf :: PrincipalArgTypeMetas -> () | |||||
Generic PrincipalArgTypeMetas Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
from :: PrincipalArgTypeMetas -> Rep PrincipalArgTypeMetas x to :: Rep PrincipalArgTypeMetas x -> PrincipalArgTypeMetas | |||||
type Rep PrincipalArgTypeMetas Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep PrincipalArgTypeMetas = D1 ('MetaData "PrincipalArgTypeMetas" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "PrincipalArgTypeMetas" 'PrefixI 'True) (S1 ('MetaSel ('Just "patmMetas") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Args) :*: S1 ('MetaSel ('Just "patmRemainder") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type))) |
data ExpandHidden Source #
ExpandLast | Add implicit arguments in the end until type is no longer hidden |
DontExpandLast | Do not append implicit arguments. |
ReallyDontExpandLast | Makes |
Instances
NFData ExpandHidden Source # | |||||
Defined in Agda.TypeChecking.Monad.Base rnf :: ExpandHidden -> () | |||||
Generic ExpandHidden Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
from :: ExpandHidden -> Rep ExpandHidden x to :: Rep ExpandHidden x -> ExpandHidden | |||||
Eq ExpandHidden Source # | |||||
Defined in Agda.TypeChecking.Monad.Base (==) :: ExpandHidden -> ExpandHidden -> Bool (/=) :: ExpandHidden -> ExpandHidden -> Bool | |||||
type Rep ExpandHidden Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep ExpandHidden = D1 ('MetaData "ExpandHidden" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "ExpandLast" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "DontExpandLast" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "ReallyDontExpandLast" 'PrefixI 'False) (U1 :: Type -> Type))) |
data ArgsCheckState a Source #
ACState | |
|
Instances
Show a => Show (ArgsCheckState a) Source # | |
Defined in Agda.TypeChecking.Monad.Base showsPrec :: Int -> ArgsCheckState a -> ShowS show :: ArgsCheckState a -> String showList :: [ArgsCheckState a] -> ShowS |
suffixNameSuggestion :: MetaNameSuggestion -> ArgName -> MetaNameSuggestion Source #
Append an ArgName
to a MetaNameSuggestion
, for computing the
name suggestions of eta-expansion metas. If the MetaNameSuggestion
is empty or an underscore, the field name is taken as the suggestion.
getMetaInfo :: MetaVariable -> Closure Range Source #
getMetaScope :: MetaVariable -> ScopeInfo Source #
getMetaEnv :: MetaVariable -> TCEnv Source #
getMetaSig :: MetaVariable -> Signature Source #
aModeToDef :: AbstractMode -> Maybe IsAbstract Source #
aDefToMode :: IsAbstract -> AbstractMode Source #
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.
InteractionPoint | |
|
Instances
HasTag InteractionPoint Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
tag :: InteractionPoint -> Maybe (Tag InteractionPoint) Source # | |||||
NFData InteractionPoint Source # | |||||
Defined in Agda.TypeChecking.Monad.Base rnf :: InteractionPoint -> () | |||||
NFData InteractionPoints Source # | |||||
Defined in Agda.TypeChecking.Monad.Base rnf :: InteractionPoints -> () | |||||
Generic InteractionPoint Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
from :: InteractionPoint -> Rep InteractionPoint x to :: Rep InteractionPoint x -> InteractionPoint | |||||
Eq InteractionPoint Source # | |||||
Defined in Agda.TypeChecking.Monad.Base (==) :: InteractionPoint -> InteractionPoint -> Bool (/=) :: InteractionPoint -> InteractionPoint -> Bool | |||||
type Tag InteractionPoint Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
type Rep InteractionPoint Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep InteractionPoint = D1 ('MetaData "InteractionPoint" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "InteractionPoint" 'PrefixI 'True) ((S1 ('MetaSel ('Just "ipRange") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range) :*: S1 ('MetaSel ('Just "ipMeta") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe MetaId))) :*: (S1 ('MetaSel ('Just "ipSolved") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool) :*: (S1 ('MetaSel ('Just "ipClause") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 IPClause) :*: S1 ('MetaSel ('Just "ipBoundary") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 IPBoundary))))) |
Which clause is an interaction point located in?
IPClause | |
| |
IPNoClause | The interaction point is not in the rhs of a clause. |
Instances
NFData IPClause Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Generic IPClause Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
| |||||
Eq IPClause Source # | |||||
type Rep IPClause Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep IPClause = D1 ('MetaData "IPClause" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "IPClause" 'PrefixI 'True) ((S1 ('MetaSel ('Just "ipcQName") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: (S1 ('MetaSel ('Just "ipcClauseNo") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Int) :*: S1 ('MetaSel ('Just "ipcType") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type))) :*: (S1 ('MetaSel ('Just "ipcWithSub") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe Substitution)) :*: (S1 ('MetaSel ('Just "ipcClause") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 SpineClause) :*: S1 ('MetaSel ('Just "ipcClosure") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Closure ()))))) :+: C1 ('MetaCons "IPNoClause" 'PrefixI 'False) (U1 :: Type -> Type)) |
type IPBoundary = IPBoundary' Term Source #
newtype IPBoundary' t Source #
IPBoundary | |
|
Instances
Functor IPBoundary' Source # | |||||
Defined in Agda.TypeChecking.Monad.Base fmap :: (a -> b) -> IPBoundary' a -> IPBoundary' b (<$) :: a -> IPBoundary' b -> IPBoundary' a # | |||||
Foldable IPBoundary' Source # | |||||
Defined in Agda.TypeChecking.Monad.Base fold :: Monoid m => IPBoundary' m -> m foldMap :: Monoid m => (a -> m) -> IPBoundary' a -> m foldMap' :: Monoid m => (a -> m) -> IPBoundary' a -> m foldr :: (a -> b -> b) -> b -> IPBoundary' a -> b foldr' :: (a -> b -> b) -> b -> IPBoundary' a -> b foldl :: (b -> a -> b) -> b -> IPBoundary' a -> b foldl' :: (b -> a -> b) -> b -> IPBoundary' a -> b foldr1 :: (a -> a -> a) -> IPBoundary' a -> a foldl1 :: (a -> a -> a) -> IPBoundary' a -> a toList :: IPBoundary' a -> [a] null :: IPBoundary' a -> Bool length :: IPBoundary' a -> Int elem :: Eq a => a -> IPBoundary' a -> Bool maximum :: Ord a => IPBoundary' a -> a minimum :: Ord a => IPBoundary' a -> a sum :: Num a => IPBoundary' a -> a product :: Num a => IPBoundary' a -> a | |||||
Traversable IPBoundary' Source # | |||||
Defined in Agda.TypeChecking.Monad.Base traverse :: Applicative f => (a -> f b) -> IPBoundary' a -> f (IPBoundary' b) sequenceA :: Applicative f => IPBoundary' (f a) -> f (IPBoundary' a) mapM :: Monad m => (a -> m b) -> IPBoundary' a -> m (IPBoundary' b) sequence :: Monad m => IPBoundary' (m a) -> m (IPBoundary' a) | |||||
ToConcrete a => ToConcrete (IPBoundary' a) Source # | |||||
Defined in Agda.Interaction.BasicOps
toConcrete :: IPBoundary' a -> AbsToCon (ConOfAbs (IPBoundary' a)) Source # bindToConcrete :: IPBoundary' a -> (ConOfAbs (IPBoundary' a) -> AbsToCon b) -> AbsToCon b Source # | |||||
Reify a => Reify (IPBoundary' a) Source # | |||||
Defined in Agda.Interaction.BasicOps
reify :: MonadReify m => IPBoundary' a -> m (ReifiesTo (IPBoundary' a)) Source # reifyWhen :: MonadReify m => Bool -> IPBoundary' a -> m (ReifiesTo (IPBoundary' a)) Source # | |||||
Instantiate t => Instantiate (IPBoundary' t) Source # | |||||
Defined in Agda.TypeChecking.Reduce instantiate' :: IPBoundary' t -> ReduceM (IPBoundary' t) Source # | |||||
InstantiateFull t => InstantiateFull (IPBoundary' t) Source # | |||||
Defined in Agda.TypeChecking.Reduce instantiateFull' :: IPBoundary' t -> ReduceM (IPBoundary' t) Source # | |||||
Normalise t => Normalise (IPBoundary' t) Source # | |||||
Defined in Agda.TypeChecking.Reduce normalise' :: IPBoundary' t -> ReduceM (IPBoundary' t) Source # | |||||
Reduce t => Reduce (IPBoundary' t) Source # | |||||
Defined in Agda.TypeChecking.Reduce reduce' :: IPBoundary' t -> ReduceM (IPBoundary' t) Source # reduceB' :: IPBoundary' t -> ReduceM (Blocked (IPBoundary' t)) Source # | |||||
Simplify t => Simplify (IPBoundary' t) Source # | |||||
Defined in Agda.TypeChecking.Reduce simplify' :: IPBoundary' t -> ReduceM (IPBoundary' t) | |||||
NFData t => NFData (IPBoundary' t) Source # | |||||
Defined in Agda.TypeChecking.Monad.Base rnf :: IPBoundary' t -> () | |||||
Generic (IPBoundary' t) Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
from :: IPBoundary' t -> Rep (IPBoundary' t) x to :: Rep (IPBoundary' t) x -> IPBoundary' t | |||||
Show t => Show (IPBoundary' t) Source # | |||||
Defined in Agda.TypeChecking.Monad.Base showsPrec :: Int -> IPBoundary' t -> ShowS show :: IPBoundary' t -> String showList :: [IPBoundary' t] -> ShowS | |||||
type ConOfAbs (IPBoundary' a) Source # | |||||
Defined in Agda.Interaction.BasicOps | |||||
type ReifiesTo (IPBoundary' a) Source # | |||||
Defined in Agda.Interaction.BasicOps | |||||
type Rep (IPBoundary' t) Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep (IPBoundary' t) = D1 ('MetaData "IPBoundary'" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'True) (C1 ('MetaCons "IPBoundary" 'PrefixI 'True) (S1 ('MetaSel ('Just "getBoundary") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Map (IntMap Bool) t)))) |
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
NFData Overapplied Source # | |||||
Defined in Agda.TypeChecking.Monad.Base rnf :: Overapplied -> () | |||||
Generic Overapplied Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
from :: Overapplied -> Rep Overapplied x to :: Rep Overapplied x -> Overapplied | |||||
Show Overapplied Source # | |||||
Defined in Agda.TypeChecking.Monad.Base showsPrec :: Int -> Overapplied -> ShowS show :: Overapplied -> String showList :: [Overapplied] -> ShowS | |||||
Eq Overapplied Source # | |||||
Defined in Agda.TypeChecking.Monad.Base (==) :: Overapplied -> Overapplied -> Bool (/=) :: Overapplied -> Overapplied -> Bool | |||||
type Rep Overapplied Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep Overapplied = D1 ('MetaData "Overapplied" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "Overapplied" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "NotOverapplied" 'PrefixI 'False) (U1 :: Type -> Type)) |
type Sections = Map ModuleName Section Source #
type RewriteRuleMap = HashMap QName RewriteRules Source #
data InstanceTable Source #
Records information about the instances in the signature. Does not deal with local instances.
InstanceTable | |
|
Instances
KillRange InstanceTable Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
EmbPrj InstanceTable Source # | |||||
Defined in Agda.TypeChecking.Serialise.Instances.Internal icode :: InstanceTable -> S Int32 Source # icod_ :: InstanceTable -> S Int32 Source # value :: Int32 -> R InstanceTable Source # | |||||
NFData InstanceTable Source # | |||||
Defined in Agda.TypeChecking.Monad.Base rnf :: InstanceTable -> () | |||||
Monoid InstanceTable Source # | |||||
Defined in Agda.TypeChecking.Monad.Base mappend :: InstanceTable -> InstanceTable -> InstanceTable mconcat :: [InstanceTable] -> InstanceTable | |||||
Semigroup InstanceTable Source # | |||||
Defined in Agda.TypeChecking.Monad.Base (<>) :: InstanceTable -> InstanceTable -> InstanceTable # sconcat :: NonEmpty InstanceTable -> InstanceTable stimes :: Integral b => b -> InstanceTable -> InstanceTable | |||||
Generic InstanceTable Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
from :: InstanceTable -> Rep InstanceTable x to :: Rep InstanceTable x -> InstanceTable | |||||
Show InstanceTable Source # | |||||
Defined in Agda.TypeChecking.Monad.Base showsPrec :: Int -> InstanceTable -> ShowS show :: InstanceTable -> String showList :: [InstanceTable] -> ShowS | |||||
type Rep InstanceTable Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep InstanceTable = D1 ('MetaData "InstanceTable" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "InstanceTable" 'PrefixI 'True) (S1 ('MetaSel ('Just "_itableTree") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (DiscrimTree QName)) :*: S1 ('MetaSel ('Just "_itableCounts") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Map QName Int)))) |
type LocalDisplayForm = Open DisplayForm Source #
data DisplayTerm Source #
DWithApp DisplayTerm [DisplayTerm] Elims |
|
DCon ConHead ConInfo [Arg DisplayTerm] |
|
DDef QName [Elim' DisplayTerm] |
|
DDot' Term Elims |
|
DTerm' Term Elims |
|
Instances
Pretty DisplayTerm Source # | |||||
Defined in Agda.TypeChecking.Monad.Base pretty :: DisplayTerm -> Doc Source # prettyPrec :: Int -> DisplayTerm -> Doc Source # prettyList :: [DisplayTerm] -> Doc Source # | |||||
NamesIn DisplayTerm Source # | |||||
Defined in Agda.Syntax.Internal.Names namesAndMetasIn' :: Monoid m => (Either QName MetaId -> m) -> DisplayTerm -> m Source # | |||||
KillRange DisplayTerm Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Reify DisplayTerm Source # | |||||
Defined in Agda.Syntax.Translation.InternalToAbstract
reify :: MonadReify m => DisplayTerm -> m (ReifiesTo DisplayTerm) Source # reifyWhen :: MonadReify m => Bool -> DisplayTerm -> m (ReifiesTo DisplayTerm) Source # | |||||
Free DisplayTerm Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
PrettyTCM DisplayTerm Source # | |||||
Defined in Agda.TypeChecking.Pretty prettyTCM :: MonadPretty m => DisplayTerm -> m Doc Source # | |||||
InstantiateFull DisplayTerm Source # | |||||
Defined in Agda.TypeChecking.Reduce | |||||
EmbPrj DisplayTerm Source # | |||||
Defined in Agda.TypeChecking.Serialise.Instances.Internal icode :: DisplayTerm -> S Int32 Source # icod_ :: DisplayTerm -> S Int32 Source # value :: Int32 -> R DisplayTerm Source # | |||||
Apply DisplayTerm Source # | |||||
Defined in Agda.TypeChecking.Substitute apply :: DisplayTerm -> Args -> DisplayTerm Source # applyE :: DisplayTerm -> Elims -> DisplayTerm Source # | |||||
Subst DisplayTerm Source # | |||||
Defined in Agda.TypeChecking.Substitute
applySubst :: Substitution' (SubstArg DisplayTerm) -> DisplayTerm -> DisplayTerm Source # | |||||
NFData DisplayTerm Source # | |||||
Defined in Agda.TypeChecking.Monad.Base rnf :: DisplayTerm -> () | |||||
Generic DisplayTerm Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
from :: DisplayTerm -> Rep DisplayTerm x to :: Rep DisplayTerm x -> DisplayTerm | |||||
Show DisplayTerm Source # | |||||
Defined in Agda.TypeChecking.Monad.Base showsPrec :: Int -> DisplayTerm -> ShowS show :: DisplayTerm -> String showList :: [DisplayTerm] -> ShowS | |||||
PrettyTCM (Elim' DisplayTerm) Source # | |||||
Defined in Agda.TypeChecking.Pretty prettyTCM :: MonadPretty m => Elim' DisplayTerm -> m Doc Source # | |||||
type ReifiesTo DisplayTerm Source # | |||||
Defined in Agda.Syntax.Translation.InternalToAbstract | |||||
type SubstArg DisplayTerm Source # | |||||
Defined in Agda.TypeChecking.Substitute | |||||
type Rep DisplayTerm Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep DisplayTerm = D1 ('MetaData "DisplayTerm" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) ((C1 ('MetaCons "DWithApp" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 DisplayTerm) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [DisplayTerm]) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Elims))) :+: C1 ('MetaCons "DCon" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ConHead) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ConInfo) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Arg DisplayTerm])))) :+: (C1 ('MetaCons "DDef" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Elim' DisplayTerm])) :+: (C1 ('MetaCons "DDot'" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Elims)) :+: C1 ('MetaCons "DTerm'" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Elims))))) |
pattern DDot :: Term -> DisplayTerm Source #
pattern DTerm :: Term -> DisplayTerm Source #
defaultDisplayForm :: QName -> [LocalDisplayForm] Source #
By default, we have no display form.
Instances
TermLike NLPType Source # | |||||
AllMetas NLPType Source # | |||||
NamesIn NLPType Source # | |||||
Defined in Agda.Syntax.Internal.Names | |||||
KillRange NLPType Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Free NLPType Source # | |||||
PrettyTCM NLPType Source # | |||||
Defined in Agda.TypeChecking.Pretty | |||||
InstantiateFull NLPType Source # | |||||
Defined in Agda.TypeChecking.Reduce | |||||
GetMatchables NLPType Source # | |||||
Defined in Agda.TypeChecking.Rewriting.NonLinPattern getMatchables :: NLPType -> [QName] Source # | |||||
NLPatVars NLPType Source # | |||||
Defined in Agda.TypeChecking.Rewriting.NonLinPattern | |||||
EmbPrj NLPType Source # | |||||
Subst NLPType Source # | |||||
Defined in Agda.TypeChecking.Substitute applySubst :: Substitution' (SubstArg NLPType) -> NLPType -> NLPType Source # | |||||
NFData NLPType Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Generic NLPType Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
| |||||
Show NLPType Source # | |||||
Match NLPType Type Source # | |||||
NLPatToTerm NLPType Type Source # | |||||
Defined in Agda.TypeChecking.Rewriting.NonLinPattern | |||||
PatternFrom Type NLPType Source # | |||||
Defined in Agda.TypeChecking.Rewriting.NonLinPattern | |||||
type SubstArg NLPType Source # | |||||
Defined in Agda.TypeChecking.Substitute | |||||
type Rep NLPType Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep NLPType = D1 ('MetaData "NLPType" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "NLPType" 'PrefixI 'True) (S1 ('MetaSel ('Just "nlpTypeSort") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 NLPSort) :*: S1 ('MetaSel ('Just "nlpTypeUnEl") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 NLPat))) |
PUniv Univ NLPat | |
PInf Univ Integer | |
PSizeUniv | |
PLockUniv | |
PLevelUniv | |
PIntervalUniv |
Instances
TermLike NLPSort Source # | |||||
AllMetas NLPSort Source # | |||||
NamesIn NLPSort Source # | |||||
Defined in Agda.Syntax.Internal.Names | |||||
KillRange NLPSort Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Free NLPSort Source # | |||||
PrettyTCM NLPSort Source # | |||||
Defined in Agda.TypeChecking.Pretty | |||||
InstantiateFull NLPSort Source # | |||||
Defined in Agda.TypeChecking.Reduce | |||||
GetMatchables NLPSort Source # | |||||
Defined in Agda.TypeChecking.Rewriting.NonLinPattern getMatchables :: NLPSort -> [QName] Source # | |||||
NLPatVars NLPSort Source # | |||||
Defined in Agda.TypeChecking.Rewriting.NonLinPattern | |||||
EmbPrj NLPSort Source # | |||||
Subst NLPSort Source # | |||||
Defined in Agda.TypeChecking.Substitute applySubst :: Substitution' (SubstArg NLPSort) -> NLPSort -> NLPSort Source # | |||||
NFData NLPSort Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Generic NLPSort Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
| |||||
Show NLPSort Source # | |||||
Match NLPSort Sort Source # | |||||
NLPatToTerm NLPSort Sort Source # | |||||
Defined in Agda.TypeChecking.Rewriting.NonLinPattern | |||||
PatternFrom Sort NLPSort Source # | |||||
Defined in Agda.TypeChecking.Rewriting.NonLinPattern | |||||
type SubstArg NLPSort Source # | |||||
Defined in Agda.TypeChecking.Substitute | |||||
type Rep NLPSort Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep NLPSort = D1 ('MetaData "NLPSort" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) ((C1 ('MetaCons "PUniv" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Univ) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 NLPat)) :+: (C1 ('MetaCons "PInf" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Univ) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Integer)) :+: C1 ('MetaCons "PSizeUniv" 'PrefixI 'False) (U1 :: Type -> Type))) :+: (C1 ('MetaCons "PLockUniv" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "PLevelUniv" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "PIntervalUniv" 'PrefixI 'False) (U1 :: Type -> Type)))) |
data RewriteRule Source #
Rewrite rules can be added independently from function clauses.
RewriteRule | |
|
Instances
NamesIn RewriteRule Source # | |||||
Defined in Agda.Syntax.Internal.Names namesAndMetasIn' :: Monoid m => (Either QName MetaId -> m) -> RewriteRule -> m Source # | |||||
KillRange RewriteRule Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
KillRange RewriteRuleMap Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
PrettyTCM RewriteRule Source # | |||||
Defined in Agda.TypeChecking.Pretty prettyTCM :: MonadPretty m => RewriteRule -> m Doc Source # | |||||
InstantiateFull RewriteRule Source # | |||||
Defined in Agda.TypeChecking.Reduce | |||||
GetMatchables RewriteRule Source # | |||||
Defined in Agda.TypeChecking.Rewriting.NonLinPattern getMatchables :: RewriteRule -> [QName] Source # | |||||
EmbPrj RewriteRule Source # | |||||
Defined in Agda.TypeChecking.Serialise.Instances.Internal icode :: RewriteRule -> S Int32 Source # icod_ :: RewriteRule -> S Int32 Source # value :: Int32 -> R RewriteRule Source # | |||||
Abstract RewriteRule Source # |
| ||||
Defined in Agda.TypeChecking.Substitute abstract :: Telescope -> RewriteRule -> RewriteRule Source # | |||||
Apply RewriteRule Source # | |||||
Defined in Agda.TypeChecking.Substitute apply :: RewriteRule -> Args -> RewriteRule Source # applyE :: RewriteRule -> Elims -> RewriteRule Source # | |||||
Subst RewriteRule Source # | |||||
Defined in Agda.TypeChecking.Substitute
applySubst :: Substitution' (SubstArg RewriteRule) -> RewriteRule -> RewriteRule Source # | |||||
NFData RewriteRule Source # | |||||
Defined in Agda.TypeChecking.Monad.Base rnf :: RewriteRule -> () | |||||
Generic RewriteRule Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
from :: RewriteRule -> Rep RewriteRule x to :: Rep RewriteRule x -> RewriteRule | |||||
Show RewriteRule Source # | |||||
Defined in Agda.TypeChecking.Monad.Base showsPrec :: Int -> RewriteRule -> ShowS show :: RewriteRule -> String showList :: [RewriteRule] -> ShowS | |||||
type SubstArg RewriteRule Source # | |||||
Defined in Agda.TypeChecking.Substitute | |||||
type Rep RewriteRule Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep RewriteRule = D1 ('MetaData "RewriteRule" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "RewriteRule" 'PrefixI 'True) ((S1 ('MetaSel ('Just "rewName") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: (S1 ('MetaSel ('Just "rewContext") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Telescope) :*: S1 ('MetaSel ('Just "rewHead") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName))) :*: ((S1 ('MetaSel ('Just "rewPats") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 PElims) :*: S1 ('MetaSel ('Just "rewRHS") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)) :*: (S1 ('MetaSel ('Just "rewType") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type) :*: S1 ('MetaSel ('Just "rewFromClause") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool))))) |
AxiomDefn AxiomData | Postulate. |
DataOrRecSigDefn DataOrRecSigData | Data or record type signature that doesn't yet have a definition. |
GeneralizableVar | Generalizable variable (introduced in |
AbstractDefn Defn | Returned by |
FunctionDefn FunctionData | |
DatatypeDefn DatatypeData | |
RecordDefn RecordData | |
ConstructorDefn ConstructorData | |
PrimitiveDefn PrimitiveData | Primitive or builtin functions. |
PrimitiveSortDefn PrimitiveSortData |
Instances
Pretty Defn Source # | |||||
NamesIn Defn Source # | |||||
Defined in Agda.Syntax.Internal.Names | |||||
KillRange Defn Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Occurs Defn Source # | |||||
InstantiateFull Defn Source # | |||||
Defined in Agda.TypeChecking.Reduce | |||||
EmbPrj Defn Source # | |||||
Abstract Defn Source # | |||||
Apply Defn Source # | |||||
NFData Defn Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Generic Defn Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
| |||||
Show Defn Source # | |||||
type Rep Defn Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep Defn = D1 ('MetaData "Defn" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (((C1 ('MetaCons "AxiomDefn" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 AxiomData)) :+: C1 ('MetaCons "DataOrRecSigDefn" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 DataOrRecSigData))) :+: (C1 ('MetaCons "GeneralizableVar" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "AbstractDefn" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Defn)) :+: C1 ('MetaCons "FunctionDefn" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 FunctionData))))) :+: ((C1 ('MetaCons "DatatypeDefn" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 DatatypeData)) :+: C1 ('MetaCons "RecordDefn" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 RecordData))) :+: (C1 ('MetaCons "ConstructorDefn" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ConstructorData)) :+: (C1 ('MetaCons "PrimitiveDefn" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 PrimitiveData)) :+: C1 ('MetaCons "PrimitiveSortDefn" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 PrimitiveSortData)))))) |
data NumGeneralizableArgs Source #
NoGeneralizableArgs | |
SomeGeneralizableArgs !Int | When lambda-lifting new args are generalizable if
|
Instances
KillRange NumGeneralizableArgs Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
EmbPrj NumGeneralizableArgs Source # | |
Defined in Agda.TypeChecking.Serialise.Instances.Internal icode :: NumGeneralizableArgs -> S Int32 Source # icod_ :: NumGeneralizableArgs -> S Int32 Source # value :: Int32 -> R NumGeneralizableArgs Source # | |
Abstract NumGeneralizableArgs Source # | |
Defined in Agda.TypeChecking.Substitute | |
Apply NumGeneralizableArgs Source # | |
Defined in Agda.TypeChecking.Substitute | |
NFData NumGeneralizableArgs Source # | |
Defined in Agda.TypeChecking.Monad.Base rnf :: NumGeneralizableArgs -> () | |
Show NumGeneralizableArgs Source # | |
Defined in Agda.TypeChecking.Monad.Base showsPrec :: Int -> NumGeneralizableArgs -> ShowS show :: NumGeneralizableArgs -> String showList :: [NumGeneralizableArgs] -> ShowS |
type CompiledRepresentation = Map BackendName [CompilerPragma] Source #
defaultDefn :: ArgInfo -> QName -> Type -> Language -> Defn -> Definition Source #
Create a definition with sensible defaults.
data ExtLamInfo Source #
Additional information for extended lambdas.
ExtLamInfo | |
|
Instances
NamesIn ExtLamInfo Source # | |||||
Defined in Agda.Syntax.Internal.Names namesAndMetasIn' :: Monoid m => (Either QName MetaId -> m) -> ExtLamInfo -> m Source # | |||||
KillRange ExtLamInfo Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
InstantiateFull ExtLamInfo Source # | |||||
Defined in Agda.TypeChecking.Reduce | |||||
EmbPrj ExtLamInfo Source # | |||||
Defined in Agda.TypeChecking.Serialise.Instances.Internal icode :: ExtLamInfo -> S Int32 Source # icod_ :: ExtLamInfo -> S Int32 Source # value :: Int32 -> R ExtLamInfo Source # | |||||
Apply ExtLamInfo Source # | |||||
Defined in Agda.TypeChecking.Substitute apply :: ExtLamInfo -> Args -> ExtLamInfo Source # applyE :: ExtLamInfo -> Elims -> ExtLamInfo Source # | |||||
NFData ExtLamInfo Source # | |||||
Defined in Agda.TypeChecking.Monad.Base rnf :: ExtLamInfo -> () | |||||
Generic ExtLamInfo Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
from :: ExtLamInfo -> Rep ExtLamInfo x to :: Rep ExtLamInfo x -> ExtLamInfo | |||||
Show ExtLamInfo Source # | |||||
Defined in Agda.TypeChecking.Monad.Base showsPrec :: Int -> ExtLamInfo -> ShowS show :: ExtLamInfo -> String showList :: [ExtLamInfo] -> ShowS | |||||
type Rep ExtLamInfo Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep ExtLamInfo = D1 ('MetaData "ExtLamInfo" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "ExtLamInfo" 'PrefixI 'True) (S1 ('MetaSel ('Just "extLamModule") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ModuleName) :*: (S1 ('MetaSel ('Just "extLamAbsurd") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool) :*: S1 ('MetaSel ('Just "extLamSys") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (Maybe System))))) |
modifySystem :: (System -> System) -> ExtLamInfo -> ExtLamInfo Source #
data Projection Source #
Additional information for projection Function
s.
Projection | |
|
Instances
Pretty Projection Source # | |||||
Defined in Agda.TypeChecking.Monad.Base pretty :: Projection -> Doc Source # prettyPrec :: Int -> Projection -> Doc Source # prettyList :: [Projection] -> Doc Source # | |||||
KillRange Projection Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
EmbPrj Projection Source # | |||||
Defined in Agda.TypeChecking.Serialise.Instances.Internal icode :: Projection -> S Int32 Source # icod_ :: Projection -> S Int32 Source # value :: Int32 -> R Projection Source # | |||||
Abstract Projection Source # | |||||
Defined in Agda.TypeChecking.Substitute abstract :: Telescope -> Projection -> Projection Source # | |||||
Apply Projection Source # | |||||
Defined in Agda.TypeChecking.Substitute apply :: Projection -> Args -> Projection Source # applyE :: Projection -> Elims -> Projection Source # | |||||
NFData Projection Source # | |||||
Defined in Agda.TypeChecking.Monad.Base rnf :: Projection -> () | |||||
Generic Projection Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
from :: Projection -> Rep Projection x to :: Rep Projection x -> Projection | |||||
Show Projection Source # | |||||
Defined in Agda.TypeChecking.Monad.Base showsPrec :: Int -> Projection -> ShowS show :: Projection -> String showList :: [Projection] -> ShowS | |||||
type Rep Projection Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep Projection = D1 ('MetaData "Projection" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "Projection" 'PrefixI 'True) ((S1 ('MetaSel ('Just "projProper") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe QName)) :*: S1 ('MetaSel ('Just "projOrig") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :*: (S1 ('MetaSel ('Just "projFromType") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Arg QName)) :*: (S1 ('MetaSel ('Just "projIndex") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Int) :*: S1 ('MetaSel ('Just "projLams") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ProjLams))))) |
Abstractions to build projection function (dropping parameters).
ProjLams | |
|
Instances
Pretty ProjLams Source # | |||||
KillRange ProjLams Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
EmbPrj ProjLams Source # | |||||
Abstract ProjLams Source # | |||||
Apply ProjLams Source # | |||||
Null ProjLams Source # | |||||
NFData ProjLams Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Generic ProjLams Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
| |||||
Show ProjLams Source # | |||||
type Rep ProjLams Source # | |||||
Defined in Agda.TypeChecking.Monad.Base |
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
CopatternMatchingAllowed EtaEquality Source # | |||||
Defined in Agda.TypeChecking.Monad.Base copatternMatchingAllowed :: EtaEquality -> Bool Source # | |||||
PatternMatchingAllowed EtaEquality Source # | |||||
Defined in Agda.TypeChecking.Monad.Base patternMatchingAllowed :: EtaEquality -> Bool Source # | |||||
KillRange EtaEquality Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
EmbPrj EtaEquality Source # | |||||
Defined in Agda.TypeChecking.Serialise.Instances.Internal icode :: EtaEquality -> S Int32 Source # icod_ :: EtaEquality -> S Int32 Source # value :: Int32 -> R EtaEquality Source # | |||||
NFData EtaEquality Source # | |||||
Defined in Agda.TypeChecking.Monad.Base rnf :: EtaEquality -> () | |||||
Generic EtaEquality Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
from :: EtaEquality -> Rep EtaEquality x to :: Rep EtaEquality x -> EtaEquality | |||||
Show EtaEquality Source # | |||||
Defined in Agda.TypeChecking.Monad.Base showsPrec :: Int -> EtaEquality -> ShowS show :: EtaEquality -> String showList :: [EtaEquality] -> ShowS | |||||
Eq EtaEquality Source # | |||||
Defined in Agda.TypeChecking.Monad.Base (==) :: EtaEquality -> EtaEquality -> Bool (/=) :: EtaEquality -> EtaEquality -> Bool | |||||
type Rep EtaEquality Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep EtaEquality = D1 ('MetaData "EtaEquality" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "Specified" 'PrefixI 'True) (S1 ('MetaSel ('Just "theEtaEquality") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 HasEta)) :+: C1 ('MetaCons "Inferred" 'PrefixI 'True) (S1 ('MetaSel ('Just "theEtaEquality") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 HasEta))) |
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? |
FunFirstOrder | Is this function |
FunErasure | Was |
FunAbstract | Is the function abstract? |
FunProj | Is this function a descendant of a field (typically, a projection)? |
Instances
KillRange FunctionFlag Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
EmbPrj FunctionFlag Source # | |||||
Defined in Agda.TypeChecking.Serialise.Instances.Internal icode :: FunctionFlag -> S Int32 Source # icod_ :: FunctionFlag -> S Int32 Source # value :: Int32 -> R FunctionFlag Source # | |||||
SmallSetElement FunctionFlag Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
NFData FunctionFlag Source # | |||||
Defined in Agda.TypeChecking.Monad.Base rnf :: FunctionFlag -> () | |||||
Bounded FunctionFlag Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Enum FunctionFlag Source # | |||||
Defined in Agda.TypeChecking.Monad.Base succ :: FunctionFlag -> FunctionFlag pred :: FunctionFlag -> FunctionFlag toEnum :: Int -> FunctionFlag fromEnum :: FunctionFlag -> Int enumFrom :: FunctionFlag -> [FunctionFlag] enumFromThen :: FunctionFlag -> FunctionFlag -> [FunctionFlag] enumFromTo :: FunctionFlag -> FunctionFlag -> [FunctionFlag] enumFromThenTo :: FunctionFlag -> FunctionFlag -> FunctionFlag -> [FunctionFlag] | |||||
Generic FunctionFlag Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
from :: FunctionFlag -> Rep FunctionFlag x to :: Rep FunctionFlag x -> FunctionFlag | |||||
Ix FunctionFlag Source # | |||||
Defined in Agda.TypeChecking.Monad.Base range :: (FunctionFlag, FunctionFlag) -> [FunctionFlag] # index :: (FunctionFlag, FunctionFlag) -> FunctionFlag -> Int # unsafeIndex :: (FunctionFlag, FunctionFlag) -> FunctionFlag -> Int inRange :: (FunctionFlag, FunctionFlag) -> FunctionFlag -> Bool # rangeSize :: (FunctionFlag, FunctionFlag) -> Int # unsafeRangeSize :: (FunctionFlag, FunctionFlag) -> Int | |||||
Show FunctionFlag Source # | |||||
Defined in Agda.TypeChecking.Monad.Base showsPrec :: Int -> FunctionFlag -> ShowS show :: FunctionFlag -> String showList :: [FunctionFlag] -> ShowS | |||||
Eq FunctionFlag Source # | |||||
Defined in Agda.TypeChecking.Monad.Base (==) :: FunctionFlag -> FunctionFlag -> Bool (/=) :: FunctionFlag -> FunctionFlag -> Bool | |||||
Ord FunctionFlag Source # | |||||
Defined in Agda.TypeChecking.Monad.Base compare :: FunctionFlag -> FunctionFlag -> Ordering (<) :: FunctionFlag -> FunctionFlag -> Bool (<=) :: FunctionFlag -> FunctionFlag -> Bool (>) :: FunctionFlag -> FunctionFlag -> Bool (>=) :: FunctionFlag -> FunctionFlag -> Bool max :: FunctionFlag -> FunctionFlag -> FunctionFlag min :: FunctionFlag -> FunctionFlag -> FunctionFlag | |||||
KillRange (SmallSet FunctionFlag) Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
type Rep FunctionFlag Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep FunctionFlag = D1 ('MetaData "FunctionFlag" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) ((C1 ('MetaCons "FunStatic" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "FunInline" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "FunMacro" 'PrefixI 'False) (U1 :: Type -> Type))) :+: ((C1 ('MetaCons "FunFirstOrder" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "FunErasure" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "FunAbstract" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "FunProj" 'PrefixI 'False) (U1 :: Type -> Type)))) |
Instances
NamesIn CompKit Source # | |||||
Defined in Agda.Syntax.Internal.Names | |||||
KillRange CompKit Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
EmbPrj CompKit Source # | |||||
NFData CompKit Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Generic CompKit Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
| |||||
Show CompKit Source # | |||||
Eq CompKit Source # | |||||
Ord CompKit Source # | |||||
type Rep CompKit Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep CompKit = D1 ('MetaData "CompKit" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "CompKit" 'PrefixI 'True) (S1 ('MetaSel ('Just "nameOfHComp") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe QName)) :*: S1 ('MetaSel ('Just "nameOfTransp") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe QName)))) |
defaultAxiom :: Defn Source #
AxiomData | |
|
Instances
NFData AxiomData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Generic AxiomData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
| |||||
Show AxiomData Source # | |||||
type Rep AxiomData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base |
data DataOrRecSigData Source #
DataOrRecSigData | |
|
Instances
Pretty DataOrRecSigData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base pretty :: DataOrRecSigData -> Doc Source # prettyPrec :: Int -> DataOrRecSigData -> Doc Source # prettyList :: [DataOrRecSigData] -> Doc Source # | |||||
NFData DataOrRecSigData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base rnf :: DataOrRecSigData -> () | |||||
Generic DataOrRecSigData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
from :: DataOrRecSigData -> Rep DataOrRecSigData x to :: Rep DataOrRecSigData x -> DataOrRecSigData | |||||
Show DataOrRecSigData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base showsPrec :: Int -> DataOrRecSigData -> ShowS show :: DataOrRecSigData -> String showList :: [DataOrRecSigData] -> ShowS | |||||
type Rep DataOrRecSigData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep DataOrRecSigData = D1 ('MetaData "DataOrRecSigData" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "DataOrRecSigData" 'PrefixI 'True) (S1 ('MetaSel ('Just "_datarecPars") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Int))) |
data FunctionData Source #
FunctionData | |
|
Instances
Pretty FunctionData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base pretty :: FunctionData -> Doc Source # prettyPrec :: Int -> FunctionData -> Doc Source # prettyList :: [FunctionData] -> Doc Source # | |||||
NFData FunctionData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base rnf :: FunctionData -> () | |||||
Generic FunctionData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
from :: FunctionData -> Rep FunctionData x to :: Rep FunctionData x -> FunctionData | |||||
Show FunctionData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base showsPrec :: Int -> FunctionData -> ShowS show :: FunctionData -> String showList :: [FunctionData] -> ShowS | |||||
type Rep FunctionData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep FunctionData = D1 ('MetaData "FunctionData" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "FunctionData" 'PrefixI 'True) (((S1 ('MetaSel ('Just "_funClauses") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Clause]) :*: (S1 ('MetaSel ('Just "_funCompiled") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe CompiledClauses)) :*: S1 ('MetaSel ('Just "_funSplitTree") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe SplitTree)))) :*: ((S1 ('MetaSel ('Just "_funTreeless") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe Compiled)) :*: S1 ('MetaSel ('Just "_funCovering") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Clause])) :*: (S1 ('MetaSel ('Just "_funInv") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 FunctionInverse) :*: S1 ('MetaSel ('Just "_funMutual") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe [QName]))))) :*: ((S1 ('MetaSel ('Just "_funProjection") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Either ProjectionLikenessMissing Projection)) :*: (S1 ('MetaSel ('Just "_funFlags") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (SmallSet FunctionFlag)) :*: S1 ('MetaSel ('Just "_funTerminates") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe Bool)))) :*: ((S1 ('MetaSel ('Just "_funExtLam") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe ExtLamInfo)) :*: S1 ('MetaSel ('Just "_funWith") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe QName))) :*: (S1 ('MetaSel ('Just "_funIsKanOp") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe QName)) :*: S1 ('MetaSel ('Just "_funOpaque") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 IsOpaque)))))) |
data DatatypeData Source #
DatatypeData | |
|
Instances
Pretty DatatypeData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base pretty :: DatatypeData -> Doc Source # prettyPrec :: Int -> DatatypeData -> Doc Source # prettyList :: [DatatypeData] -> Doc Source # | |||||
NFData DatatypeData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base rnf :: DatatypeData -> () | |||||
Generic DatatypeData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
from :: DatatypeData -> Rep DatatypeData x to :: Rep DatatypeData x -> DatatypeData | |||||
Show DatatypeData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base showsPrec :: Int -> DatatypeData -> ShowS show :: DatatypeData -> String showList :: [DatatypeData] -> ShowS | |||||
type Rep DatatypeData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep DatatypeData = D1 ('MetaData "DatatypeData" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "DatatypeData" 'PrefixI 'True) (((S1 ('MetaSel ('Just "_dataPars") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Nat) :*: S1 ('MetaSel ('Just "_dataIxs") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Nat)) :*: (S1 ('MetaSel ('Just "_dataClause") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe Clause)) :*: (S1 ('MetaSel ('Just "_dataCons") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [QName]) :*: S1 ('MetaSel ('Just "_dataSort") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Sort)))) :*: ((S1 ('MetaSel ('Just "_dataMutual") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe [QName])) :*: S1 ('MetaSel ('Just "_dataAbstr") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 IsAbstract)) :*: (S1 ('MetaSel ('Just "_dataPathCons") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [QName]) :*: (S1 ('MetaSel ('Just "_dataTranspIx") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe QName)) :*: S1 ('MetaSel ('Just "_dataTransp") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe QName))))))) |
data RecordData Source #
RecordData | |
|
Instances
Pretty RecordData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base pretty :: RecordData -> Doc Source # prettyPrec :: Int -> RecordData -> Doc Source # prettyList :: [RecordData] -> Doc Source # | |||||
NFData RecordData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base rnf :: RecordData -> () | |||||
Generic RecordData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
from :: RecordData -> Rep RecordData x to :: Rep RecordData x -> RecordData | |||||
Show RecordData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base showsPrec :: Int -> RecordData -> ShowS show :: RecordData -> String showList :: [RecordData] -> ShowS | |||||
type Rep RecordData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep RecordData = D1 ('MetaData "RecordData" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "RecordData" 'PrefixI 'True) (((S1 ('MetaSel ('Just "_recPars") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Nat) :*: (S1 ('MetaSel ('Just "_recClause") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe Clause)) :*: S1 ('MetaSel ('Just "_recConHead") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ConHead))) :*: (S1 ('MetaSel ('Just "_recNamedCon") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool) :*: (S1 ('MetaSel ('Just "_recFields") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Dom QName]) :*: S1 ('MetaSel ('Just "_recTel") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Telescope)))) :*: ((S1 ('MetaSel ('Just "_recMutual") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe [QName])) :*: (S1 ('MetaSel ('Just "_recEtaEquality'") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 EtaEquality) :*: S1 ('MetaSel ('Just "_recPatternMatching") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 PatternOrCopattern))) :*: ((S1 ('MetaSel ('Just "_recInduction") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe Induction)) :*: S1 ('MetaSel ('Just "_recTerminates") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe Bool))) :*: (S1 ('MetaSel ('Just "_recAbstr") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 IsAbstract) :*: S1 ('MetaSel ('Just "_recComp") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 CompKit)))))) |
data ConstructorData Source #
ConstructorData | |
|
Instances
Pretty ConstructorData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base pretty :: ConstructorData -> Doc Source # prettyPrec :: Int -> ConstructorData -> Doc Source # prettyList :: [ConstructorData] -> Doc Source # | |||||
NFData ConstructorData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base rnf :: ConstructorData -> () | |||||
Generic ConstructorData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
from :: ConstructorData -> Rep ConstructorData x to :: Rep ConstructorData x -> ConstructorData | |||||
Show ConstructorData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base showsPrec :: Int -> ConstructorData -> ShowS show :: ConstructorData -> String showList :: [ConstructorData] -> ShowS | |||||
type Rep ConstructorData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep ConstructorData = D1 ('MetaData "ConstructorData" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "ConstructorData" 'PrefixI 'True) (((S1 ('MetaSel ('Just "_conPars") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Int) :*: S1 ('MetaSel ('Just "_conArity") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Int)) :*: (S1 ('MetaSel ('Just "_conSrcCon") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ConHead) :*: (S1 ('MetaSel ('Just "_conData") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Just "_conAbstr") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 IsAbstract)))) :*: ((S1 ('MetaSel ('Just "_conComp") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 CompKit) :*: (S1 ('MetaSel ('Just "_conProj") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe [QName])) :*: S1 ('MetaSel ('Just "_conForced") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [IsForced]))) :*: (S1 ('MetaSel ('Just "_conErased") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe [Bool])) :*: (S1 ('MetaSel ('Just "_conErasure") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Bool) :*: S1 ('MetaSel ('Just "_conInline") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Bool)))))) |
data PrimitiveData Source #
PrimitiveData | |
|
Instances
Pretty PrimitiveData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base pretty :: PrimitiveData -> Doc Source # prettyPrec :: Int -> PrimitiveData -> Doc Source # prettyList :: [PrimitiveData] -> Doc Source # | |||||
NFData PrimitiveData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base rnf :: PrimitiveData -> () | |||||
Generic PrimitiveData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
from :: PrimitiveData -> Rep PrimitiveData x to :: Rep PrimitiveData x -> PrimitiveData | |||||
Show PrimitiveData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base showsPrec :: Int -> PrimitiveData -> ShowS show :: PrimitiveData -> String showList :: [PrimitiveData] -> ShowS | |||||
type Rep PrimitiveData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep PrimitiveData = D1 ('MetaData "PrimitiveData" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "PrimitiveData" 'PrefixI 'True) ((S1 ('MetaSel ('Just "_primAbstr") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 IsAbstract) :*: (S1 ('MetaSel ('Just "_primName") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 PrimitiveId) :*: S1 ('MetaSel ('Just "_primClauses") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Clause]))) :*: (S1 ('MetaSel ('Just "_primInv") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 FunctionInverse) :*: (S1 ('MetaSel ('Just "_primCompiled") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe CompiledClauses)) :*: S1 ('MetaSel ('Just "_primOpaque") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 IsOpaque))))) |
data PrimitiveSortData Source #
Instances
Pretty PrimitiveSortData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base pretty :: PrimitiveSortData -> Doc Source # prettyPrec :: Int -> PrimitiveSortData -> Doc Source # prettyList :: [PrimitiveSortData] -> Doc Source # | |||||
NFData PrimitiveSortData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base rnf :: PrimitiveSortData -> () | |||||
Generic PrimitiveSortData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
from :: PrimitiveSortData -> Rep PrimitiveSortData x to :: Rep PrimitiveSortData x -> PrimitiveSortData | |||||
Show PrimitiveSortData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base showsPrec :: Int -> PrimitiveSortData -> ShowS show :: PrimitiveSortData -> String showList :: [PrimitiveSortData] -> ShowS | |||||
type Rep PrimitiveSortData Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep PrimitiveSortData = D1 ('MetaData "PrimitiveSortData" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "PrimitiveSortData" 'PrefixI 'True) (S1 ('MetaSel ('Just "_primSortName") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 BuiltinSort) :*: S1 ('MetaSel ('Just "_primSortSort") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Sort))) |
pattern DataOrRecSig :: Int -> Defn Source #
pattern PrimitiveSort :: BuiltinSort -> Sort -> Defn Source #
axiomConstTransp :: Defn -> Bool Source #
datarecPars :: Defn -> Int Source #
data ProjectionLikenessMissing Source #
Indicates the reason behind a function having not been marked projection-like.
MaybeProjection | Projection-likeness analysis has not run on this function yet. It may do so in the future. |
NeverProjection | The user has requested that this function be not be marked projection-like. The analysis may already have run on this function, but the results have been discarded, and it will not be run again. |
Instances
Pretty ProjectionLikenessMissing Source # | |||||
Defined in Agda.TypeChecking.Monad.Base pretty :: ProjectionLikenessMissing -> Doc Source # prettyPrec :: Int -> ProjectionLikenessMissing -> Doc Source # prettyList :: [ProjectionLikenessMissing] -> Doc Source # | |||||
KillRange ProjectionLikenessMissing Source # | |||||
EmbPrj ProjectionLikenessMissing Source # | |||||
Defined in Agda.TypeChecking.Serialise.Instances.Internal icode :: ProjectionLikenessMissing -> S Int32 Source # icod_ :: ProjectionLikenessMissing -> S Int32 Source # value :: Int32 -> R ProjectionLikenessMissing Source # | |||||
NFData ProjectionLikenessMissing Source # | |||||
Defined in Agda.TypeChecking.Monad.Base rnf :: ProjectionLikenessMissing -> () | |||||
Bounded ProjectionLikenessMissing Source # | |||||
Enum ProjectionLikenessMissing Source # | |||||
Defined in Agda.TypeChecking.Monad.Base succ :: ProjectionLikenessMissing -> ProjectionLikenessMissing pred :: ProjectionLikenessMissing -> ProjectionLikenessMissing toEnum :: Int -> ProjectionLikenessMissing fromEnum :: ProjectionLikenessMissing -> Int enumFrom :: ProjectionLikenessMissing -> [ProjectionLikenessMissing] enumFromThen :: ProjectionLikenessMissing -> ProjectionLikenessMissing -> [ProjectionLikenessMissing] enumFromTo :: ProjectionLikenessMissing -> ProjectionLikenessMissing -> [ProjectionLikenessMissing] enumFromThenTo :: ProjectionLikenessMissing -> ProjectionLikenessMissing -> ProjectionLikenessMissing -> [ProjectionLikenessMissing] | |||||
Generic ProjectionLikenessMissing Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
from :: ProjectionLikenessMissing -> Rep ProjectionLikenessMissing x to :: Rep ProjectionLikenessMissing x -> ProjectionLikenessMissing | |||||
Show ProjectionLikenessMissing Source # | |||||
Defined in Agda.TypeChecking.Monad.Base showsPrec :: Int -> ProjectionLikenessMissing -> ShowS show :: ProjectionLikenessMissing -> String showList :: [ProjectionLikenessMissing] -> ShowS | |||||
type Rep ProjectionLikenessMissing Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep ProjectionLikenessMissing = D1 ('MetaData "ProjectionLikenessMissing" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "MaybeProjection" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "NeverProjection" 'PrefixI 'False) (U1 :: Type -> Type)) |
type FunctionInverse = FunctionInverse' Clause Source #
funClauses :: Defn -> [Clause] Source #
funCompiled :: Defn -> Maybe CompiledClauses Source #
funCovering :: Defn -> [Clause] Source #
funInv :: Defn -> FunctionInverse Source #
funProjection :: Defn -> Either ProjectionLikenessMissing Projection Source #
funTerminates :: Defn -> Maybe Bool Source #
dataAbstr :: Defn -> IsAbstract Source #
dataPathCons :: Defn -> [QName] Source #
recConHead :: Defn -> ConHead Source #
recNamedCon :: Defn -> Bool Source #
recEtaEquality' :: Defn -> EtaEquality Source #
recTerminates :: Defn -> Maybe Bool Source #
recAbstr :: Defn -> IsAbstract Source #
conAbstr :: Defn -> IsAbstract Source #
conErasure :: Defn -> Bool Source #
primAbstr :: Defn -> IsAbstract Source #
primName :: Defn -> PrimitiveId Source #
primClauses :: Defn -> [Clause] Source #
primInv :: Defn -> FunctionInverse Source #
primCompiled :: Defn -> Maybe CompiledClauses Source #
primOpaque :: Defn -> IsOpaque Source #
data BuiltinSort Source #
Instances
KillRange BuiltinSort Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
EmbPrj BuiltinSort Source # | |||||
Defined in Agda.TypeChecking.Serialise.Instances.Internal icode :: BuiltinSort -> S Int32 Source # icod_ :: BuiltinSort -> S Int32 Source # value :: Int32 -> R BuiltinSort Source # | |||||
NFData BuiltinSort Source # | |||||
Defined in Agda.TypeChecking.Monad.Base rnf :: BuiltinSort -> () | |||||
Generic BuiltinSort Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
from :: BuiltinSort -> Rep BuiltinSort x to :: Rep BuiltinSort x -> BuiltinSort | |||||
Show BuiltinSort Source # | |||||
Defined in Agda.TypeChecking.Monad.Base showsPrec :: Int -> BuiltinSort -> ShowS show :: BuiltinSort -> String showList :: [BuiltinSort] -> ShowS | |||||
Eq BuiltinSort Source # | |||||
Defined in Agda.TypeChecking.Monad.Base (==) :: BuiltinSort -> BuiltinSort -> Bool (/=) :: BuiltinSort -> BuiltinSort -> Bool | |||||
type Rep BuiltinSort Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep BuiltinSort = D1 ('MetaData "BuiltinSort" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) ((C1 ('MetaCons "SortUniv" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Univ)) :+: C1 ('MetaCons "SortOmega" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Univ))) :+: (C1 ('MetaCons "SortIntervalUniv" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "SortLevelUniv" 'PrefixI 'False) (U1 :: Type -> Type))) |
data BuiltinDescriptor Source #
BuiltinData (TCM Type) [BuiltinId] | |
BuiltinDataCons (TCM Type) | |
BuiltinPrim PrimitiveId (Term -> TCM ()) | |
BuiltinSort BuiltinSort | |
BuiltinPostulate Relevance (TCM Type) | |
BuiltinUnknown (Maybe (TCM Type)) (Term -> Type -> TCM ()) | Builtin of any kind.
Type can be checked ( |
primSortName :: Defn -> BuiltinSort Source #
primSortSort :: Defn -> Sort Source #
data FunctionInverse' c Source #
Instances
DropArgs FunctionInverse Source # | |||||
Defined in Agda.TypeChecking.DropArgs dropArgs :: Int -> FunctionInverse -> FunctionInverse Source # | |||||
InstantiateFull FunctionInverse Source # | |||||
Defined in Agda.TypeChecking.Reduce | |||||
Abstract FunctionInverse Source # | |||||
Defined in Agda.TypeChecking.Substitute abstract :: Telescope -> FunctionInverse -> FunctionInverse Source # | |||||
Apply FunctionInverse Source # | |||||
Defined in Agda.TypeChecking.Substitute apply :: FunctionInverse -> Args -> FunctionInverse Source # applyE :: FunctionInverse -> Elims -> FunctionInverse Source # | |||||
Functor FunctionInverse' Source # | |||||
Defined in Agda.TypeChecking.Monad.Base fmap :: (a -> b) -> FunctionInverse' a -> FunctionInverse' b (<$) :: a -> FunctionInverse' b -> FunctionInverse' a # | |||||
Pretty c => Pretty (FunctionInverse' c) Source # | |||||
Defined in Agda.TypeChecking.Monad.Base pretty :: FunctionInverse' c -> Doc Source # prettyPrec :: Int -> FunctionInverse' c -> Doc Source # prettyList :: [FunctionInverse' c] -> Doc Source # | |||||
NamesIn a => NamesIn (FunctionInverse' a) Source # | |||||
Defined in Agda.Syntax.Internal.Names namesAndMetasIn' :: Monoid m => (Either QName MetaId -> m) -> FunctionInverse' a -> m Source # | |||||
KillRange c => KillRange (FunctionInverse' c) Source # | |||||
Defined in Agda.TypeChecking.Monad.Base killRange :: KillRangeT (FunctionInverse' c) Source # | |||||
EmbPrj a => EmbPrj (FunctionInverse' a) Source # | |||||
Defined in Agda.TypeChecking.Serialise.Instances.Internal icode :: FunctionInverse' a -> S Int32 Source # icod_ :: FunctionInverse' a -> S Int32 Source # value :: Int32 -> R (FunctionInverse' a) Source # | |||||
NFData c => NFData (FunctionInverse' c) Source # | |||||
Defined in Agda.TypeChecking.Monad.Base rnf :: FunctionInverse' c -> () | |||||
Generic (FunctionInverse' c) Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
from :: FunctionInverse' c -> Rep (FunctionInverse' c) x to :: Rep (FunctionInverse' c) x -> FunctionInverse' c | |||||
Show c => Show (FunctionInverse' c) Source # | |||||
Defined in Agda.TypeChecking.Monad.Base showsPrec :: Int -> FunctionInverse' c -> ShowS show :: FunctionInverse' c -> String showList :: [FunctionInverse' c] -> ShowS | |||||
type Rep (FunctionInverse' c) Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep (FunctionInverse' c) = D1 ('MetaData "FunctionInverse'" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "NotInjective" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "Inverse" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (InversionMap c)))) |
recRecursive :: Defn -> Bool Source #
Is the record type recursive?
recRecursive_ :: RecordData -> Bool Source #
recEtaEquality :: Defn -> HasEta Source #
_recEtaEquality :: RecordData -> HasEta Source #
emptyFunctionData :: HasOptions m => m FunctionData Source #
A template for creating Function
definitions, with sensible
defaults.
emptyFunctionData_ :: Bool -> FunctionData Source #
emptyFunction :: HasOptions m => m Defn Source #
emptyFunction_ :: Bool -> Defn Source #
funFlag_ :: FunctionFlag -> Lens' FunctionData Bool Source #
funMacro_ :: Lens' FunctionData Bool Source #
funFirstOrder :: Lens' Defn Bool Source #
Toggle the FunFirstOrder
flag.
funErasure :: Lens' Defn Bool Source #
Toggle the FunErasure
flag.
funAbstract :: Lens' Defn Bool Source #
Toggle the FunAbstract
flag.
funAbstr :: Lens' Defn IsAbstract Source #
Toggle the FunAbstract
flag.
funAbstract_ :: Lens' FunctionData Bool Source #
Toggle the FunAbstract
flag.
funAbstr_ :: Lens' FunctionData IsAbstract Source #
Toggle the FunAbstract
flag.
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
Null Simplification Source # | |||||
Defined in Agda.TypeChecking.Monad.Base empty :: Simplification Source # null :: Simplification -> Bool Source # | |||||
NFData Simplification Source # | |||||
Defined in Agda.TypeChecking.Monad.Base rnf :: Simplification -> () | |||||
Monoid Simplification Source # | |||||
Defined in Agda.TypeChecking.Monad.Base mappend :: Simplification -> Simplification -> Simplification mconcat :: [Simplification] -> Simplification | |||||
Semigroup Simplification Source # | |||||
Defined in Agda.TypeChecking.Monad.Base (<>) :: Simplification -> Simplification -> Simplification # sconcat :: NonEmpty Simplification -> Simplification stimes :: Integral b => b -> Simplification -> Simplification | |||||
Generic Simplification Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
from :: Simplification -> Rep Simplification x to :: Rep Simplification x -> Simplification | |||||
Show Simplification Source # | |||||
Defined in Agda.TypeChecking.Monad.Base showsPrec :: Int -> Simplification -> ShowS show :: Simplification -> String showList :: [Simplification] -> ShowS | |||||
Eq Simplification Source # | |||||
Defined in Agda.TypeChecking.Monad.Base (==) :: Simplification -> Simplification -> Bool (/=) :: Simplification -> Simplification -> Bool | |||||
type Rep Simplification Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep Simplification = D1 ('MetaData "Simplification" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "YesSimplification" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "NoSimplification" 'PrefixI 'False) (U1 :: Type -> Type)) |
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
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
SmallSetElement AllowedReduction Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
NFData AllowedReduction Source # | |||||
Defined in Agda.TypeChecking.Monad.Base rnf :: AllowedReduction -> () | |||||
Bounded AllowedReduction Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Enum AllowedReduction Source # | |||||
Defined in Agda.TypeChecking.Monad.Base succ :: AllowedReduction -> AllowedReduction pred :: AllowedReduction -> AllowedReduction toEnum :: Int -> AllowedReduction fromEnum :: AllowedReduction -> Int enumFrom :: AllowedReduction -> [AllowedReduction] enumFromThen :: AllowedReduction -> AllowedReduction -> [AllowedReduction] enumFromTo :: AllowedReduction -> AllowedReduction -> [AllowedReduction] enumFromThenTo :: AllowedReduction -> AllowedReduction -> AllowedReduction -> [AllowedReduction] | |||||
Generic AllowedReduction Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
from :: AllowedReduction -> Rep AllowedReduction x to :: Rep AllowedReduction x -> AllowedReduction | |||||
Ix AllowedReduction Source # | |||||
Defined in Agda.TypeChecking.Monad.Base range :: (AllowedReduction, AllowedReduction) -> [AllowedReduction] # index :: (AllowedReduction, AllowedReduction) -> AllowedReduction -> Int # unsafeIndex :: (AllowedReduction, AllowedReduction) -> AllowedReduction -> Int inRange :: (AllowedReduction, AllowedReduction) -> AllowedReduction -> Bool # rangeSize :: (AllowedReduction, AllowedReduction) -> Int # unsafeRangeSize :: (AllowedReduction, AllowedReduction) -> Int | |||||
Show AllowedReduction Source # | |||||
Defined in Agda.TypeChecking.Monad.Base showsPrec :: Int -> AllowedReduction -> ShowS show :: AllowedReduction -> String showList :: [AllowedReduction] -> ShowS | |||||
Eq AllowedReduction Source # | |||||
Defined in Agda.TypeChecking.Monad.Base (==) :: AllowedReduction -> AllowedReduction -> Bool (/=) :: AllowedReduction -> AllowedReduction -> Bool | |||||
Ord AllowedReduction Source # | |||||
Defined in Agda.TypeChecking.Monad.Base compare :: AllowedReduction -> AllowedReduction -> Ordering (<) :: AllowedReduction -> AllowedReduction -> Bool (<=) :: AllowedReduction -> AllowedReduction -> Bool (>) :: AllowedReduction -> AllowedReduction -> Bool (>=) :: AllowedReduction -> AllowedReduction -> Bool max :: AllowedReduction -> AllowedReduction -> AllowedReduction min :: AllowedReduction -> AllowedReduction -> AllowedReduction | |||||
type Rep AllowedReduction Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep AllowedReduction = D1 ('MetaData "AllowedReduction" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (((C1 ('MetaCons "ProjectionReductions" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "InlineReductions" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "CopatternReductions" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "FunctionReductions" 'PrefixI 'False) (U1 :: Type -> Type))) :+: ((C1 ('MetaCons "RecursiveReductions" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "LevelReductions" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "TypeLevelReductions" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "UnconfirmedReductions" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "NonTerminatingReductions" 'PrefixI 'False) (U1 :: Type -> Type))))) |
allReductions :: AllowedReductions Source #
Not quite all reductions (skip non-terminating reductions)
data ReduceDefs Source #
OnlyReduceDefs (Set QName) | |
DontReduceDefs (Set QName) |
Instances
NFData ReduceDefs Source # | |||||
Defined in Agda.TypeChecking.Monad.Base rnf :: ReduceDefs -> () | |||||
Generic ReduceDefs Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
from :: ReduceDefs -> Rep ReduceDefs x to :: Rep ReduceDefs x -> ReduceDefs | |||||
type Rep ReduceDefs Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep ReduceDefs = D1 ('MetaData "ReduceDefs" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "OnlyReduceDefs" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Set QName))) :+: C1 ('MetaCons "DontReduceDefs" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Set QName)))) |
locallyReduceDefs :: MonadTCEnv m => ReduceDefs -> m a -> m a Source #
locallyTC :: MonadTCEnv m => Lens' TCEnv a -> (a -> a) -> m b -> m b Source #
Modify the lens-indicated part of the TCEnv
in a subcomputation.
locallyReduceAllDefs :: MonadTCEnv m => m a -> m a Source #
shouldReduceDef :: MonadTCEnv m => QName -> m Bool Source #
asksTC :: MonadTCEnv m => (TCEnv -> a) -> m a Source #
toReduceDefs :: (Bool, [QName]) -> ReduceDefs Source #
fromReduceDefs :: ReduceDefs -> (Bool, [QName]) Source #
locallyReconstructed :: MonadTCEnv m => m a -> m a Source #
eReconstructed :: Lens' TCEnv Bool Source #
isReconstructed :: MonadTCEnv m => m Bool Source #
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 #
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?
defForced :: Definition -> [IsForced] Source #
type InversionMap c = Map TermHead [c] Source #
Instances
Pretty TermHead Source # | |||||
KillRange TermHead Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
EmbPrj TermHead Source # | |||||
NFData TermHead Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Generic TermHead Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
| |||||
Show TermHead Source # | |||||
Eq TermHead Source # | |||||
Ord TermHead Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
type Rep TermHead Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep TermHead = D1 ('MetaData "TermHead" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) ((C1 ('MetaCons "SortHead" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "PiHead" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "ConsHead" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: (C1 ('MetaCons "VarHead" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Nat)) :+: C1 ('MetaCons "UnknownHead" 'PrefixI 'False) (U1 :: Type -> Type)))) |
itableCounts :: Lens' InstanceTable (Map QName Int) Source #
pattern SortProp :: BuiltinSort Source #
pattern SortSet :: BuiltinSort Source #
pattern SortStrictSet :: BuiltinSort Source #
pattern SortPropOmega :: BuiltinSort Source #
pattern SortSetOmega :: BuiltinSort Source #
pattern SortStrictSetOmega :: BuiltinSort Source #
ifTopLevelAndHighlightingLevelIsOr :: MonadTCEnv tcm => HighlightingLevel -> Bool -> tcm () -> tcm () Source #
ifTopLevelAndHighlightingLevelIs l b m
runs m
when we're
type-checking the top-level module (or before we've started doing
this) and either the highlighting level is at least l
or b
is
True
.
ifTopLevelAndHighlightingLevelIs :: MonadTCEnv tcm => HighlightingLevel -> tcm () -> tcm () Source #
ifTopLevelAndHighlightingLevelIs l m
runs m
when we're
type-checking the top-level module (or before we've started doing
this) and the highlighting level is at least l
.
type LetBindings = Map Name (Open LetBinding) Source #
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
NFData AbstractMode Source # | |||||
Defined in Agda.TypeChecking.Monad.Base rnf :: AbstractMode -> () | |||||
Generic AbstractMode Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
from :: AbstractMode -> Rep AbstractMode x to :: Rep AbstractMode x -> AbstractMode | |||||
Show AbstractMode Source # | |||||
Defined in Agda.TypeChecking.Monad.Base showsPrec :: Int -> AbstractMode -> ShowS show :: AbstractMode -> String showList :: [AbstractMode] -> ShowS | |||||
Eq AbstractMode Source # | |||||
Defined in Agda.TypeChecking.Monad.Base (==) :: AbstractMode -> AbstractMode -> Bool (/=) :: AbstractMode -> AbstractMode -> Bool | |||||
type Rep AbstractMode Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep AbstractMode = D1 ('MetaData "AbstractMode" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "AbstractMode" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "ConcreteMode" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "IgnoreAbstractMode" 'PrefixI 'False) (U1 :: Type -> Type))) |
data UnquoteFlags Source #
UnquoteFlags | |
|
Instances
NFData UnquoteFlags Source # | |||||
Defined in Agda.TypeChecking.Monad.Base rnf :: UnquoteFlags -> () | |||||
Generic UnquoteFlags Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
from :: UnquoteFlags -> Rep UnquoteFlags x to :: Rep UnquoteFlags x -> UnquoteFlags | |||||
type Rep UnquoteFlags Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep UnquoteFlags = D1 ('MetaData "UnquoteFlags" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "UnquoteFlags" 'PrefixI 'True) (S1 ('MetaSel ('Just "_unquoteNormalise") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool))) |
unquoteNormalise :: Lens' UnquoteFlags Bool Source #
eUnquoteNormalise :: Lens' TCEnv Bool Source #
eAnonymousModules :: Lens' TCEnv [(ModuleName, Nat)] Source #
eTerminationCheck :: Lens' TCEnv (TerminationCheck ()) Source #
eSolvingConstraints :: Lens' TCEnv Bool Source #
eCheckingWhere :: Lens' TCEnv Bool Source #
eWorkingOnTypes :: Lens' TCEnv Bool Source #
eAssignMetas :: Lens' TCEnv Bool Source #
eQuantity :: Lens' TCEnv Quantity Source #
Note that this lens does not satisfy all lens laws: If hard
compile-time mode is enabled, then quantities other than zero are
replaced by __IMPOSSIBLE__
.
eHardCompileTimeMode :: Lens' TCEnv Bool Source #
eSplitOnStrict :: Lens' TCEnv Bool Source #
eDisplayFormsEnabled :: Lens' TCEnv Bool Source #
eFoldLetBindings :: Lens' TCEnv Bool Source #
eExpandLastBool :: Lens' TCEnv Bool Source #
isExpandLast :: ExpandHidden -> Bool Source #
toExpandLast :: Bool -> ExpandHidden Source #
eInjectivityDepth :: Lens' TCEnv Int Source #
eCompareBlocked :: Lens' TCEnv Bool Source #
ePrintDomainFreePi :: Lens' TCEnv Bool Source #
ePrintMetasBare :: Lens' TCEnv Bool Source #
eInsideDotPattern :: Lens' TCEnv Bool Source #
eInstanceDepth :: Lens' TCEnv Int Source #
eIsDebugPrinting :: Lens' TCEnv Bool Source #
eCallByNeed :: Lens' TCEnv Bool Source #
eCheckpoints :: Lens' TCEnv (Map CheckpointId Substitution) Source #
eGeneralizedVars :: Lens' TCEnv (Map QName GeneralizedValue) Source #
eConflComputingOverlap :: Lens' TCEnv Bool Source #
eCurrentlyElaborating :: Lens' TCEnv Bool Source #
currentModality :: MonadTCEnv m => m Modality Source #
The current modality.
Note that the returned cohesion component is always unitCohesion
.
onLetBindingType :: (Dom Type -> Dom Type) -> LetBinding -> LetBinding Source #
isDontExpandLast :: ExpandHidden -> Bool Source #
data CandidateKind Source #
Instances
NFData CandidateKind Source # | |||||
Defined in Agda.TypeChecking.Monad.Base rnf :: CandidateKind -> () | |||||
Generic CandidateKind Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
from :: CandidateKind -> Rep CandidateKind x to :: Rep CandidateKind x -> CandidateKind | |||||
Show CandidateKind Source # | |||||
Defined in Agda.TypeChecking.Monad.Base showsPrec :: Int -> CandidateKind -> ShowS show :: CandidateKind -> String showList :: [CandidateKind] -> ShowS | |||||
Eq CandidateKind | |||||
Defined in Agda.TypeChecking.Substitute (==) :: CandidateKind -> CandidateKind -> Bool (/=) :: CandidateKind -> CandidateKind -> Bool | |||||
Ord CandidateKind | |||||
Defined in Agda.TypeChecking.Substitute compare :: CandidateKind -> CandidateKind -> Ordering (<) :: CandidateKind -> CandidateKind -> Bool (<=) :: CandidateKind -> CandidateKind -> Bool (>) :: CandidateKind -> CandidateKind -> Bool (>=) :: CandidateKind -> CandidateKind -> Bool max :: CandidateKind -> CandidateKind -> CandidateKind min :: CandidateKind -> CandidateKind -> CandidateKind | |||||
type Rep CandidateKind Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep CandidateKind = D1 ('MetaData "CandidateKind" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "LocalCandidate" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "GlobalCandidate" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName))) |
data TerminationError Source #
Information about a mutual block which did not pass the termination checker.
TerminationError | |
|
Instances
NFData TerminationError Source # | |||||
Defined in Agda.TypeChecking.Monad.Base rnf :: TerminationError -> () | |||||
Generic TerminationError Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
from :: TerminationError -> Rep TerminationError x to :: Rep TerminationError x -> TerminationError | |||||
Show TerminationError Source # | |||||
Defined in Agda.TypeChecking.Monad.Base showsPrec :: Int -> TerminationError -> ShowS show :: TerminationError -> String showList :: [TerminationError] -> ShowS | |||||
type Rep TerminationError Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep TerminationError = D1 ('MetaData "TerminationError" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "TerminationError" 'PrefixI 'True) (S1 ('MetaSel ('Just "termErrFunctions") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [QName]) :*: S1 ('MetaSel ('Just "termErrCalls") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [CallInfo]))) |
data IllegalRewriteRuleReason Source #
LHSNotDefinitionOrConstructor | |
VariablesNotBoundByLHS IntSet | |
VariablesBoundMoreThanOnce IntSet | |
LHSReduces Term Term | |
HeadSymbolIsProjection QName | |
HeadSymbolIsProjectionLikeFunction QName | |
HeadSymbolIsTypeConstructor QName | |
HeadSymbolContainsMetas QName | |
ConstructorParametersNotGeneral ConHead Args | |
ContainsUnsolvedMetaVariables (Set MetaId) | |
BlockedOnProblems (Set ProblemId) | |
RequiresDefinitions (Set QName) | |
DoesNotTargetRewriteRelation | |
BeforeFunctionDefinition | |
BeforeMutualFunctionDefinition QName | |
DuplicateRewriteRule |
Instances
EmbPrj IllegalRewriteRuleReason Source # | |||||
Defined in Agda.TypeChecking.Serialise.Instances.Errors icode :: IllegalRewriteRuleReason -> S Int32 Source # icod_ :: IllegalRewriteRuleReason -> S Int32 Source # value :: Int32 -> R IllegalRewriteRuleReason Source # | |||||
NFData IllegalRewriteRuleReason Source # | |||||
Defined in Agda.TypeChecking.Monad.Base rnf :: IllegalRewriteRuleReason -> () | |||||
Generic IllegalRewriteRuleReason Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
from :: IllegalRewriteRuleReason -> Rep IllegalRewriteRuleReason x to :: Rep IllegalRewriteRuleReason x -> IllegalRewriteRuleReason | |||||
Show IllegalRewriteRuleReason Source # | |||||
Defined in Agda.TypeChecking.Monad.Base showsPrec :: Int -> IllegalRewriteRuleReason -> ShowS show :: IllegalRewriteRuleReason -> String showList :: [IllegalRewriteRuleReason] -> ShowS | |||||
type Rep IllegalRewriteRuleReason Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep IllegalRewriteRuleReason = D1 ('MetaData "IllegalRewriteRuleReason" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) ((((C1 ('MetaCons "LHSNotDefinitionOrConstructor" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "VariablesNotBoundByLHS" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 IntSet))) :+: (C1 ('MetaCons "VariablesBoundMoreThanOnce" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 IntSet)) :+: C1 ('MetaCons "LHSReduces" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)))) :+: ((C1 ('MetaCons "HeadSymbolIsProjection" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "HeadSymbolIsProjectionLikeFunction" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName))) :+: (C1 ('MetaCons "HeadSymbolIsTypeConstructor" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "HeadSymbolContainsMetas" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName))))) :+: (((C1 ('MetaCons "ConstructorParametersNotGeneral" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ConHead) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Args)) :+: C1 ('MetaCons "ContainsUnsolvedMetaVariables" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Set MetaId)))) :+: (C1 ('MetaCons "BlockedOnProblems" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Set ProblemId))) :+: C1 ('MetaCons "RequiresDefinitions" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Set QName))))) :+: ((C1 ('MetaCons "DoesNotTargetRewriteRelation" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BeforeFunctionDefinition" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "BeforeMutualFunctionDefinition" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "DuplicateRewriteRule" 'PrefixI 'False) (U1 :: Type -> Type))))) |
warningName :: Warning -> WarningName Source #
isSourceCodeWarning :: WarningName -> Bool Source #
Should warnings of that type be serialized?
Only when changes in the source code can silence or influence the warning.
tcWarningOrigin :: TCWarning -> SrcFile Source #
Information about a call.
CallInfo | |
|
Instances
Pretty CallInfo Source # | We only | ||||
HasRange CallInfo Source # | |||||
PrettyTCM CallInfo Source # | |||||
Defined in Agda.TypeChecking.Pretty.Call | |||||
NFData CallInfo Source # | |||||
Defined in Agda.TypeChecking.Monad.Base | |||||
Generic CallInfo Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
| |||||
Show CallInfo Source # | |||||
type Rep CallInfo Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep CallInfo = D1 ('MetaData "CallInfo" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "CallInfo" 'PrefixI 'True) (S1 ('MetaSel ('Just "callInfoTarget") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Just "callInfoCall") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Closure Term)))) |
data ErasedDatatypeReason Source #
The reason for an ErasedDatatype
error.
SeveralConstructors | There are several constructors. |
NoErasedMatches | The flag |
NoK | The K rule is not activated. |
Instances
NFData ErasedDatatypeReason Source # | |||||
Defined in Agda.TypeChecking.Monad.Base rnf :: ErasedDatatypeReason -> () | |||||
Generic ErasedDatatypeReason Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
from :: ErasedDatatypeReason -> Rep ErasedDatatypeReason x to :: Rep ErasedDatatypeReason x -> ErasedDatatypeReason | |||||
Show ErasedDatatypeReason Source # | |||||
Defined in Agda.TypeChecking.Monad.Base showsPrec :: Int -> ErasedDatatypeReason -> ShowS show :: ErasedDatatypeReason -> String showList :: [ErasedDatatypeReason] -> ShowS | |||||
type Rep ErasedDatatypeReason Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep ErasedDatatypeReason = D1 ('MetaData "ErasedDatatypeReason" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "SeveralConstructors" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "NoErasedMatches" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "NoK" 'PrefixI 'False) (U1 :: Type -> Type))) |
data SplitError Source #
Error when splitting a pattern variable into possible constructor patterns.
NotADatatype (Closure Type) | Neither data type nor record. |
BlockedType Blocker (Closure Type) | Type could not be sufficiently reduced. |
ErasedDatatype ErasedDatatypeReason (Closure Type) | Data type, but in erased position. |
CoinductiveDatatype (Closure Type) | Split on codata not allowed. UNUSED, but keep! -- | NoRecordConstructor Type -- ^ record type, but no constructor |
UnificationStuck | |
| |
CosplitCatchall | Copattern split with a catchall |
CosplitNoTarget | We do not know the target type of the clause. |
CosplitNoRecordType (Closure Type) | Target type is not a record type. |
CannotCreateMissingClause QName (Telescope, [NamedArg DeBruijnPattern]) Doc (Closure (Abs Type)) | |
GenericSplitError String |
Instances
PrettyTCM SplitError Source # | |||||
Defined in Agda.TypeChecking.Errors prettyTCM :: MonadPretty m => SplitError -> m Doc Source # | |||||
NFData SplitError Source # | |||||
Defined in Agda.TypeChecking.Monad.Base rnf :: SplitError -> () | |||||
Generic SplitError Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
from :: SplitError -> Rep SplitError x to :: Rep SplitError x -> SplitError | |||||
Show SplitError Source # | |||||
Defined in Agda.TypeChecking.Monad.Base showsPrec :: Int -> SplitError -> ShowS show :: SplitError -> String showList :: [SplitError] -> ShowS | |||||
type Rep SplitError Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep SplitError = D1 ('MetaData "SplitError" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (((C1 ('MetaCons "NotADatatype" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Closure Type))) :+: C1 ('MetaCons "BlockedType" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Blocker) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Closure Type)))) :+: (C1 ('MetaCons "ErasedDatatype" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ErasedDatatypeReason) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Closure Type))) :+: (C1 ('MetaCons "CoinductiveDatatype" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Closure Type))) :+: C1 ('MetaCons "UnificationStuck" 'PrefixI 'True) ((S1 ('MetaSel ('Just "cantSplitBlocker") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe Blocker)) :*: (S1 ('MetaSel ('Just "cantSplitConName") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Just "cantSplitTel") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Telescope))) :*: (S1 ('MetaSel ('Just "cantSplitConIdx") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Args) :*: (S1 ('MetaSel ('Just "cantSplitGivenIdx") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Args) :*: S1 ('MetaSel ('Just "cantSplitFailures") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [UnificationFailure]))))))) :+: ((C1 ('MetaCons "CosplitCatchall" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "CosplitNoTarget" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "CosplitNoRecordType" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Closure Type))) :+: (C1 ('MetaCons "CannotCreateMissingClause" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Telescope, [NamedArg DeBruijnPattern]))) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Doc) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Closure (Abs Type))))) :+: C1 ('MetaCons "GenericSplitError" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String)))))) |
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
PrettyTCM UnificationFailure Source # | |||||
Defined in Agda.TypeChecking.Errors prettyTCM :: MonadPretty m => UnificationFailure -> m Doc Source # | |||||
NFData UnificationFailure Source # | |||||
Defined in Agda.TypeChecking.Monad.Base rnf :: UnificationFailure -> () | |||||
Generic UnificationFailure Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
from :: UnificationFailure -> Rep UnificationFailure x to :: Rep UnificationFailure x -> UnificationFailure | |||||
Show UnificationFailure Source # | |||||
Defined in Agda.TypeChecking.Monad.Base showsPrec :: Int -> UnificationFailure -> ShowS show :: UnificationFailure -> String showList :: [UnificationFailure] -> ShowS | |||||
type Rep UnificationFailure Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep UnificationFailure = D1 ('MetaData "UnificationFailure" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) ((C1 ('MetaCons "UnifyIndicesNotVars" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Telescope) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Args)))) :+: C1 ('MetaCons "UnifyRecursiveEq" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Telescope) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Int) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)))) :+: (C1 ('MetaCons "UnifyReflexiveEq" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Telescope) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term))) :+: C1 ('MetaCons "UnifyUnusableModality" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Telescope) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Int) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Modality)))))) |
data NegativeUnification Source #
Instances
PrettyTCM NegativeUnification Source # | |||||
Defined in Agda.TypeChecking.Errors prettyTCM :: MonadPretty m => NegativeUnification -> m Doc Source # | |||||
NFData NegativeUnification Source # | |||||
Defined in Agda.TypeChecking.Monad.Base rnf :: NegativeUnification -> () | |||||
Generic NegativeUnification Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
from :: NegativeUnification -> Rep NegativeUnification x to :: Rep NegativeUnification x -> NegativeUnification | |||||
Show NegativeUnification Source # | |||||
Defined in Agda.TypeChecking.Monad.Base showsPrec :: Int -> NegativeUnification -> ShowS show :: NegativeUnification -> String showList :: [NegativeUnification] -> ShowS | |||||
type Rep NegativeUnification Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep NegativeUnification = D1 ('MetaData "NegativeUnification" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "UnifyConflict" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Telescope) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term))) :+: C1 ('MetaCons "UnifyCycle" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Telescope) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Int) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)))) |
data UnquoteError Source #
BadVisibility String (Arg Term) | |
ConInsteadOfDef QName String String | |
DefInsteadOfCon QName String String | |
NonCanonical String Term | |
BlockedOnMeta TCState Blocker | |
PatLamWithoutClauses Term | |
UnquotePanic String |
Instances
NFData UnquoteError Source # | |||||
Defined in Agda.TypeChecking.Monad.Base rnf :: UnquoteError -> () | |||||
Generic UnquoteError Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
from :: UnquoteError -> Rep UnquoteError x to :: Rep UnquoteError x -> UnquoteError | |||||
Show UnquoteError Source # | |||||
Defined in Agda.TypeChecking.Monad.Base showsPrec :: Int -> UnquoteError -> ShowS show :: UnquoteError -> String showList :: [UnquoteError] -> ShowS | |||||
type Rep UnquoteError Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep UnquoteError = D1 ('MetaData "UnquoteError" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) ((C1 ('MetaCons "BadVisibility" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Arg Term))) :+: (C1 ('MetaCons "ConInsteadOfDef" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String))) :+: C1 ('MetaCons "DefInsteadOfCon" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String))))) :+: ((C1 ('MetaCons "NonCanonical" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)) :+: C1 ('MetaCons "BlockedOnMeta" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 TCState) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Blocker))) :+: (C1 ('MetaCons "PatLamWithoutClauses" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)) :+: C1 ('MetaCons "UnquotePanic" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String))))) |
data IncorrectTypeForRewriteRelationReason Source #
Instances
NFData IncorrectTypeForRewriteRelationReason Source # | |||||
Defined in Agda.TypeChecking.Monad.Base rnf :: IncorrectTypeForRewriteRelationReason -> () | |||||
Generic IncorrectTypeForRewriteRelationReason Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
| |||||
Show IncorrectTypeForRewriteRelationReason Source # | |||||
Defined in Agda.TypeChecking.Monad.Base showsPrec :: Int -> IncorrectTypeForRewriteRelationReason -> ShowS show :: IncorrectTypeForRewriteRelationReason -> String showList :: [IncorrectTypeForRewriteRelationReason] -> ShowS | |||||
type Rep IncorrectTypeForRewriteRelationReason Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep IncorrectTypeForRewriteRelationReason = D1 ('MetaData "IncorrectTypeForRewriteRelationReason" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "ShouldAcceptAtLeastTwoArguments" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "FinalTwoArgumentsNotVisible" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "TypeDoesNotEndInSort" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Telescope)))) |
data InvalidFileNameReason Source #
Extra information for InvalidFileName
error.
Instances
NFData InvalidFileNameReason Source # | |||||
Defined in Agda.TypeChecking.Monad.Base rnf :: InvalidFileNameReason -> () | |||||
Generic InvalidFileNameReason Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
from :: InvalidFileNameReason -> Rep InvalidFileNameReason x to :: Rep InvalidFileNameReason x -> InvalidFileNameReason | |||||
Show InvalidFileNameReason Source # | |||||
Defined in Agda.TypeChecking.Monad.Base showsPrec :: Int -> InvalidFileNameReason -> ShowS show :: InvalidFileNameReason -> String showList :: [InvalidFileNameReason] -> ShowS | |||||
type Rep InvalidFileNameReason Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep InvalidFileNameReason = D1 ('MetaData "InvalidFileNameReason" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "DoesNotCorrespondToValidModuleName" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "RootNameModuleNotAQualifiedModuleName" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Text))) |
data InductionAndEta Source #
Instances
NFData DataOrRecordE Source # | |||||
Defined in Agda.TypeChecking.Monad.Base rnf :: DataOrRecordE -> () | |||||
NFData InductionAndEta Source # | |||||
Defined in Agda.TypeChecking.Monad.Base rnf :: InductionAndEta -> () | |||||
Generic InductionAndEta Source # | |||||
Defined in Agda.TypeChecking.Monad.Base
from :: InductionAndEta -> Rep InductionAndEta x to :: Rep InductionAndEta x -> InductionAndEta | |||||
Show InductionAndEta Source # | |||||
Defined in Agda.TypeChecking.Monad.Base showsPrec :: Int -> InductionAndEta -> ShowS show :: InductionAndEta -> String showList :: [InductionAndEta] -> ShowS | |||||
type Rep InductionAndEta Source # | |||||
Defined in Agda.TypeChecking.Monad.Base type Rep InductionAndEta = D1 ('MetaData "InductionAndEta" "Agda.TypeChecking.Monad.Base" "Agda-2.6.20240714-inplace" 'False) (C1 ('MetaCons "InductionAndEta" 'PrefixI 'True) (S1 ('MetaSel ('Just "recordInduction") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe Induction)) :*: S1 ('MetaSel ('Just "recordEtaEquality") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 EtaEquality))) |
sizedTypesOption :: HasOptions m => m Bool Source #
guardednessOption :: HasOptions m => m Bool Source #
withoutKOption :: HasOptions m => m Bool Source #
cubicalOption :: HasOptions m => m (Maybe Cubical) Source #
cubicalCompatibleOption :: HasOptions m => m Bool Source #
enableCaching :: HasOptions m => m Bool Source #
Environment of the reduce monad.
fmapReduce :: (a -> b) -> ReduceM a -> ReduceM b Source #
runReduceF :: (a -> ReduceM b) -> TCM (a -> b) Source #
newtype BlockT (m :: Type -> Type) a Source #
Instances
MonadTrans BlockT Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
HasOptions m => HasOptions (BlockT m) Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
Monad m => MonadBlock (BlockT m) Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
MonadReduce m => MonadReduce (BlockT m) Source # | |
Defined in Agda.TypeChecking.Monad.Base liftReduce :: ReduceM a -> BlockT m a Source # | |
MonadTCEnv m => MonadTCEnv (BlockT m) Source # | |
MonadTCM m => MonadTCM (BlockT m) Source # | |
MonadTCState m => MonadTCState (BlockT m) Source # | |
ReadTCState m => ReadTCState (BlockT m) Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
HasBuiltins m => HasBuiltins (BlockT m) Source # | |
Defined in Agda.TypeChecking.Monad.Builtin getBuiltinThing :: SomeBuiltin -> BlockT m (Maybe (Builtin PrimFun)) Source # | |
MonadAddContext m => MonadAddContext (BlockT m) Source # | |
Defined in Agda.TypeChecking.Monad.Context addCtx :: Name -> Dom Type -> BlockT m a -> BlockT m a Source # addLetBinding' :: Origin -> Name -> Term -> Dom Type -> BlockT m a -> BlockT m a Source # updateContext :: Substitution -> (Context -> Context) -> BlockT m a -> BlockT m a Source # withFreshName :: Range -> ArgName -> (Name -> BlockT m a) -> BlockT m a Source # | |
MonadDebug m => MonadDebug (BlockT m) Source # | |
Defined in Agda.TypeChecking.Monad.Debug formatDebugMessage :: VerboseKey -> VerboseLevel -> TCM Doc -> BlockT m String Source # traceDebugMessage :: VerboseKey -> VerboseLevel -> String -> BlockT m a -> BlockT m a Source # verboseBracket :: VerboseKey -> VerboseLevel -> String -> BlockT m a -> BlockT m a Source # getVerbosity :: BlockT m Verbosity Source # getProfileOptions :: BlockT m ProfileOptions Source # isDebugPrinting :: BlockT m Bool Source # nowDebugPrinting :: BlockT m a -> BlockT m a Source # | |
PureTCM m => PureTCM (BlockT m) Source # | |
Defined in Agda.TypeChecking.Monad.Pure | |
HasConstInfo m => HasConstInfo (BlockT m) Source # | |
Defined in Agda.TypeChecking.Monad.Signature getConstInfo :: QName -> BlockT m Definition Source # getConstInfo' :: QName -> BlockT m (Either SigError Definition) Source # getRewriteRulesFor :: QName -> BlockT m RewriteRules Source # | |
MonadIO m => MonadIO (BlockT m) Source # | |
Defined in Agda.TypeChecking.Monad.Base | |
Monad m => Applicative (BlockT m) Source # | |
Functor m => Functor (BlockT m) Source # | |
Monad m => Monad (BlockT m) Source # | |
MonadFail m => MonadFail (BlockT m) Source # | |
Defined in Agda.TypeChecking.Monad.Base |
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.
setTCLens :: MonadTCState m => Lens' TCState a -> a -> m () infix 4 Source #
Overwrite the part of the TCState
focused on by the lens.
setTCLens' :: MonadTCState m => Lens' TCState a -> a -> m () Source #
Overwrite the part of the TCState
focused on by the lens
(strictly).
modifyTCLens :: MonadTCState m => Lens' TCState a -> (a -> a) -> m () Source #
Modify the part of the TCState
focused on by the lens.
modifyTCLens' :: MonadTCState m => Lens' TCState a -> (a -> a) -> m () Source #
Modify the part of the TCState
focused on by the lens
(strictly).
modifyTCLensM :: MonadTCState m => Lens' TCState a -> (a -> m a) -> m () Source #
Modify a part of the state monadically.
stateTCLens :: MonadTCState m => Lens' TCState a -> (a -> (r, a)) -> m r Source #
Modify the part of the TCState
focused on by the lens, and return some result.
runBlocked :: Monad m => BlockT m a -> m (Either Blocker a) Source #
returnTCMT :: forall (m :: Type -> Type) a. Applicative m => a -> TCMT m a Source #
bindTCMT :: forall (m :: Type -> Type) a b. Monad m => TCMT m a -> (a -> TCMT m b) -> TCMT m b Source #
thenTCMT :: forall (m :: Type -> Type) a b. Applicative m => TCMT m a -> TCMT m b -> TCMT m b Source #
apTCMT :: forall (m :: Type -> Type) a b. Applicative m => TCMT m (a -> b) -> TCMT m a -> TCMT m b Source #
internalError :: (HasCallStack, MonadTCM tcm) => String -> tcm 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.
typeError' :: MonadTCError m => CallStack -> TypeError -> m a Source #
locatedTypeError :: MonadTCError m => (a -> TypeError) -> HasCallStack => a -> m b Source #
Utility function for 1-arg constructed type errors.
Note that the HasCallStack
constraint is on the *resulting* function.
genericDocError :: (HasCallStack, MonadTCError m) => Doc -> m a Source #
typeError'_ :: (MonadTCEnv m, ReadTCState m) => CallStack -> TypeError -> m TCErr Source #
typeError_ :: (HasCallStack, 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, except that it might raise IOException
s) 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.
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
doExpandLast :: TCM a -> TCM a Source #
Restore setting for ExpandLast
to default.
workOnTypes :: (MonadTCEnv m, HasOptions m, MonadDebug m) => m a -> m a Source #
Modify the context whenever going from the l.h.s. (term side) of the typing judgement to the r.h.s. (type side).
withCurrentModule :: MonadTCEnv m => ModuleName -> m a -> m a Source #
Set the name of the current module.
getAnonymousVariables :: MonadTCEnv m => ModuleName -> m Nat Source #
Get the number of variables bound by anonymous modules.
withAnonymousModule :: ModuleName -> Nat -> TCM a -> TCM a Source #
Add variables bound by an anonymous module.
withEnv :: MonadTCEnv m => TCEnv -> m a -> m a Source #
Set the current environment to the given
withHighlightingLevel :: HighlightingLevel -> TCM a -> TCM a Source #
Set highlighting level
dontExpandLast :: TCM a -> TCM a Source #
reallyDontExpandLast :: TCM a -> TCM a Source #
performedSimplification :: MonadTCEnv m => m a -> m a Source #
If the reduced did a proper match (constructor or literal pattern), then record this as simplification step.
performedSimplification' :: MonadTCEnv m => Simplification -> m a -> m a Source #
getSimplification :: MonadTCEnv m => m Simplification Source #
updateAllowedReductions :: (AllowedReductions -> AllowedReductions) -> TCEnv -> TCEnv Source #
Lens for AllowedReductions
.
modifyAllowedReductions :: MonadTCEnv m => (AllowedReductions -> AllowedReductions) -> m a -> m a Source #
putAllowedReductions :: MonadTCEnv m => AllowedReductions -> m a -> m a Source #
onlyReduceProjections :: MonadTCEnv m => m a -> m a Source #
Reduce Def f vs
only if f
is a projection.
allowAllReductions :: MonadTCEnv m => m a -> m a Source #
Allow all reductions except for non-terminating functions (default).
allowNonTerminatingReductions :: MonadTCEnv m => m a -> m a Source #
Allow all reductions including non-terminating functions.
onlyReduceTypes :: MonadTCEnv m => m a -> m a Source #
Allow all reductions when reducing types. Otherwise only allow inlined functions to be unfolded.
typeLevelReductions :: MonadTCEnv m => m a -> m a Source #
Update allowed reductions when working on types
callByName :: TCM a -> TCM a Source #
Don't use call-by-need evaluation for the given computation.
dontFoldLetBindings :: MonadTCEnv m => m a -> m a Source #
Don't fold let bindings when printing. This is a bit crude since it disables any folding of let
bindings at all. In many cases it's better to use removeLetBinding
before printing to drop
the let bindings that should not be folded.
removeLetBinding :: MonadTCEnv m => Name -> m a -> m a Source #
Remove a let bound variable.
newtype BuiltinAccess a Source #
The trivial implementation of HasBuiltins
, using a constant TCState
.
This may be used instead of TCMT
/ReduceM
where builtins must be accessed
in a pure context.
BuiltinAccess | |
|
Instances
HasBuiltins BuiltinAccess Source # | |
Defined in Agda.TypeChecking.Monad.Builtin getBuiltinThing :: SomeBuiltin -> BuiltinAccess (Maybe (Builtin PrimFun)) Source # | |
Applicative BuiltinAccess Source # | |
Defined in Agda.TypeChecking.Monad.Builtin pure :: a -> BuiltinAccess a (<*>) :: BuiltinAccess (a -> b) -> BuiltinAccess a -> BuiltinAccess b # liftA2 :: (a -> b -> c) -> BuiltinAccess a -> BuiltinAccess b -> BuiltinAccess c (*>) :: BuiltinAccess a -> BuiltinAccess b -> BuiltinAccess b (<*) :: BuiltinAccess a -> BuiltinAccess b -> BuiltinAccess a | |
Functor BuiltinAccess Source # | |
Defined in Agda.TypeChecking.Monad.Builtin fmap :: (a -> b) -> BuiltinAccess a -> BuiltinAccess b (<$) :: a -> BuiltinAccess b -> BuiltinAccess a # | |
Monad BuiltinAccess Source # | |
Defined in Agda.TypeChecking.Monad.Builtin (>>=) :: BuiltinAccess a -> (a -> BuiltinAccess b) -> BuiltinAccess b (>>) :: BuiltinAccess a -> BuiltinAccess b -> BuiltinAccess b return :: a -> BuiltinAccess a | |
MonadFail BuiltinAccess Source # | |
Defined in Agda.TypeChecking.Monad.Builtin fail :: String -> BuiltinAccess a |
runBuiltinAccess :: TCState -> BuiltinAccess a -> a Source #
Run a BuiltinAccess
monad.
litType :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => Literal -> m Type Source #
primZero :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primSuc :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primNat :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primWord64 :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primFloat :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primChar :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primString :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primQName :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaMeta :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
setBuiltinThings :: BuiltinThings PrimFun -> TCM () Source #
bindPrimitive :: PrimitiveId -> PrimFun -> TCM () Source #
bindBuiltinRewriteRelation :: QName -> TCM () Source #
Add one (more) relation symbol to the rewrite relations.
getBuiltinRewriteRelations :: (HasBuiltins m, MonadTCError m) => m (Set QName) Source #
Get the currently registered rewrite relation symbols.
getBuiltinRewriteRelations' :: HasBuiltins m => m (Maybe (Set QName)) Source #
Get the currently registered rewrite relation symbols, if any.
getBuiltin :: (HasBuiltins m, MonadTCError m) => BuiltinId -> m Term Source #
getPrimitive' :: HasBuiltins m => PrimitiveId -> m (Maybe PrimFun) Source #
getPrimitive :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => PrimitiveId -> m PrimFun Source #
getPrimitiveTerm :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => PrimitiveId -> m Term Source #
getPrimitiveTerm' :: HasBuiltins m => PrimitiveId -> m (Maybe Term) Source #
getPrimitiveName' :: HasBuiltins m => PrimitiveId -> m (Maybe QName) Source #
getTerm :: (HasBuiltins m, IsBuiltin a) => String -> a -> m Term Source #
getTerm use name
looks up name
as a primitive or builtin, and
throws an error otherwise.
The use
argument describes how the name is used for the sake of
the error message.
primInteger :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primIntegerPos :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primIntegerNegSuc :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primUnit :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primUnitUnit :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primBool :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primTrue :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primFalse :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primSigma :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primList :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primNil :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primCons :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primIO :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primMaybe :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primNothing :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primJust :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primPath :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primPathP :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primIntervalUniv :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primInterval :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primIZero :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primIOne :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primIsOne :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primItIsOne :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primIsOne1 :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primIsOne2 :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primIsOneEmpty :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primSub :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primSubIn :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primTrans :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primId :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primEquiv :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primEquivFun :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primEquivProof :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primTranspProof :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
prim_glue :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
prim_unglue :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
prim_glueU :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
prim_unglueU :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primSizeUniv :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primSize :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primSizeLt :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primSizeSuc :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primSizeInf :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primSizeMax :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primInf :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primSharp :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primFlat :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primEquality :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primRefl :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primLevel :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primLevelUniv :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primProp :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primSet :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primStrictSet :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primPropOmega :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primSetOmega :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primSSetOmega :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primFromNat :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primFromNeg :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primFromString :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primArgInfo :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primArgArgInfo :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primArg :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primArgArg :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAbs :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAbsAbs :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaTerm :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaTermVar :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaTermLam :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaTermExtLam :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaTermDef :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaTermCon :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaTermPi :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaTermSort :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaTermLit :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaTermUnsupported :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaTermMeta :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaErrorPart :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaErrorPartString :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaErrorPartTerm :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaErrorPartPatt :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaErrorPartName :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primHiding :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primHidden :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primInstance :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primVisible :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primRelevance :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primRelevant :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primIrrelevant :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primQuantity :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primQuantity0 :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primQuantityω :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primModality :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primModalityConstructor :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAssoc :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAssocLeft :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAssocRight :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAssocNon :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primPrecedence :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primPrecRelated :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primPrecUnrelated :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primFixity :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primFixityFixity :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaLiteral :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaLitNat :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaLitWord64 :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaLitFloat :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaLitString :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaLitChar :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaLitQName :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaLitMeta :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaSort :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaSortSet :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaSortLit :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaSortProp :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaSortPropLit :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaSortInf :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaSortUnsupported :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaDefinition :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaDefinitionFunDef :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaDefinitionDataDef :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaDefinitionRecordDef :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaDefinitionPostulate :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaDefinitionPrimitive :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaDefinitionDataConstructor :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaClause :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaClauseClause :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaClauseAbsurd :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaPattern :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaPatCon :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaPatVar :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaPatDot :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaPatLit :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaPatProj :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaPatAbsurd :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaBlocker :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaBlockerAny :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaBlockerAll :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaBlockerMeta :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaTCM :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaTCMReturn :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaTCMBind :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaTCMUnify :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaTCMTypeError :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaTCMInferType :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaTCMCheckType :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaTCMNormalise :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaTCMReduce :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaTCMCatchError :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaTCMGetContext :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaTCMExtendContext :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaTCMInContext :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaTCMFreshName :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaTCMDeclareDef :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaTCMDeclarePostulate :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaTCMDeclareData :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaTCMDefineData :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaTCMDefineFun :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaTCMGetType :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaTCMGetDefinition :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaTCMQuoteTerm :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaTCMUnquoteTerm :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaTCMQuoteOmegaTerm :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaTCMCommit :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaTCMIsMacro :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaTCMBlock :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaTCMFormatErrorParts :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaTCMDebugPrint :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaTCMWithNormalisation :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaTCMWithReconstructed :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaTCMWithExpandLast :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaTCMWithReduceDefs :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaTCMAskNormalisation :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaTCMAskReconstructed :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaTCMAskExpandLast :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaTCMAskReduceDefs :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaTCMNoConstraints :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaTCMWorkOnTypes :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaTCMRunSpeculative :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaTCMExec :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaTCMGetInstances :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaTCMSolveInstances :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaTCMPragmaForeign :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
primAgdaTCMPragmaCompile :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => m Term Source #
data CoinductionKit Source #
The coinductive primitives.
CoinductionKit | |
|
coinductionKit' :: TCM CoinductionKit Source #
Tries to build a CoinductionKit
.
Sort primitives.
SortKit | |
|
sortKit :: (HasBuiltins m, MonadTCError m, HasOptions m) => m SortKit Source #
infallibleSortKit :: HasBuiltins m => m SortKit Source #
Compute a SortKit
in contexts that do not support failure (e.g.
Reify
). This should only be used when we are sure that the
primitive sorts have been bound, i.e. because it is "after" type
checking.
getPrimName :: Term -> QName Source #
isPrimitive :: HasBuiltins m => PrimitiveId -> QName -> m Bool Source #
intervalSort :: Sort Source #
intervalView' :: HasBuiltins m => m (Term -> IntervalView) Source #
intervalView :: HasBuiltins m => Term -> m IntervalView Source #
intervalUnview :: HasBuiltins m => IntervalView -> m Term Source #
intervalUnview' :: HasBuiltins m => m (IntervalView -> Term) Source #
pathView :: HasBuiltins m => Type -> m PathView Source #
Check whether the type is actually an path (lhs ≡ rhs) and extract lhs, rhs, and their type.
Precondition: type is reduced.
idViewAsPath :: HasBuiltins m => Type -> m PathView Source #
Non dependent Path
boldPathView :: Type -> PathView Source #
primEqualityName :: TCM QName Source #
Get the name of the equality type.
equalityView :: Type -> TCM EqualityView Source #
Check whether the type is actually an equality (lhs ≡ rhs) and extract lhs, rhs, and their type.
Precondition: type is reduced.
class EqualityUnview a where Source #
Revert the EqualityView
.
Postcondition: type is reduced.
equalityUnview :: a -> Type Source #
Instances
constrainedPrims :: [PrimitiveId] Source #
Primitives with typechecking constrants.
getNameOfConstrained :: HasBuiltins m => PrimitiveId -> m (Maybe QName) Source #
isInteractionMeta :: ReadTCState m => MetaId -> m (Maybe InteractionId) Source #
Does the meta variable correspond to an interaction point?
Time: O(log n)
where n
is the number of interaction metas.
notSoPrettySigCubicalNotErasure :: QName -> String Source #
Generates an error message corresponding to
SigCubicalNotErasure
for a given QName
.
inFreshModuleIfFreeParams :: TCM a -> TCM a Source #
Unless all variables in the context are module parameters, create a fresh module to capture the non-module parameters. Used when unquoting to make sure generated definitions work properly.
lookupSection :: (Functor m, ReadTCState m) => ModuleName -> m Telescope Source #
Lookup a section telescope.
If it doesn't exist, like in hierarchical top-level modules, the section telescope is empty.
defaultGetVerbosity :: HasOptions m => m Verbosity Source #
defaultGetProfileOptions :: HasOptions m => m ProfileOptions Source #
defaultIsDebugPrinting :: MonadTCEnv m => m Bool Source #
defaultNowDebugPrinting :: MonadTCEnv m => m a -> m a Source #
displayDebugMessage :: MonadDebug m => VerboseKey -> VerboseLevel -> String -> m () Source #
Print a debug message if switched on.
catchAndPrintImpossible :: (CatchImpossible m, Monad m) => VerboseKey -> VerboseLevel -> m String -> m String Source #
During printing, catch internal errors of kind Impossible
and print them.
hasVerbosity :: MonadDebug m => VerboseKey -> VerboseLevel -> m Bool Source #
Check whether a certain verbosity level is activated.
Precondition: The level must be non-negative.
alwaysReportSLn :: MonadDebug m => VerboseKey -> VerboseLevel -> String -> m () Source #
Conditionally println debug string. Works regardless of the debug flag.
reportResult :: MonadDebug m => VerboseKey -> VerboseLevel -> (a -> TCM Doc) -> m a -> m a Source #
Debug print the result of a computation.
alwaysReportSDoc :: MonadDebug m => VerboseKey -> VerboseLevel -> TCM Doc -> m () Source #
Conditionally render debug Doc
and print it. Works regardless of the debug flag.
unlessDebugPrinting :: MonadDebug m => m () -> m () Source #
Debug print some lines if the verbosity level for the given
VerboseKey
is at least VerboseLevel
.
Note: In the presence of OverloadedStrings
, just
@
traceS key level "Literate string"
gives an
Ambiguous type variable error in
GHC@.
Use the legacy functions traceSLn
and traceSDoc
instead then.
traceS :: MonadDebug m => VerboseKey -> VerboseLevel -> a -> m c -> m c Source #
Instances
TraceS Doc Source # | |
Defined in Agda.TypeChecking.Monad.Debug traceS :: MonadDebug m => VerboseKey -> VerboseLevel -> Doc -> m c -> m c Source # | |
TraceS String Source # | |
Defined in Agda.TypeChecking.Monad.Debug traceS :: MonadDebug m => VerboseKey -> VerboseLevel -> String -> m c -> m c Source # | |
TraceS (TCM Doc) Source # | |
Defined in Agda.TypeChecking.Monad.Debug traceS :: MonadDebug m => VerboseKey -> VerboseLevel -> TCM Doc -> m c -> m c Source # | |
TraceS [Doc] Source # | |
Defined in Agda.TypeChecking.Monad.Debug traceS :: MonadDebug m => VerboseKey -> VerboseLevel -> [Doc] -> m c -> m c Source # | |
TraceS [TCM Doc] Source # | |
Defined in Agda.TypeChecking.Monad.Debug traceS :: MonadDebug m => VerboseKey -> VerboseLevel -> [TCM Doc] -> m c -> m c Source # | |
TraceS [String] Source # | |
Defined in Agda.TypeChecking.Monad.Debug traceS :: MonadDebug m => VerboseKey -> VerboseLevel -> [String] -> m c -> m c Source # |
traceSDoc :: MonadDebug m => VerboseKey -> VerboseLevel -> TCM Doc -> m a -> m a Source #
Conditionally render debug Doc
, print it, and then continue.
traceSLn :: MonadDebug m => VerboseKey -> VerboseLevel -> String -> m a -> m a Source #
openVerboseBracket :: MonadDebug m => VerboseKey -> VerboseLevel -> String -> m () Source #
closeVerboseBracket :: MonadDebug m => VerboseKey -> VerboseLevel -> m () Source #
closeVerboseBracketException :: MonadDebug m => VerboseKey -> VerboseLevel -> m () Source #
hasExactVerbosity :: MonadDebug m => VerboseKey -> VerboseLevel -> m Bool Source #
Check whether a certain verbosity level is activated (exact match).
whenExactVerbosity :: MonadDebug m => VerboseKey -> VerboseLevel -> m () -> m () Source #
Run a computation if a certain verbosity level is activated (exact match).
__CRASH_WHEN__ :: (HasCallStack, MonadTCM m, MonadDebug m) => VerboseKey -> VerboseLevel -> m () Source #
hasProfileOption :: MonadDebug m => ProfileOption -> m Bool Source #
Check whether a certain profile option is activated.
whenProfile :: MonadDebug m => ProfileOption -> m () -> m () Source #
Run some code when the given profiling option is active.
resetAllState :: TCM () Source #
Resets all of the type checking state.
Keep only Benchmark
and backend information.
updatePersistentState :: (PersistentTCState -> PersistentTCState) -> TCState -> TCState Source #
localTCStateSaving :: TCM a -> TCM (a, TCState) Source #
Same as localTCState
but also returns the state in which we were just
before reverting it.
localTCStateSavingWarnings :: TCM a -> TCM a Source #
Same as localTCState
but keep all warnings.
speculateTCState :: TCM (a, SpeculateResult) -> TCM a Source #
Allow rolling back the state changes of a TCM computation.
speculateTCState_ :: TCM SpeculateResult -> TCM () Source #
freshTCM :: TCM a -> TCM (Either TCErr a) Source #
A fresh TCM instance.
The computation is run in a fresh state, with the exception that the persistent state is preserved. If the computation changes the state, then these changes are ignored, except for changes to the persistent state. (Changes to the persistent state are also ignored if errors other than type errors or IO exceptions are encountered.)
modifyPersistentState :: (PersistentTCState -> PersistentTCState) -> TCM () Source #
modifyScope :: MonadTCState m => (ScopeInfo -> ScopeInfo) -> m () Source #
Modify the current scope.
modifyScope_ :: MonadTCState m => (ScopeInfo -> ScopeInfo) -> m () Source #
Modify the current scope without updating the inverse maps.
locallyScope :: ReadTCState m => Lens' ScopeInfo a -> (a -> a) -> m b -> m b Source #
Run a computation in a modified scope.
withScope_ :: ReadTCState m => ScopeInfo -> m a -> m a Source #
Same as withScope
, but discard the scope from the computation.
localScope :: TCM a -> TCM a Source #
Discard any changes to the scope by a computation.
notInScopeError :: QName -> TCM a Source #
Scope error.
printScope :: String -> Int -> String -> TCM () Source #
Debug print the scope.
notInScopeWarning :: QName -> TCM () Source #
modifySignature :: MonadTCState m => (Signature -> Signature) -> m () Source #
modifyImportedSignature :: MonadTCState m => (Signature -> Signature) -> m () Source #
getSignature :: ReadTCState m => m Signature Source #
modifyGlobalDefinition :: MonadTCState m => QName -> (Definition -> Definition) -> m () Source #
Update a possibly imported definition. Warning: changes made to imported definitions (during type checking) will not persist outside the current module. This function is currently used to update the compiled representation of a function during compilation.
updateDefinition :: QName -> (Definition -> Definition) -> Signature -> Signature Source #
setSignature :: MonadTCState m => Signature -> m () Source #
withSignature :: (ReadTCState m, MonadTCState m) => Signature -> m a -> m a Source #
Run some computation in a different signature, restore original signature.
addRewriteRulesFor :: QName -> RewriteRules -> [QName] -> Signature -> Signature Source #
updateTheDef :: (Defn -> Defn) -> Definition -> Definition Source #
updateDefCopatternLHS :: (Bool -> Bool) -> Definition -> Definition Source #
modifyRecEta :: MonadTCState m => QName -> (EtaEquality -> EtaEquality) -> m () Source #
lookupDefinition :: QName -> Signature -> Maybe Definition Source #
updateDefinitions :: (Definitions -> Definitions) -> Signature -> Signature Source #
updateDefType :: (Type -> Type) -> Definition -> Definition Source #
updateDefArgOccurrences :: ([Occurrence] -> [Occurrence]) -> Definition -> Definition Source #
updateDefPolarity :: ([Polarity] -> [Polarity]) -> Definition -> Definition Source #
updateDefCompiledRep :: (CompiledRepresentation -> CompiledRepresentation) -> Definition -> Definition Source #
addCompilerPragma :: BackendName -> CompilerPragma -> Definition -> Definition Source #
updateCompiledClauses :: (Maybe CompiledClauses -> Maybe CompiledClauses) -> Defn -> Defn Source #
updateDefBlocked :: (Blocked_ -> Blocked_) -> Definition -> Definition Source #
setTopLevelModule :: TopLevelModuleName -> TCM () Source #
Set the top-level module. This affects the global module id of freshly generated names.
currentTopLevelModule :: (MonadTCEnv m, ReadTCState m) => m (Maybe TopLevelModuleName) Source #
The name of the current top-level module, if any.
withTopLevelModule :: TopLevelModuleName -> TCM a -> TCM a Source #
Use a different top-level module for a computation. Used when generating names for imported modules.
addForeignCode :: BackendName -> String -> TCM () Source #
appInteractionOutputCallback :: Response -> TCM () Source #
getPatternSyns :: ReadTCState m => m PatternSynDefns Source #
setPatternSyns :: PatternSynDefns -> TCM () Source #
modifyPatternSyns :: (PatternSynDefns -> PatternSynDefns) -> TCM () Source #
Lens for stPatternSyns
.
getPatternSynImports :: ReadTCState m => m PatternSynDefns Source #
updateInstanceDefs :: (TempInstanceTable -> TempInstanceTable) -> TCState -> TCState Source #
Lens for stInstanceDefs
.
modifyInstanceDefs :: (TempInstanceTable -> TempInstanceTable) -> TCM () Source #
getAnonInstanceDefs :: TCM (Set QName) Source #
clearUnknownInstance :: QName -> TCM () Source #
Remove an instance from the set of unresolved instances.
addUnknownInstance :: QName -> TCM () Source #
Add an instance whose type is still unresolved.
getCurrentRange :: MonadTCEnv m => m Range Source #
data BoundedSize Source #
Result of querying whether size variable i
is bounded by another
size.
Instances
Show BoundedSize Source # | |
Defined in Agda.TypeChecking.Monad.SizedTypes showsPrec :: Int -> BoundedSize -> ShowS show :: BoundedSize -> String showList :: [BoundedSize] -> ShowS | |
Eq BoundedSize Source # | |
Defined in Agda.TypeChecking.Monad.SizedTypes (==) :: BoundedSize -> BoundedSize -> Bool (/=) :: BoundedSize -> BoundedSize -> Bool |
class IsSizeType a where Source #
Check if a type is the primSize
type. The argument should be reduce
d.
isSizeType :: (HasOptions m, HasBuiltins m) => a -> m (Maybe BoundedSize) Source #
Instances
IsSizeType Term Source # | |
Defined in Agda.TypeChecking.Monad.SizedTypes isSizeType :: (HasOptions m, HasBuiltins m) => Term -> m (Maybe BoundedSize) Source # | |
IsSizeType CompareAs Source # | |
Defined in Agda.TypeChecking.Monad.SizedTypes isSizeType :: (HasOptions m, HasBuiltins m) => CompareAs -> m (Maybe BoundedSize) Source # | |
IsSizeType a => IsSizeType (Dom a) Source # | |
Defined in Agda.TypeChecking.Monad.SizedTypes isSizeType :: (HasOptions m, HasBuiltins m) => Dom a -> m (Maybe BoundedSize) Source # | |
IsSizeType a => IsSizeType (Type' a) Source # | |
Defined in Agda.TypeChecking.Monad.SizedTypes isSizeType :: (HasOptions m, HasBuiltins m) => Type' a -> m (Maybe BoundedSize) Source # | |
IsSizeType a => IsSizeType (b, a) Source # | |
Defined in Agda.TypeChecking.Monad.SizedTypes isSizeType :: (HasOptions m, HasBuiltins m) => (b, a) -> m (Maybe BoundedSize) Source # |
isSizeTypeTest :: (HasOptions m, HasBuiltins m) => m (Term -> Maybe BoundedSize) Source #
getBuiltinSize :: HasBuiltins m => m (Maybe QName, Maybe QName) Source #
getBuiltinDefName :: HasBuiltins m => BuiltinId -> m (Maybe QName) Source #
isSizeNameTest :: (HasOptions m, HasBuiltins m) => m (QName -> Bool) Source #
isSizeNameTestRaw :: (HasOptions m, HasBuiltins m) => m (QName -> Bool) Source #
haveSizedTypes :: TCM Bool Source #
Test whether OPTIONS --sized-types and whether the size built-ins are defined.
haveSizeLt :: TCM Bool Source #
Test whether the SIZELT builtin is defined.
builtinSizeHook :: BuiltinId -> QName -> Type -> TCM () Source #
Add polarity info to a SIZE builtin.
sizeSucName :: (HasBuiltins m, HasOptions m) => m (Maybe QName) Source #
The name of SIZESUC
.
sizeMax :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => List1 Term -> m Term Source #
Transform list of terms into a term build from binary maximum.
sizeView :: (HasBuiltins m, MonadTCEnv m, ReadTCState m) => Term -> m SizeView Source #
Expects argument to be reduce
d.
data ProjectedVar Source #
A de Bruijn index under some projections.
ProjectedVar | |
|
Instances
Show ProjectedVar Source # | |
Defined in Agda.TypeChecking.Monad.SizedTypes showsPrec :: Int -> ProjectedVar -> ShowS show :: ProjectedVar -> String showList :: [ProjectedVar] -> ShowS | |
Eq ProjectedVar Source # | Ignore |
Defined in Agda.TypeChecking.Monad.SizedTypes (==) :: ProjectedVar -> ProjectedVar -> Bool (/=) :: ProjectedVar -> ProjectedVar -> Bool |
viewProjectedVar :: Term -> Maybe ProjectedVar Source #
data DeepSizeView Source #
A deep view on sizes.
Instances
Pretty DeepSizeView Source # | |
Defined in Agda.TypeChecking.Monad.SizedTypes pretty :: DeepSizeView -> Doc Source # prettyPrec :: Int -> DeepSizeView -> Doc Source # prettyList :: [DeepSizeView] -> Doc Source # | |
Show DeepSizeView Source # | |
Defined in Agda.TypeChecking.Monad.SizedTypes showsPrec :: Int -> DeepSizeView -> ShowS show :: DeepSizeView -> String showList :: [DeepSizeView] -> ShowS |
data SizeViewComparable a Source #
Instances
Functor SizeViewComparable Source # | |
Defined in Agda.TypeChecking.Monad.SizedTypes fmap :: (a -> b) -> SizeViewComparable a -> SizeViewComparable b (<$) :: a -> SizeViewComparable b -> SizeViewComparable a # |
sizeViewComparable :: DeepSizeView -> DeepSizeView -> SizeViewComparable () Source #
sizeViewComparable v w
checks whether v >= w
(then Left
)
or v <= w
(then Right
). If uncomparable, it returns NotComparable
.
sizeViewSuc_ :: QName -> DeepSizeView -> DeepSizeView Source #
sizeViewPred :: Nat -> DeepSizeView -> DeepSizeView Source #
sizeViewPred k v
decrements v
by k
(must be possible!).
sizeViewOffset :: DeepSizeView -> Maybe Offset Source #
sizeViewOffset v
returns the number of successors or Nothing when infty.
removeSucs :: (DeepSizeView, DeepSizeView) -> (DeepSizeView, DeepSizeView) Source #
Remove successors common to both sides.
unDeepSizeView :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => DeepSizeView -> m Term Source #
type SizeMaxView = List1 DeepSizeView Source #
type SizeMaxView' = [DeepSizeView] Source #
maxViewMax :: SizeMaxView -> SizeMaxView -> SizeMaxView Source #
maxViewCons :: DeepSizeView -> SizeMaxView -> SizeMaxView Source #
maxViewCons v ws = max v ws
. It only adds v
to ws
if it is not
subsumed by an element of ws
.
sizeViewComparableWithMax :: DeepSizeView -> SizeMaxView -> SizeViewComparable SizeMaxView' Source #
sizeViewComparableWithMax v ws
tries to find w
in ws
that compares with v
and singles this out.
Precondition: v /= DSizeInv
.
maxViewSuc_ :: QName -> SizeMaxView -> SizeMaxView Source #
unMaxView :: (HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) => SizeMaxView -> m Term Source #
modifyContextInfo :: MonadTCEnv tcm => (forall e. Dom e -> Dom e) -> tcm a -> tcm a Source #
Modify the Dom
part of context entries.
unsafeInTopContext :: (MonadTCEnv m, ReadTCState m) => m a -> m a Source #
Change to top (=empty) context, but don't update the checkpoints. Totally not safe!
unsafeEscapeContext :: MonadTCM tcm => Int -> tcm a -> tcm a Source #
Delete the last n
bindings from the context.
Doesn't update checkpoints! Use escapeContext
or `updateContext
rho (drop n)` instead, for an appropriate substitution rho
.
escapeContext :: MonadAddContext m => Impossible -> Int -> m a -> m a Source #
Delete the last n
bindings from the context. Any occurrences of
these variables are replaced with the given err
.
checkpoint :: (MonadDebug tcm, MonadTCM tcm, MonadFresh CheckpointId tcm, ReadTCState tcm) => Substitution -> tcm a -> tcm a Source #
Add a new checkpoint. Do not use directly!
getContextTelescope :: (Applicative m, MonadTCEnv m) => m Telescope Source #
Get the current context as a Telescope
.
checkpointSubstitution' :: MonadTCEnv tcm => CheckpointId -> tcm (Maybe Substitution) Source #
Get the substitution from the context at a given checkpoint to the current context.
getModuleParameterSub :: (MonadTCEnv m, ReadTCState m) => ModuleName -> m (Maybe Substitution) Source #
Get substitution Γ ⊢ ρ : Γm
where Γ
is the current context
and Γm
is the module parameter telescope of module m
.
Returns Nothing
in case the we don't have a checkpoint for m
.
defaultAddCtx :: MonadAddContext m => Name -> Dom Type -> m a -> m a Source #
Default implementation of addCtx in terms of updateContext
withFreshName_ :: MonadAddContext m => ArgName -> (Name -> m a) -> m a Source #
withShadowingNameTCM :: Name -> TCM b -> TCM b Source #
Run the given TCM action, and register the given variable as being shadowed by all the names with the same root that are added to the context during this TCM action.
defaultAddLetBinding' :: (ReadTCState m, MonadTCEnv m) => Origin -> Name -> Term -> Dom Type -> m a -> m a Source #
Add a let bound variable
addRecordNameContext :: (MonadAddContext m, MonadFresh NameId m) => Dom Type -> m b -> m b Source #
class AddContext b where Source #
Various specializations of addCtx
.
addContext :: MonadAddContext m => b -> m a -> m a Source #
contextSize :: b -> Nat Source #
Instances
AddContext Name Source # | |
Defined in Agda.TypeChecking.Monad.Context addContext :: MonadAddContext m => Name -> m a -> m a Source # contextSize :: Name -> Nat Source # | |
AddContext Telescope Source # | |
Defined in Agda.TypeChecking.Monad.Context addContext :: MonadAddContext m => Telescope -> m a -> m a Source # contextSize :: Telescope -> Nat Source # | |
AddContext String Source # | |
Defined in Agda.TypeChecking.Monad.Context addContext :: MonadAddContext m => String -> m a -> m a Source # contextSize :: String -> Nat Source # | |
AddContext (Dom Type) Source # | |
Defined in Agda.TypeChecking.Monad.Context addContext :: MonadAddContext m => Dom Type -> m a -> m a Source # | |
AddContext (Dom (Name, Type)) Source # | |
Defined in Agda.TypeChecking.Monad.Context addContext :: MonadAddContext m => Dom (Name, Type) -> m a -> m a Source # | |
AddContext (Dom (String, Type)) Source # | |
Defined in Agda.TypeChecking.Monad.Context addContext :: MonadAddContext m => Dom (String, Type) -> m a -> m a Source # | |
AddContext (KeepNames Telescope) Source # | |
Defined in Agda.TypeChecking.Monad.Context addContext :: MonadAddContext m => KeepNames Telescope -> m a -> m a Source # | |
AddContext a => AddContext [a] Source # | |
Defined in Agda.TypeChecking.Monad.Context addContext :: MonadAddContext m => [a] -> m a0 -> m a0 Source # contextSize :: [a] -> Nat Source # | |
AddContext (Name, Dom Type) Source # | |
Defined in Agda.TypeChecking.Monad.Context addContext :: MonadAddContext m => (Name, Dom Type) -> m a -> m a Source # | |
AddContext (KeepNames String, Dom Type) Source # | |
Defined in Agda.TypeChecking.Monad.Context addContext :: MonadAddContext m => (KeepNames String, Dom Type) -> m a -> m a Source # | |
AddContext (List1 Name, Dom Type) Source # | |
Defined in Agda.TypeChecking.Monad.Context | |
AddContext (List1 (Arg Name), Type) Source # | |
Defined in Agda.TypeChecking.Monad.Context | |
AddContext (List1 (NamedArg Name), Type) Source # | |
Defined in Agda.TypeChecking.Monad.Context | |
AddContext (List1 (WithHiding Name), Dom Type) Source # | |
Defined in Agda.TypeChecking.Monad.Context addContext :: MonadAddContext m => (List1 (WithHiding Name), Dom Type) -> m a -> m a Source # contextSize :: (List1 (WithHiding Name), Dom Type) -> Nat Source # | |
AddContext (Text, Dom Type) Source # | |
Defined in Agda.TypeChecking.Monad.Context addContext :: MonadAddContext m => (Text, Dom Type) -> m a -> m a Source # | |
AddContext (String, Dom Type) Source # | |
Defined in Agda.TypeChecking.Monad.Context addContext :: MonadAddContext m => (String, Dom Type) -> m a -> m a Source # | |
AddContext ([Name], Dom Type) Source # | |
Defined in Agda.TypeChecking.Monad.Context addContext :: MonadAddContext m => ([Name], Dom Type) -> m a -> m a Source # | |
AddContext ([Arg Name], Type) Source # | |
Defined in Agda.TypeChecking.Monad.Context addContext :: MonadAddContext m => ([Arg Name], Type) -> m a -> m a Source # | |
AddContext ([NamedArg Name], Type) Source # | |
Defined in Agda.TypeChecking.Monad.Context addContext :: MonadAddContext m => ([NamedArg Name], Type) -> m a -> m a Source # | |
AddContext ([WithHiding Name], Dom Type) Source # | |
Defined in Agda.TypeChecking.Monad.Context addContext :: MonadAddContext m => ([WithHiding Name], Dom Type) -> m a -> m a Source # contextSize :: ([WithHiding Name], Dom Type) -> Nat Source # |
Wrapper to tell addContext
not to mark names as
NotInScope
. Used when adding a user-provided, but already type
checked, telescope to the context.
Instances
AddContext (KeepNames Telescope) Source # | |
Defined in Agda.TypeChecking.Monad.Context addContext :: MonadAddContext m => KeepNames Telescope -> m a -> m a Source # | |
AddContext (KeepNames String, Dom Type) Source # | |
Defined in Agda.TypeChecking.Monad.Context addContext :: MonadAddContext m => (KeepNames String, Dom Type) -> m a -> m a Source # |
underAbstraction' :: (Subst a, MonadAddContext m, AddContext (name, Dom Type)) => (String -> name) -> Dom Type -> Abs a -> (a -> m b) -> m b Source #
underAbstraction :: (Subst a, MonadAddContext m) => Dom Type -> Abs a -> (a -> m b) -> m b Source #
Go under an abstraction. Do not extend context in case of NoAbs
.
underAbstractionAbs' :: (Subst a, MonadAddContext m, AddContext (name, Dom Type)) => (String -> name) -> Dom Type -> Abs a -> (a -> m b) -> m b Source #
underAbstractionAbs :: (Subst a, MonadAddContext m) => Dom Type -> Abs a -> (a -> m b) -> m b Source #
underAbstraction_ :: (Subst a, MonadAddContext m) => Abs a -> (a -> m b) -> m b Source #
Go under an abstract without worrying about the type to add to the context.
mapAbstraction :: (Subst a, Subst b, MonadAddContext m) => Dom Type -> (a -> m b) -> Abs a -> m (Abs b) Source #
Map a monadic function on the thing under the abstraction, adding the abstracted variable to the context.
mapAbstraction_ :: (Subst a, Subst b, MonadAddContext m) => (a -> m b) -> Abs a -> m (Abs b) Source #
getLetBindings :: MonadTCEnv tcm => tcm [(Name, LetBinding)] Source #
addLetBinding :: MonadAddContext m => ArgInfo -> Origin -> Name -> Term -> Type -> m a -> m a Source #
Add a let bound variable
removeLetBindingsFrom :: MonadTCEnv m => Name -> m a -> m a Source #
Remove a let bound variable and all let bindings introduced after it. For instance before printing its body to avoid folding the binding itself, or using bindings defined later. Relies on the invariant that names introduced later are sorted after earlier names.
getContextSize :: (Applicative m, MonadTCEnv m) => m Nat Source #
Get the size of the current context.
getContextArgs :: (Applicative m, MonadTCEnv m) => m Args Source #
Generate [var (n - 1), ..., var 0]
for all declarations in the context.
getContextTerms :: (Applicative m, MonadTCEnv m) => m [Term] Source #
Generate [var (n - 1), ..., var 0]
for all declarations in the context.
getContextNames :: (Applicative m, MonadTCEnv m) => m [Name] Source #
Get the names of all declarations in the context.
lookupBV_ :: Nat -> Context -> Maybe ContextEntry Source #
get type of bound variable (i.e. deBruijn index)
lookupBV' :: MonadTCEnv m => Nat -> m (Maybe ContextEntry) Source #
getVarInfo :: (MonadFail m, MonadTCEnv m) => Name -> m (Term, Dom Type) Source #
Get the term corresponding to a named variable. If it is a lambda bound variable the deBruijn index is returned and if it is a let bound variable its definition is returned.
workOnTypes' :: MonadTCEnv m => Bool -> m a -> m a Source #
Internal workhorse, expects value of --experimental-irrelevance flag as argument.
applyQuantityToJudgement :: (MonadTCEnv tcm, LensQuantity q) => q -> tcm a -> tcm a Source #
Apply the quantity to the quantity annotation of the (typing/equality) judgement.
Precondition: The quantity must not be
.Quantity1
something
applyRelevanceToContext :: (MonadTCEnv tcm, LensRelevance r) => r -> tcm a -> tcm a Source #
(Conditionally) wake up irrelevant variables and make them relevant. For instance, in an irrelevant function argument otherwise irrelevant variables may be used, so they are awoken before type checking the argument.
Also allow the use of irrelevant definitions.
applyRelevanceToContextOnly :: MonadTCEnv tcm => Relevance -> tcm a -> tcm a Source #
(Conditionally) wake up irrelevant variables and make them relevant. For instance, in an irrelevant function argument otherwise irrelevant variables may be used, so they are awoken before type checking the argument.
Precondition: Relevance /= Relevant
applyRelevanceToJudgementOnly :: MonadTCEnv tcm => Relevance -> tcm a -> tcm a Source #
Apply relevance rel
the the relevance annotation of the (typing/equality)
judgement. This is part of the work done when going into a rel
-context.
Precondition: Relevance /= Relevant
applyRelevanceToContextFunBody :: (MonadTCM tcm, LensRelevance r) => r -> tcm a -> tcm a Source #
Like applyRelevanceToContext
, but only act on context if
--irrelevant-projections
.
See issue #2170.
applyCohesionToContext :: (MonadTCEnv tcm, LensCohesion m) => m -> tcm a -> tcm a Source #
Apply inverse composition with the given cohesion to the typing context.
applyCohesionToContextOnly :: MonadTCEnv tcm => Cohesion -> tcm a -> tcm a Source #
splittableCohesion :: (HasOptions m, LensCohesion a) => a -> m Bool Source #
Can we split on arguments of the given cohesion?
applyModalityToContext :: (MonadTCEnv tcm, LensModality m) => m -> tcm a -> tcm a Source #
(Conditionally) wake up irrelevant variables and make them relevant. For instance, in an irrelevant function argument otherwise irrelevant variables may be used, so they are awoken before type checking the argument.
Also allow the use of irrelevant definitions.
This function might also do something for other modalities.
applyModalityToContextOnly :: MonadTCEnv tcm => Modality -> tcm a -> tcm a Source #
(Conditionally) wake up irrelevant variables and make them relevant. For instance, in an irrelevant function argument otherwise irrelevant variables may be used, so they are awoken before type checking the argument.
This function might also do something for other modalities, but not for quantities.
Precondition: Modality /= Relevant
applyModalityToJudgementOnly :: MonadTCEnv tcm => Modality -> tcm a -> tcm a Source #
Apply the relevance and quantity components of the modality to the modality annotation of the (typing/equality) judgement.
Precondition: The relevance component must not be Relevant
.
applyModalityToContextFunBody :: (MonadTCM tcm, LensModality r) => r -> tcm a -> tcm a Source #
Like applyModalityToContext
, but only act on context (for Relevance) if
--irrelevant-projections
.
See issue #2170.
wakeIrrelevantVars :: MonadTCEnv tcm => tcm a -> tcm a Source #
Wake up irrelevant variables and make them relevant. This is used
when type checking terms in a hole, in which case you want to be able to
(for instance) infer the type of an irrelevant variable. In the course
of type checking an irrelevant function argument applyRelevanceToContext
is used instead, which also sets the context relevance to Irrelevant
.
This is not the right thing to do when type checking interactively in a
hole since it also marks all metas created during type checking as
irrelevant (issue #2568).
Also set the current quantity to 0.
withClosure :: (MonadTCEnv m, ReadTCState m) => Closure a -> (a -> m b) -> m (Closure b) Source #
mapClosure :: (MonadTCEnv m, ReadTCState m) => (a -> m b) -> Closure a -> m (Closure b) Source #
solvingProblems :: MonadConstraint m => Set ProblemId -> m a -> m a Source #
isProblemSolved :: (MonadTCEnv m, ReadTCState m) => ProblemId -> m Bool Source #
isProblemSolved' :: (MonadTCEnv m, ReadTCState m) => Bool -> ProblemId -> m Bool Source #
isProblemCompletelySolved :: (MonadTCEnv m, ReadTCState m) => ProblemId -> m Bool Source #
getAllConstraints :: ReadTCState m => m Constraints Source #
isBlockingConstraint :: Constraint -> Bool Source #
A problem is considered solved if there are no unsolved blocking constraints belonging to it. There's no really good principle for what constraints are blocking and which are not, but the general idea is that nothing bad should happen if you assume a non-blocking constraint is solvable, but it turns out it isn't. For instance, assuming an equality constraint between two types that turns out to be false can lead to ill typed terms in places where we don't expect them.
getConstraintsForProblem :: ReadTCState m => ProblemId -> m Constraints Source #
getAwakeConstraints :: ReadTCState m => m Constraints Source #
Get the awake constraints
dropConstraints :: MonadConstraint m => (ProblemConstraint -> Bool) -> m () Source #
takeConstraints :: MonadConstraint m => (ProblemConstraint -> Bool) -> m Constraints Source #
Takes out all constraints matching given filter. Danger! The taken constraints need to be solved or put back at some point.
putConstraintsToSleep :: MonadConstraint m => (ProblemConstraint -> Bool) -> m () Source #
putAllConstraintsToSleep :: MonadConstraint m => m () Source #
data ConstraintStatus Source #
Instances
Show ConstraintStatus Source # | |
Defined in Agda.TypeChecking.Monad.Constraints showsPrec :: Int -> ConstraintStatus -> ShowS show :: ConstraintStatus -> String showList :: [ConstraintStatus] -> ShowS | |
Eq ConstraintStatus Source # | |
Defined in Agda.TypeChecking.Monad.Constraints (==) :: ConstraintStatus -> ConstraintStatus -> Bool (/=) :: ConstraintStatus -> ConstraintStatus -> Bool |
holdConstraints :: (ConstraintStatus -> ProblemConstraint -> Bool) -> TCM a -> TCM a Source #
Suspend constraints matching the predicate during the execution of the second argument. Caution: held sleeping constraints will not be woken up by events that would normally trigger a wakeup call.
takeAwakeConstraint :: MonadConstraint m => m (Maybe ProblemConstraint) Source #
takeAwakeConstraint' :: MonadConstraint m => (ProblemConstraint -> Bool) -> m (Maybe ProblemConstraint) Source #
withConstraint :: MonadConstraint m => (Constraint -> m a) -> ProblemConstraint -> m a Source #
buildProblemConstraint :: (MonadTCEnv m, ReadTCState m) => Set ProblemId -> Blocker -> Constraint -> m ProblemConstraint Source #
buildProblemConstraint_ :: (MonadTCEnv m, ReadTCState m) => Blocker -> Constraint -> m ProblemConstraint Source #
buildConstraint :: Blocker -> Constraint -> TCM ProblemConstraint Source #
addConstraint' :: Blocker -> Constraint -> TCM () Source #
Add new a constraint
addConstraintTo :: Lens' TCState Constraints -> Blocker -> Constraint -> TCM () Source #
addAwakeConstraint' :: Blocker -> Constraint -> TCM () Source #
nowSolvingConstraints :: MonadTCEnv m => m a -> m a Source #
Start solving constraints
isSolvingConstraints :: MonadTCEnv m => m Bool Source #
catchConstraint :: MonadConstraint m => Constraint -> m () -> m () Source #
Add constraint if the action raises a pattern violation
canDropRecursiveInstance :: (ReadTCState m, HasOptions m) => m Bool Source #
wakeConstraints' :: MonadConstraint m => (ProblemConstraint -> WakeUp) -> m () Source #
Wake constraints matching the given predicate (and aren't instance
constraints if shouldPostponeInstanceSearch
).
mapAwakeConstraints :: (Constraints -> Constraints) -> TCState -> TCState Source #
mapSleepingConstraints :: (Constraints -> Constraints) -> TCState -> TCState Source #
inMutualBlock :: (MutualId -> TCM a) -> TCM a Source #
Pass the current mutual block id or create a new mutual block if we are not already inside on.
setMutualBlockInfo :: MutualId -> MutualInfo -> TCM () Source #
Set the mutual block info for a block, possibly overwriting the existing one.
insertMutualBlockInfo :: MutualId -> MutualInfo -> TCM () Source #
Set the mutual block info for a block if non-existing.
currentOrFreshMutualBlock :: TCM MutualId Source #
Get the current mutual block, if any, otherwise a fresh mutual block is returned.
lookupMutualBlock :: ReadTCState tcm => MutualId -> tcm MutualBlock Source #
checkPragmaOptionConsistency :: PragmaOptions -> PragmaOptions -> TCM () Source #
Check that that you don't turn on inconsistent options. For instance, if --a implies --b and you have both --a and --no-b. Only warn for things that have changed compared to the old options.
setCommandLineOptions' Source #
:: AbsolutePath | The base directory of relative paths. |
-> CommandLineOptions | |
-> TCM () |
:: AbsolutePath | The base directory of relative paths. |
-> CommandLineOptions | |
-> TCM CommandLineOptions |
:: [FilePath] | New include directories. |
-> AbsolutePath | The base directory of relative paths. |
-> TCM () |
Makes the given directories absolute and stores them as include directories.
If the include directories change, then the state is reset (completely, except for the include directories and some other things).
An empty list is interpreted as ["."]
.
getAgdaLibFilesWithoutTopLevelModuleName Source #
:: AbsolutePath | The file. |
-> TCM [AgdaLibFile] |
Returns potential library files for a file without a known top-level module name.
Once the top-level module name is known one can use
checkLibraryFileNotTooFarDown
to check that the potential library
files were not located too far down the directory hierarchy.
Nothing is returned if optUseLibs
is False
.
checkLibraryFileNotTooFarDown :: TopLevelModuleName -> AgdaLibFile -> TCM () Source #
Checks that a library file for the module A.B.C
(say) in the
directory dirAB
is located at least two directories above the
file (not in dir/A
or dirAB
).
:: AbsolutePath | The file name. |
-> TopLevelModuleName | The top-level module name. |
-> TCM [OptionsPragma] |
Returns the library options for a given file.
:: AbsolutePath | The base directory of relative paths. |
-> CommandLineOptions | |
-> TCM CommandLineOptions |
setOptionsFromPragma :: OptionsPragma -> TCM () Source #
Set pragma options without checking for consistency.
setOptionsFromPragma' :: Bool -> OptionsPragma -> TCM () Source #
checkAndSetOptionsFromPragma :: OptionsPragma -> TCM () Source #
Set pragma options and check them for consistency.
enableDisplayForms :: MonadTCEnv m => m a -> m a Source #
Disable display forms.
disableDisplayForms :: MonadTCEnv m => m a -> m a Source #
Disable display forms.
displayFormsEnabled :: MonadTCEnv m => m Bool Source #
Check if display forms are enabled.
isPropEnabled :: HasOptions m => m Bool Source #
isLevelUniverseEnabled :: HasOptions m => m Bool Source #
isTwoLevelEnabled :: HasOptions m => m Bool Source #
hasUniversePolymorphism :: HasOptions m => m Bool Source #
showImplicitArguments :: HasOptions m => m Bool Source #
showGeneralizedArguments :: HasOptions m => m Bool Source #
showIrrelevantArguments :: HasOptions m => m Bool Source #
showIdentitySubstitutions :: HasOptions m => m Bool Source #
withShowAllArguments :: ReadTCState m => m a -> m a Source #
Switch on printing of implicit and irrelevant arguments. E.g. for reification in with-function generation.
Restores all PragmaOptions
after completion.
Thus, do not attempt to make persistent PragmaOptions
changes in a withShowAllArguments
bracket.
withShowAllArguments' :: ReadTCState m => Bool -> m a -> m a Source #
withPragmaOptions :: ReadTCState m => (PragmaOptions -> PragmaOptions) -> m a -> m a Source #
Change PragmaOptions
for a computation and restore afterwards.
withoutPrintingGeneralization :: ReadTCState m => m a -> m a Source #
positivityCheckEnabled :: HasOptions m => m Bool Source #
typeInType :: HasOptions m => m Bool Source #
etaEnabled :: HasOptions m => m Bool Source #
maxInstanceSearchDepth :: HasOptions m => m Int Source #
maxInversionDepth :: HasOptions m => m Int Source #
getLanguage :: HasOptions m => m Language Source #
Returns the Language
currently in effect.
setHardCompileTimeModeIfErased' Source #
:: LensQuantity q | |
=> q | The quantity. |
-> TCM a | Continuation. |
-> TCM a |
If the quantity is "erased", then hard compile-time mode is enabled when the continuation is run.
Precondition: The quantity must not be
.Quantity1
something
setRunTimeModeUnlessInHardCompileTimeMode Source #
Use run-time mode in the continuation unless the current mode is the hard compile-time mode.
warnForPlentyInHardCompileTimeMode :: Erased -> TCM () Source #
Warn if the user explicitly wrote @ω
or @plenty
but the
current mode is the hard compile-time mode.
addConstant :: QName -> Definition -> TCM () Source #
Add a constant to the signature. Lifts the definition to top level.
addConstant' :: QName -> ArgInfo -> QName -> Type -> Defn -> TCM () Source #
A combination of addConstant
and defaultDefn
. The Language
does not need to be supplied.
setTerminates :: MonadTCState m => QName -> Bool -> m () Source #
Set termination info of a defined function symbol.
setCompiledClauses :: QName -> CompiledClauses -> TCM () Source #
Set CompiledClauses of a defined function symbol.
modifyFunClauses :: QName -> ([Clause] -> [Clause]) -> TCM () Source #
Modify the clauses of a function.
addClauses :: (MonadConstraint m, MonadTCState m) => QName -> [Clause] -> m () Source #
Lifts clauses to the top-level and adds them to definition.
Also adjusts the funCopatternLHS
field if necessary.
mkPragma :: String -> TCM CompilerPragma Source #
getUniqueCompilerPragma :: BackendName -> QName -> TCM (Maybe CompilerPragma) Source #
setFunctionFlag :: FunctionFlag -> Bool -> QName -> TCM () Source #
markStatic :: QName -> TCM () Source #
markInline :: Bool -> QName -> TCM () Source #
markInjective :: QName -> TCM () Source #
markFirstOrder :: QName -> TCM () Source #
unionSignatures :: [Signature] -> Signature Source #
addSection :: ModuleName -> TCM () Source #
Add a section to the signature.
The current context will be stored as the cumulative module parameters for this section.
getSection :: (Functor m, ReadTCState m) => ModuleName -> m (Maybe Section) Source #
Get a section.
Why Maybe? The reason is that we look up all prefixes of a module to compute number of parameters, and for hierarchical top-level modules, A.B.C say, A and A.B do not exist.
setModuleCheckpoint :: ModuleName -> TCM () Source #
Sets the checkpoint for the given module to the current checkpoint.
addDisplayForms :: QName -> TCM () Source #
Add display forms for a name f
copied by a module application. Essentially if f
can reduce to
λ xs → A.B.C.f vs
by unfolding module application copies (defCopy
), then we add a display form
A.B.C.f vs ==> f xs
projectionArgs :: Definition -> Int Source #
Number of dropped initial arguments of a projection(-like) function.
addDisplayForm :: QName -> DisplayForm -> TCM () Source #
Add a display form to a definition (could be in this or imported signature).
:: ModuleName | Name of new module defined by the module macro. |
-> Telescope | Parameters of new module. |
-> ModuleName | Name of old module applied to arguments. |
-> Args | Arguments of module application. |
-> ScopeCopyInfo | Imported names and modules |
-> TCM () |
Module application (followed by module parameter abstraction).
applySection' :: ModuleName -> Telescope -> ModuleName -> Args -> ScopeCopyInfo -> TCM () Source #
hasLoopingDisplayForm :: QName -> TCM Bool Source #
Check if a display form is looping.
getDisplayForms :: (HasConstInfo m, ReadTCState m) => QName -> m [LocalDisplayForm] Source #
hasDisplayForms :: (HasConstInfo m, ReadTCState m) => QName -> m Bool Source #
chaseDisplayForms :: QName -> TCM (Set QName) Source #
Find all names used (recursively) by display forms of a given name.
canonicalName :: HasConstInfo m => QName -> m QName Source #
singleConstructorType :: QName -> TCM Bool Source #
Does the given constructor come from a single-constructor type?
Precondition: The name has to refer to a constructor.
prettySigCubicalNotErasure :: MonadPretty m => QName -> m Doc Source #
Generates an error message corresponding to
SigCubicalNotErasure
for a given QName
.
sigError :: (HasCallStack, MonadDebug m) => m a -> SigError -> m a Source #
An eliminator for SigError
. All constructors except for
SigAbstract
are assumed to be impossible.
getOriginalConstInfo :: (ReadTCState m, HasConstInfo m) => QName -> m Definition Source #
The computation getConstInfo
sometimes tweaks the returned
Definition
, depending on the current Language
and the
Language
of the Definition
. This variant of getConstInfo
does
not perform any tweaks.
defaultGetRewriteRulesFor :: (ReadTCState m, MonadTCEnv m) => QName -> m RewriteRules Source #
getOriginalProjection :: HasConstInfo m => QName -> m QName Source #
Get the original name of the projection (the current one could be from a module application).
isProjection :: HasConstInfo m => QName -> m (Maybe Projection) Source #
Is it the name of a record projection?
defaultGetConstInfo :: (HasOptions m, MonadDebug m, MonadTCEnv m) => TCState -> TCEnv -> QName -> m (Either SigError Definition) Source #
alwaysMakeAbstract :: Definition -> Maybe Definition Source #
Return the abstract view of a definition, regardless of whether the definition would be treated abstractly.
getConInfo :: HasConstInfo m => ConHead -> m Definition Source #
getPolarity' :: HasConstInfo m => Comparison -> QName -> m [Polarity] Source #
Look up polarity of a definition and compose with polarity
represented by Comparison
.
setPolarity :: (MonadTCState m, MonadDebug m) => QName -> [Polarity] -> m () Source #
Set the polarity of a definition.
getForcedArgs :: HasConstInfo m => QName -> m [IsForced] Source #
Look up the forced arguments of a definition.
getArgOccurrence :: QName -> Nat -> TCM Occurrence Source #
Get argument occurrence info for argument i
of definition d
(never fails).
setArgOccurrences :: MonadTCState m => QName -> [Occurrence] -> m () Source #
Sets the defArgOccurrences
for the given identifier (which
should already exist in the signature).
modifyArgOccurrences :: MonadTCState m => QName -> ([Occurrence] -> [Occurrence]) -> m () Source #
getCompiled :: HasConstInfo m => QName -> m (Maybe Compiled) Source #
setErasedConArgs :: QName -> [Bool] -> TCM () Source #
getCompiledArgUse :: HasConstInfo m => QName -> m (Maybe [ArgUsage]) Source #
getMutual :: QName -> TCM (Maybe [QName]) Source #
Get the mutually recursive identifiers of a symbol from the signature.
getMutual_ :: Defn -> Maybe [QName] Source #
Get the mutually recursive identifiers from a Definition
.
setMutual :: QName -> [QName] -> TCM () Source #
Set the mutually recursive identifiers.
TODO: This produces data of quadratic size (which has to be processed upon serialization). Presumably qs is usually short, but in some cases (for instance for generated code) it may be long. It would be better to assign a unique identifier to each SCC, and store the names separately.
mutuallyRecursive :: QName -> QName -> TCM Bool Source #
Check whether two definitions are mutually recursive.
definitelyNonRecursive_ :: Defn -> Bool Source #
A functiondatarecord definition is nonRecursive if it is not even mutually recursive with itself.
getCurrentModuleFreeVars :: TCM Nat Source #
Get the number of parameters to the current module.
getDefModule :: HasConstInfo m => QName -> m (Either SigError ModuleName) Source #
getDefFreeVars :: (Functor m, Applicative m, ReadTCState m, MonadTCEnv m) => QName -> m Nat Source #
Compute the number of free variables of a defined name. This is the sum of number of parameters shared with the current module and the number of anonymous variables (if the name comes from a let-bound module).
getModuleFreeVars :: (Functor m, Applicative m, MonadTCEnv m, ReadTCState m) => ModuleName -> m Nat Source #
freeVarsToApply :: (Functor m, HasConstInfo m, HasOptions m, ReadTCState m, MonadTCEnv m, MonadDebug m) => QName -> m Args Source #
moduleParamsToApply :: (Functor m, Applicative m, HasOptions m, MonadTCEnv m, ReadTCState m, MonadDebug m) => ModuleName -> m Args Source #
Compute the context variables to apply a definition to.
We have to insert the module telescope of the common prefix of the current module and the module where the definition comes from. (Properly raised to the current context.)
Example:
module M₁ Γ where
module M₁ Δ where
f = ...
module M₃ Θ where
... M₁.M₂.f [insert Γ raised by Θ]
instantiateDef :: (Functor m, HasConstInfo m, HasOptions m, ReadTCState m, MonadTCEnv m, MonadDebug m) => Definition -> m Definition Source #
Instantiate a closed definition with the correct part of the current context.
instantiateRewriteRule :: (Functor m, HasConstInfo m, HasOptions m, ReadTCState m, MonadTCEnv m, MonadDebug m) => RewriteRule -> m RewriteRule Source #
instantiateRewriteRules :: (Functor m, HasConstInfo m, HasOptions m, ReadTCState m, MonadTCEnv m, MonadDebug m) => RewriteRules -> m RewriteRules Source #
inAbstractMode :: MonadTCEnv m => m a -> m a Source #
Enter abstract mode. Abstract definition in the current module are transparent.
inConcreteMode :: MonadTCEnv m => m a -> m a Source #
Not in abstract mode. All abstract definitions are opaque.
ignoreAbstractMode :: MonadTCEnv m => m a -> m a Source #
Ignore abstract mode. All abstract definitions are transparent.
underOpaqueId :: MonadTCEnv m => OpaqueId -> m a -> m a Source #
Go under the given opaque block. The unfolding set will turn opaque definitions transparent.
inConcreteOrAbstractMode :: (MonadTCEnv m, HasConstInfo m) => QName -> (Definition -> m a) -> m a Source #
Enter the reducibility environment associated with a definition: The environment will have the same concreteness as the name, and we will be in the opaque block enclosing the name, if any.
relOfConst :: HasConstInfo m => QName -> m Relevance Source #
Get relevance of a constant.
modalityOfConst :: HasConstInfo m => QName -> m Modality Source #
Get modality of a constant.
droppedPars :: Definition -> Int Source #
The number of dropped parameters for a definition. 0 except for projection(-like) functions and constructors.
isProjection_ :: Defn -> Maybe Projection Source #
isRelevantProjection :: HasConstInfo m => QName -> m (Maybe Projection) Source #
Is it the name of a non-irrelevant record projection?
isStaticFun :: Defn -> Bool Source #
Is it a function marked STATIC?
isInlineFun :: Defn -> Bool Source #
Is it a function marked INLINE?
isProperProjection :: Defn -> Bool Source #
Returns True
if we are dealing with a proper projection,
i.e., not a projection-like function nor a record field value
(projection applied to argument).
usesCopatterns :: (HasConstInfo m, HasBuiltins m) => QName -> m Bool Source #
Check whether a definition uses copatterns.
allMetaClasses :: [MetaClass] Source #
All possible metavariable classes.
dontAssignMetas :: (MonadTCEnv m, HasOptions m, MonadDebug m) => m a -> m a Source #
Switch off assignment of metas.
isRemoteMeta :: ReadTCState m => m (MetaId -> Bool) Source #
Is the meta-variable from another top-level module?
nextLocalMeta :: ReadTCState m => m MetaId Source #
If another meta-variable is created, then it will get this
MetaId
(unless the state is changed too much, for instance by
setTopLevelModule
).
data LocalMetaStores Source #
Pairs of local meta-stores.
LocalMetaStores | |
|
metasCreatedBy :: ReadTCState m => m a -> m (a, LocalMetaStores) Source #
Run a computation and record which new metas it created.
lookupLocalMeta' :: ReadTCState m => MetaId -> m (Maybe MetaVariable) Source #
Find information about the given local meta-variable, if any.
lookupLocalMeta :: (HasCallStack, MonadDebug m, ReadTCState m) => MetaId -> m MetaVariable Source #
Find information about the given local meta-variable.
lookupMetaInstantiation :: ReadTCState m => MetaId -> m MetaInstantiation Source #
Find the meta-variable's instantiation.
lookupMetaJudgement :: ReadTCState m => MetaId -> m (Judgement MetaId) Source #
Find the meta-variable's judgement.
lookupMetaModality :: ReadTCState m => MetaId -> m Modality Source #
Find the meta-variable's modality.
updateMetaVarTCM :: HasCallStack => MetaId -> (MetaVariable -> MetaVariable) -> TCM () Source #
Update the information associated with a local meta-variable.
isOpenMeta :: MetaInstantiation -> Bool Source #
insertMetaVar :: MetaId -> MetaVariable -> TCM () Source #
Insert a new meta-variable with associated information into the local meta store.
getMetaPriority :: (HasCallStack, MonadDebug m, ReadTCState m) => MetaId -> m MetaPriority Source #
Returns the MetaPriority
of the given local meta-variable.
isSortMeta :: ReadTCState m => MetaId -> m Bool Source #
isSortJudgement :: Judgement a -> Bool Source #
isSortMeta_ :: MetaVariable -> Bool Source #
metaInstantiationToMetaKind :: MetaInstantiation -> MetaKind Source #
If a meta variable is still open, what is its kind?
getMetaType :: ReadTCState m => MetaId -> m Type Source #
getMetaContextArgs :: MonadTCEnv m => MetaVariable -> m Args Source #
Compute the context variables that a local meta-variable should be applied to, accounting for pruning.
getMetaTypeInContext :: (HasBuiltins m, HasCallStack, MonadDebug m, MonadReduce m, MonadTCEnv m, ReadTCState m) => MetaId -> m Type Source #
Given a local meta-variable, return the type applied to the current context.
isGeneralizableMeta :: (HasCallStack, MonadDebug m, ReadTCState m) => MetaId -> m DoGeneralize Source #
Is it a local meta-variable that might be generalized?
class IsInstantiatedMeta a where Source #
Check whether all metas are instantiated. Precondition: argument is a meta (in some form) or a list of metas.
isInstantiatedMeta :: (MonadFail m, ReadTCState m) => a -> m Bool Source #
Instances
IsInstantiatedMeta MetaId Source # | |
Defined in Agda.TypeChecking.Monad.MetaVars isInstantiatedMeta :: (MonadFail m, ReadTCState m) => MetaId -> m Bool Source # | |
IsInstantiatedMeta Level Source # | |
Defined in Agda.TypeChecking.Monad.MetaVars isInstantiatedMeta :: (MonadFail m, ReadTCState m) => Level -> m Bool Source # | |
IsInstantiatedMeta PlusLevel Source # | |
Defined in Agda.TypeChecking.Monad.MetaVars isInstantiatedMeta :: (MonadFail m, ReadTCState m) => PlusLevel -> m Bool Source # | |
IsInstantiatedMeta Term Source # | |
Defined in Agda.TypeChecking.Monad.MetaVars isInstantiatedMeta :: (MonadFail m, ReadTCState m) => Term -> m Bool Source # | |
IsInstantiatedMeta a => IsInstantiatedMeta (Arg a) Source # | |
Defined in Agda.TypeChecking.Monad.MetaVars isInstantiatedMeta :: (MonadFail m, ReadTCState m) => Arg a -> m Bool Source # | |
IsInstantiatedMeta a => IsInstantiatedMeta (Abs a) Source # | Does not worry about raising. |
Defined in Agda.TypeChecking.Monad.MetaVars isInstantiatedMeta :: (MonadFail m, ReadTCState m) => Abs a -> m Bool Source # | |
IsInstantiatedMeta a => IsInstantiatedMeta (Maybe a) Source # | |
Defined in Agda.TypeChecking.Monad.MetaVars isInstantiatedMeta :: (MonadFail m, ReadTCState m) => Maybe a -> m Bool Source # | |
IsInstantiatedMeta a => IsInstantiatedMeta [a] Source # | |
Defined in Agda.TypeChecking.Monad.MetaVars isInstantiatedMeta :: (MonadFail m, ReadTCState m) => [a] -> m Bool Source # |
isInstantiatedMeta' :: (MonadFail m, ReadTCState m) => MetaId -> m (Maybe Term) Source #
constraintMetas :: Constraint -> TCM (Set MetaId) Source #
Returns all metavariables in a constraint. Slightly complicated by the fact that blocked terms are represented by two meta variables. To find the second one we need to look up the meta listeners for the one in the UnBlock constraint. This is used for the purpose of deciding if a metavariable is constrained or if it can be generalized over (see Agda.TypeChecking.Generalize).
getMetaListeners :: (HasCallStack, MonadDebug m, ReadTCState m) => MetaId -> m [Listener] Source #
Get the listeners for a local meta-variable.
createMetaInfo :: (MonadTCEnv m, ReadTCState m) => m MetaInfo Source #
Create MetaInfo
in the current environment.
createMetaInfo' :: (MonadTCEnv m, ReadTCState m) => RunMetaOccursCheck -> m MetaInfo Source #
setValueMetaName :: MonadMetaSolver m => Term -> MetaNameSuggestion -> m () Source #
setMetaNameSuggestion :: MonadMetaSolver m => MetaId -> MetaNameSuggestion -> m () Source #
getMetaNameSuggestion :: (HasCallStack, MonadDebug m, ReadTCState m) => MetaId -> m MetaNameSuggestion Source #
setMetaGeneralizableArgInfo :: MonadMetaSolver m => MetaId -> ArgInfo -> m () Source #
Change the ArgInfo that will be used when generalizing over this local meta-variable.
updateMetaVarRange :: MonadMetaSolver m => MetaId -> Range -> m () Source #
setMetaOccursCheck :: MonadMetaSolver m => MetaId -> RunMetaOccursCheck -> m () Source #
findInteractionPoint_ :: Range -> InteractionPoints -> Maybe InteractionId Source #
Find an interaction point by Range
by searching the whole map.
Issue 3000: Don't consider solved interaction points.
O(n): linear in the number of registered interaction points.
connectInteractionPoint :: MonadInteractionPoints m => InteractionId -> MetaId -> m () Source #
Hook up a local meta-variable to an interaction point.
removeInteractionPoint :: MonadInteractionPoints m => InteractionId -> m () Source #
Mark an interaction point as solved.
getInteractionPoints :: ReadTCState m => m [InteractionId] Source #
Get a list of interaction ids.
getInteractionMetas :: ReadTCState m => m [MetaId] Source #
Get all metas that correspond to unsolved interaction ids.
getUniqueMetasRanges :: (HasCallStack, MonadDebug m, ReadTCState m) => [MetaId] -> m [Range] Source #
getUnsolvedMetas :: (HasCallStack, MonadDebug m, ReadTCState m) => m [Range] Source #
getOpenMetas :: ReadTCState m => m [MetaId] Source #
getUnsolvedInteractionMetas :: (HasCallStack, MonadDebug m, ReadTCState m) => m [Range] Source #
getInteractionIdsAndMetas :: ReadTCState m => m [(InteractionId, MetaId)] Source #
Get all metas that correspond to unsolved interaction ids.
lookupInteractionPoint :: (MonadFail m, ReadTCState m, MonadError TCErr m) => InteractionId -> m InteractionPoint Source #
Get the information associated to an interaction point.
lookupInteractionId :: (MonadFail m, ReadTCState m, MonadError TCErr m, MonadTCEnv m) => InteractionId -> m MetaId Source #
Get MetaId
for an interaction point.
Precondition: interaction point is connected.
lookupInteractionMeta :: ReadTCState m => InteractionId -> m (Maybe MetaId) Source #
Check whether an interaction id is already associated with a meta variable.
newMeta :: MonadMetaSolver m => Frozen -> MetaInfo -> MetaPriority -> Permutation -> Judgement a -> m MetaId Source #
Generate new meta variable.
newMetaTCM' :: MetaInstantiation -> Frozen -> MetaInfo -> MetaPriority -> Permutation -> Judgement a -> TCM MetaId Source #
Generate a new meta variable with some instantiation given.
For instance, the instantiation could be a PostponedTypeCheckingProblem
.
getInteractionScope :: (MonadDebug m, MonadFail m, ReadTCState m, MonadError TCErr m, MonadTCEnv m) => InteractionId -> m ScopeInfo Source #
withMetaInfo' :: (MonadTCEnv m, ReadTCState m, MonadTrace m) => MetaVariable -> m a -> m a Source #
withMetaInfo :: (MonadTCEnv m, ReadTCState m, MonadTrace m) => Closure Range -> m a -> m a Source #
listenToMeta :: MonadMetaSolver m => Listener -> MetaId -> m () Source #
listenToMeta l m
: register l
as a listener to m
. This is done
when the type of l is blocked by m
.
unlistenToMeta :: MonadMetaSolver m => Listener -> MetaId -> m () Source #
Unregister a listener.
clearMetaListeners :: MonadMetaSolver m => MetaId -> m () Source #
etaExpandMetaSafe :: MonadMetaSolver m => MetaId -> m () Source #
Do safe eta-expansions for meta (SingletonRecords,Levels
).
etaExpandListeners :: MonadMetaSolver m => MetaId -> m () Source #
Eta expand metavariables listening on the current meta.
wakeupListener :: MonadMetaSolver m => Listener -> m () Source #
Wake up a meta listener and let it do its thing
solveAwakeConstraints :: MonadConstraint m => m () Source #
solveAwakeConstraints' :: MonadConstraint m => Bool -> m () Source #
freezeMetas :: MonadTCState m => LocalMetaStore -> m (Set MetaId) Source #
Freeze the given meta-variables (but only if they are open) and return those that were not already frozen.
unfreezeMetas :: TCM () Source #
Thaw all open meta variables.
isFrozen :: (HasCallStack, MonadDebug m, ReadTCState m) => MetaId -> m Bool Source #
withFrozenMetas :: (MonadMetaSolver m, MonadTCState m) => m a -> m a Source #
class UnFreezeMeta a where Source #
Unfreeze a meta and its type if this is a meta again. Does not unfreeze deep occurrences of meta-variables or remote meta-variables.
unfreezeMeta :: MonadMetaSolver m => a -> m () Source #
Instances
UnFreezeMeta MetaId Source # | |
Defined in Agda.TypeChecking.Monad.MetaVars unfreezeMeta :: MonadMetaSolver m => MetaId -> m () Source # | |
UnFreezeMeta Level Source # | |
Defined in Agda.TypeChecking.Monad.MetaVars unfreezeMeta :: MonadMetaSolver m => Level -> m () Source # | |
UnFreezeMeta PlusLevel Source # | |
Defined in Agda.TypeChecking.Monad.MetaVars unfreezeMeta :: MonadMetaSolver m => PlusLevel -> m () Source # | |
UnFreezeMeta Sort Source # | |
Defined in Agda.TypeChecking.Monad.MetaVars unfreezeMeta :: MonadMetaSolver m => Sort -> m () Source # | |
UnFreezeMeta Term Source # | |
Defined in Agda.TypeChecking.Monad.MetaVars unfreezeMeta :: MonadMetaSolver m => Term -> m () Source # | |
UnFreezeMeta Type Source # | |
Defined in Agda.TypeChecking.Monad.MetaVars unfreezeMeta :: MonadMetaSolver m => Type -> m () Source # | |
UnFreezeMeta a => UnFreezeMeta (Abs a) Source # | |
Defined in Agda.TypeChecking.Monad.MetaVars unfreezeMeta :: MonadMetaSolver m => Abs a -> m () Source # | |
UnFreezeMeta a => UnFreezeMeta [a] Source # | |
Defined in Agda.TypeChecking.Monad.MetaVars unfreezeMeta :: MonadMetaSolver m => [a] -> m () Source # |
data CheckResult where Source #
The result and associated parameters of a type-checked file, when invoked directly via interaction or a backend. Note that the constructor is not exported.
pattern CheckResult :: Interface -> [TCWarning] -> ModuleCheckMode -> Source -> CheckResult | Flattened unidirectional pattern for |
crInterface :: CheckResult -> Interface Source #
crWarnings :: CheckResult -> [TCWarning] Source #
crMode :: CheckResult -> ModuleCheckMode Source #
activeBackendMayEraseType :: QName -> TCM Bool Source #
Ask the active backend whether a type may be erased. See issue #3732.
backendInteraction :: AbsolutePath -> [Backend] -> TCM () -> (AbsolutePath -> TCM CheckResult) -> TCM () Source #
parseBackendOptions :: [Backend] -> [String] -> CommandLineOptions -> OptM ([Backend], CommandLineOptions) Source #
callBackend :: String -> IsMain -> CheckResult -> TCM () Source #
Call the compilerMain
function of the given backend.
lookupBackend :: BackendName -> TCM (Maybe Backend) Source #
Look for a backend of the given name.