Agda
Safe HaskellNone
LanguageHaskell2010

Agda.TypeChecking.Monad.Base

Synopsis

Documentation

pattern Axiom :: Bool -> Defn Source #

pattern Constructor :: Int -> Int -> ConHead -> QName -> IsAbstract -> CompKit -> Maybe [QName] -> [IsForced] -> Maybe [Bool] -> Bool -> Bool -> Defn Source #

pattern DDot :: Term -> DisplayTerm Source #

pattern DTerm :: Term -> DisplayTerm Source #

pattern DataOrRecSig :: Int -> Defn Source #

pattern Datatype :: Nat -> Nat -> Maybe Clause -> [QName] -> Sort -> Maybe [QName] -> IsAbstract -> [QName] -> Maybe QName -> Maybe QName -> Defn Source #

pattern PProp :: NLPat -> NLPSort Source #

pattern PSSet :: NLPat -> NLPSort Source #

pattern PType :: NLPat -> NLPSort Source #

absurdLambdaName :: String Source #

Name of absurdLambda definitions.

allReductions :: AllowedReductions Source #

Not quite all reductions (skip non-terminating reductions)

apReduce :: ReduceM (a -> b) -> ReduceM a -> ReduceM b Source #

apTCMT :: forall (m :: Type -> Type) a b. Applicative m => TCMT m (a -> b) -> TCMT m a -> TCMT m b Source #

asksTC :: MonadTCEnv m => (TCEnv -> a) -> m a Source #

bindReduce :: ReduceM a -> (a -> ReduceM b) -> ReduceM b Source #

bindTCMT :: forall (m :: Type -> Type) a b. Monad m => TCMT m a -> (a -> TCMT m b) -> TCMT m b Source #

catchError_ :: TCM a -> (TCErr -> TCM a) -> TCM a Source #

Preserve the state of the failing computation.

currentModality :: MonadTCEnv m => m Modality Source #

The current modality. Note that the returned cohesion component is always unitCohesion.

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?

defaultDefn :: ArgInfo -> QName -> Type -> Language -> Defn -> Definition Source #

Create a definition with sensible defaults.

defaultDisplayForm :: QName -> [LocalDisplayForm] Source #

By default, we have no display form.

defaultInteractionOutputCallback :: InteractionOutputCallback Source #

The default InteractionOutputCallback function prints certain things to stdout (other things generate internal errors).

dirToCmp :: (Comparison -> a -> a -> c) -> CompareDirection -> a -> a -> c Source #

Turn a Comparison function into a CompareDirection function.

Property: dirToCmp f (fromCmp cmp) = f cmp

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__.

emptyFunctionData :: HasOptions m => m FunctionData Source #

A template for creating Function definitions, with sensible defaults.

extendedLambdaName :: String Source #

Base name for extended lambda patterns

finally_ :: TCM a -> TCM b -> TCM a Source #

Execute a finalizer even when an exception is thrown. Does not catch any errors. In case both the regular computation and the finalizer throw an exception, the one of the finalizer is propagated.

flipCmp :: CompareDirection -> CompareDirection Source #

Flip the direction of comparison.

fmapReduce :: (a -> b) -> ReduceM a -> ReduceM b Source #

fmapTCMT :: forall (m :: Type -> Type) a b. Functor m => (a -> b) -> TCMT m a -> TCMT m b Source #

forkTCM :: TCM () -> TCM () Source #

Runs the given computation in a separate thread, with a copy of the current state and environment.

Note that Agda sometimes uses actual, mutable state. If the computation given to forkTCM tries to modify this state, then bad things can happen, because accesses are not mutually exclusive. The forkTCM function has been added mainly to allow the thread to read (a snapshot of) the current state in a convenient way.

Note also that exceptions which are raised in the thread are not propagated to the parent, so the thread should not do anything important.

funProj :: Lens' Defn Bool Source #

Toggle the FunProj flag.

generalizedFieldName :: String Source #

Base name for generalized variable projections

getGeneralizedFieldName :: QName -> Maybe String Source #

Check whether we have a generalized variable field

getsTC :: ReadTCState m => (TCState -> a) -> m a Source #

iFullHash :: Interface -> Hash Source #

Combines the source hash and the (full) hashes of the imported modules.

ifTopLevelAndHighlightingLevelIs :: MonadTCEnv tcm => HighlightingLevel -> tcm () -> tcm () Source #

ifTopLevelAndHighlightingLevelIs l m runs m when we're type-checking the top-level module (or before we've started doing this) and the highlighting level is at least l.

ifTopLevelAndHighlightingLevelIsOr :: MonadTCEnv tcm => HighlightingLevel -> Bool -> tcm () -> tcm () Source #

ifTopLevelAndHighlightingLevelIs l b m runs m when we're type-checking the top-level module (or before we've started doing this) and either the highlighting level is at least l or b is True.

initPreScopeState :: PreScopeState Source #

Empty state of type checker.

intSignature :: Lens' Interface Signature Source #

A lens for the iSignature field of the Interface type.

isAbsurdLambdaName :: QName -> Bool Source #

Check whether we have an definition from an absurd lambda.

isEmptyFunction :: Defn -> Bool Source #

Checking whether we are dealing with a function yet to be defined.

isExtendedLambdaName :: QName -> Bool Source #

Check whether we have an definition from an extended lambda.

isSourceCodeWarning :: WarningName -> Bool Source #

Should warnings of that type be serialized?

Only when changes in the source code can silence or influence the warning.

locallyTC :: MonadTCEnv m => Lens' TCEnv a -> (a -> a) -> m b -> m b Source #

Modify the lens-indicated part of the TCEnv in a subcomputation.

locatedTypeError :: MonadTCError m => (a -> TypeError) -> HasCallStack => a -> m b Source #

Utility function for 1-arg constructed type errors. Note that the HasCallStack constraint is on the *resulting* function.

mapTCMT :: (forall a1. m a1 -> n a1) -> TCMT m a -> TCMT n a Source #

modifyTC' :: MonadTCState m => (TCState -> TCState) -> m () Source #

A variant of modifyTC in which the computation is strict in the new state.

modifyTCLens :: MonadTCState m => Lens' TCState a -> (a -> a) -> m () Source #

Modify the part of the TCState focused on by the lens.

modifyTCLens' :: MonadTCState m => Lens' TCState a -> (a -> a) -> m () Source #

Modify the part of the TCState focused on by the lens (strictly).

modifyTCLensM :: (MonadTCState m, ReadTCState m) => Lens' TCState a -> (a -> m a) -> m () Source #

Modify a part of the state monadically.

This is an instance of %==.

patternInTeleName :: String Source #

Base name for patterns in telescopes

projArgInfo :: Projection -> ArgInfo Source #

The info of the principal (record) argument.

projDropPars :: Projection -> ProjOrigin -> Term Source #

Building the projection function (which drops the parameters).

pureTCM :: forall (m :: Type -> Type) a. MonadIO m => (TCState -> TCEnv -> a) -> TCMT m a Source #

recRecursive :: Defn -> Bool Source #

Is the record type recursive?

redBind :: ReduceM (Reduced a a') -> (a -> ReduceM b) -> (a' -> ReduceM (Reduced b b')) -> ReduceM (Reduced b b') Source #

Conceptually: redBind m f k = either (return . Left . f) k =<< m

returnTCMT :: forall (m :: Type -> Type) a. Applicative m => a -> TCMT m a Source #

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 IOExceptions) in the initial environment.

runTCM :: MonadIO m => TCEnv -> TCState -> TCMT m a -> m (a, TCState) Source #

Running the type checking monad (most general form).

runTCMTop :: TCM a -> IO (Either TCErr a) Source #

Running the type checking monad on toplevel (with initial state).

runTCMTop' :: MonadIO m => TCMT m a -> m a Source #

setEtaEquality :: EtaEquality -> HasEta -> EtaEquality Source #

Make sure we do not overwrite a user specification.

setTCLens :: MonadTCState m => Lens' TCState a -> a -> m () infix 4 Source #

Overwrite the part of the TCState focused on by the lens.

setTCLens' :: MonadTCState m => Lens' TCState a -> a -> m () Source #

Overwrite the part of the TCState focused on by the lens (strictly).

srcFilePath :: MonadFileId m => SourceFile -> m AbsolutePath Source #

Get the file name of a SourceFile.

srcFromPath :: MonadFileId m => AbsolutePath -> m SourceFile Source #

Get make a SourceFile from a file name.

stAreWeCaching :: Lens' TCState Bool Source #

use areWeCaching from the Caching module instead.

stCurrentModule :: Lens' TCState (Maybe (ModuleName, TopLevelModuleName)) Source #

Note that the lens is "strict".

stateTCLens :: (MonadTCState m, ReadTCState m) => Lens' TCState a -> (a -> (r, a)) -> m r Source #

Modify the part of the TCState focused on by the lens, and return some result.

stateTCLensM :: (MonadTCState m, ReadTCState m) => Lens' TCState a -> (a -> m (r, a)) -> m r Source #

Modify a part of the state monadically, and return some result.

This is an instance of %%=.

suffixNameSuggestion :: MetaNameSuggestion -> ArgName -> MetaNameSuggestion Source #

Append an ArgName to a MetaNameSuggestion, for computing the name suggestions of eta-expansion metas. If the MetaNameSuggestion is empty or an underscore, the field name is taken as the suggestion.

thenTCMT :: forall (m :: Type -> Type) a b. Applicative m => TCMT m a -> TCMT m b -> TCMT m b Source #

topLevelModuleFilePath :: ModuleToSource -> TopLevelModuleName -> AbsolutePath Source #

Lookup the path of a top level module name, which must be a known one.

unKleisli :: (a -> ReduceM b) -> ReduceM (a -> b) Source #

unionBuiltin :: Builtin a -> Builtin a -> Builtin a Source #

Union two Builtins. Only defined for BuiltinRewriteRelations.

viewTC :: MonadTCEnv m => Lens' TCEnv a -> m a Source #

data AbstractMode Source #

Constructors

AbstractMode

Abstract things in the current module can be accessed.

ConcreteMode

No abstract things can be accessed.

IgnoreAbstractMode

All abstract things can be accessed.

Instances

Instances details
NFData AbstractMode Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: AbstractMode -> () #

Generic AbstractMode Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep AbstractMode 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep AbstractMode = D1 ('MetaData "AbstractMode" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-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)))
Show AbstractMode Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Eq AbstractMode Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep AbstractMode Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep AbstractMode = D1 ('MetaData "AbstractMode" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-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 AllowedReduction Source #

Controlling reduce.

Constructors

ProjectionReductions

(Projection and) projection-like functions may be reduced.

InlineReductions

Functions marked INLINE may be reduced.

CopatternReductions

Copattern definitions may be reduced.

FunctionReductions

Non-recursive functions and primitives may be reduced.

RecursiveReductions

Even recursive functions may be reduced.

LevelReductions

Reduce Level terms.

TypeLevelReductions

Allow allReductions in types, even if not allowed at term level (used by confluence checker)

UnconfirmedReductions

Functions whose termination has not (yet) been confirmed.

NonTerminatingReductions

Functions that have failed termination checking.

Instances

Instances details
SmallSetElement AllowedReduction Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

NFData AllowedReduction Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: AllowedReduction -> () #

Bounded AllowedReduction Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Enum AllowedReduction Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Generic AllowedReduction Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep AllowedReduction 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep AllowedReduction = D1 ('MetaData "AllowedReduction" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (((C1 ('MetaCons "ProjectionReductions" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "InlineReductions" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "CopatternReductions" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "FunctionReductions" 'PrefixI 'False) (U1 :: Type -> Type))) :+: ((C1 ('MetaCons "RecursiveReductions" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "LevelReductions" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "TypeLevelReductions" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "UnconfirmedReductions" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "NonTerminatingReductions" 'PrefixI 'False) (U1 :: Type -> Type)))))
Ix AllowedReduction Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Show AllowedReduction Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Eq AllowedReduction Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Ord AllowedReduction Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep AllowedReduction Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep AllowedReduction = D1 ('MetaData "AllowedReduction" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (((C1 ('MetaCons "ProjectionReductions" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "InlineReductions" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "CopatternReductions" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "FunctionReductions" 'PrefixI 'False) (U1 :: Type -> Type))) :+: ((C1 ('MetaCons "RecursiveReductions" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "LevelReductions" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "TypeLevelReductions" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "UnconfirmedReductions" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "NonTerminatingReductions" 'PrefixI 'False) (U1 :: Type -> Type)))))

data ArgsCheckState a Source #

Constructors

ACState 

Fields

Instances

Instances details
Show a => Show (ArgsCheckState a) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

data AxiomData Source #

Constructors

AxiomData 

Fields

Instances

Instances details
NFData AxiomData Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: AxiomData -> () #

Generic AxiomData Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep AxiomData 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep AxiomData = D1 ('MetaData "AxiomData" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "AxiomData" 'PrefixI 'True) (S1 ('MetaSel ('Just "_axiomConstTransp") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool)))
Show AxiomData Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep AxiomData Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep AxiomData = D1 ('MetaData "AxiomData" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "AxiomData" 'PrefixI 'True) (S1 ('MetaSel ('Just "_axiomConstTransp") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool)))

type Backend' opts env menv mod def = Backend'_boot Definition TCM opts env menv mod def Source #

newtype BlockT (m :: Type -> Type) a Source #

Constructors

BlockT 

Fields

Instances

Instances details
MonadTrans BlockT Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

lift :: Monad m => m a -> BlockT m a #

HasOptions m => HasOptions (BlockT m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Monad m => MonadBlock (BlockT m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

MonadReduce m => MonadReduce (BlockT m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

liftReduce :: ReduceM a -> BlockT m a Source #

MonadTCEnv m => MonadTCEnv (BlockT m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

askTC :: BlockT m TCEnv Source #

localTC :: (TCEnv -> TCEnv) -> BlockT m a -> BlockT m a Source #

MonadTCM m => MonadTCM (BlockT m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

liftTCM :: TCM a -> BlockT m a Source #

MonadTCState m => MonadTCState (BlockT m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

ReadTCState m => ReadTCState (BlockT m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

getTCState :: BlockT m TCState Source #

locallyTCState :: Lens' TCState a -> (a -> a) -> BlockT m b -> BlockT m b Source #

withTCState :: (TCState -> TCState) -> BlockT m a -> BlockT m a Source #

HasBuiltins m => HasBuiltins (BlockT m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Builtin

MonadAddContext m => MonadAddContext (BlockT m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Context

Methods

addCtx :: Name -> Dom Type -> BlockT m a -> BlockT m a Source #

addLetBinding' :: Origin -> Name -> Term -> Dom Type -> BlockT m a -> BlockT m a Source #

updateContext :: Substitution -> (Context -> Context) -> BlockT m a -> BlockT m a Source #

withFreshName :: Range -> ArgName -> (Name -> BlockT m a) -> BlockT m a Source #

MonadDebug m => MonadDebug (BlockT m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Debug

PureTCM m => PureTCM (BlockT m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Pure

HasConstInfo m => HasConstInfo (BlockT m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Signature

Monad m => Applicative (BlockT m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

pure :: a -> BlockT m a #

(<*>) :: BlockT m (a -> b) -> BlockT m a -> BlockT m b #

liftA2 :: (a -> b -> c) -> BlockT m a -> BlockT m b -> BlockT m c #

(*>) :: BlockT m a -> BlockT m b -> BlockT m b #

(<*) :: BlockT m a -> BlockT m b -> BlockT m a #

Functor m => Functor (BlockT m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

fmap :: (a -> b) -> BlockT m a -> BlockT m b #

(<$) :: a -> BlockT m b -> BlockT m a #

Monad m => Monad (BlockT m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

(>>=) :: BlockT m a -> (a -> BlockT m b) -> BlockT m b #

(>>) :: BlockT m a -> BlockT m b -> BlockT m b #

return :: a -> BlockT m a #

MonadFail m => MonadFail (BlockT m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

fail :: String -> BlockT m a #

MonadIO m => MonadIO (BlockT m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

liftIO :: IO a -> BlockT m a #

data Builtin pf Source #

Constructors

Builtin Term 
Prim pf 
BuiltinRewriteRelations (Set QName)

BUILTIN REWRITE. We can have several rewrite relations.

Instances

Instances details
Functor Builtin Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

fmap :: (a -> b) -> Builtin a -> Builtin b #

(<$) :: a -> Builtin b -> Builtin a #

Foldable Builtin Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

fold :: Monoid m => Builtin m -> m #

foldMap :: Monoid m => (a -> m) -> Builtin a -> m #

foldMap' :: Monoid m => (a -> m) -> Builtin a -> m #

foldr :: (a -> b -> b) -> b -> Builtin a -> b #

foldr' :: (a -> b -> b) -> b -> Builtin a -> b #

foldl :: (b -> a -> b) -> b -> Builtin a -> b #

foldl' :: (b -> a -> b) -> b -> Builtin a -> b #

foldr1 :: (a -> a -> a) -> Builtin a -> a #

foldl1 :: (a -> a -> a) -> Builtin a -> a #

toList :: Builtin a -> [a] #

null :: Builtin a -> Bool #

length :: Builtin a -> Int #

elem :: Eq a => a -> Builtin a -> Bool #

maximum :: Ord a => Builtin a -> a #

minimum :: Ord a => Builtin a -> a #

sum :: Num a => Builtin a -> a #

product :: Num a => Builtin a -> a #

Traversable Builtin Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

traverse :: Applicative f => (a -> f b) -> Builtin a -> f (Builtin b) #

sequenceA :: Applicative f => Builtin (f a) -> f (Builtin a) #

mapM :: Monad m => (a -> m b) -> Builtin a -> m (Builtin b) #

sequence :: Monad m => Builtin (m a) -> m (Builtin a) #

NamesIn a => NamesIn (Builtin a) Source # 
Instance details

Defined in Agda.Syntax.Internal.Names

Methods

namesAndMetasIn' :: Monoid m => (Either QName MetaId -> m) -> Builtin a -> m Source #

InstantiateFull a => InstantiateFull (Builtin a) Source # 
Instance details

Defined in Agda.TypeChecking.Reduce

EmbPrj a => EmbPrj (Builtin a) Source # 
Instance details

Defined in Agda.TypeChecking.Serialise.Instances.Internal

NFData pf => NFData (Builtin pf) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: Builtin pf -> () #

Generic (Builtin pf) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep (Builtin pf) 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep (Builtin pf) = D1 ('MetaData "Builtin" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "Builtin" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)) :+: (C1 ('MetaCons "Prim" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 pf)) :+: C1 ('MetaCons "BuiltinRewriteRelations" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Set QName)))))

Methods

from :: Builtin pf -> Rep (Builtin pf) x #

to :: Rep (Builtin pf) x -> Builtin pf #

Show pf => Show (Builtin pf) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

showsPrec :: Int -> Builtin pf -> ShowS #

show :: Builtin pf -> String #

showList :: [Builtin pf] -> ShowS #

type Rep (Builtin pf) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep (Builtin pf) = D1 ('MetaData "Builtin" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "Builtin" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)) :+: (C1 ('MetaCons "Prim" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 pf)) :+: C1 ('MetaCons "BuiltinRewriteRelations" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Set QName)))))

data BuiltinDescriptor Source #

Constructors

BuiltinData (TCM Type) [BuiltinId] 
BuiltinDataCons (TCM Type) 
BuiltinPrim PrimitiveId (Term -> TCM ()) 
BuiltinSort BuiltinSort 
BuiltinPostulate Relevance (TCM Type) 
BuiltinUnknown (Maybe (TCM Type)) (Term -> Type -> TCM ())

Builtin of any kind. Type can be checked (Just t) or inferred (Nothing). The second argument is the hook for the verification function.

data BuiltinSort Source #

Instances

Instances details
KillRange BuiltinSort Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

EmbPrj BuiltinSort Source # 
Instance details

Defined in Agda.TypeChecking.Serialise.Instances.Internal

NFData BuiltinSort Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: BuiltinSort -> () #

Generic BuiltinSort Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep BuiltinSort 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep BuiltinSort = D1 ('MetaData "BuiltinSort" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) ((C1 ('MetaCons "SortUniv" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Univ)) :+: C1 ('MetaCons "SortOmega" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Univ))) :+: (C1 ('MetaCons "SortIntervalUniv" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "SortLevelUniv" 'PrefixI 'False) (U1 :: Type -> Type)))
Show BuiltinSort Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Eq BuiltinSort Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep BuiltinSort Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep BuiltinSort = D1 ('MetaData "BuiltinSort" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) ((C1 ('MetaCons "SortUniv" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Univ)) :+: C1 ('MetaCons "SortOmega" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Univ))) :+: (C1 ('MetaCons "SortIntervalUniv" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "SortLevelUniv" 'PrefixI 'False) (U1 :: Type -> Type)))

type 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.

data Call Source #

Constructors

CheckClause Type SpineClause 
CheckLHS SpineLHS 
CheckPattern Pattern Telescope Type 
CheckPatternLinearityType Name 
CheckPatternLinearityValue Name 
CheckLetBinding LetBinding 
InferExpr Expr 
CheckExprCall Comparison Expr Type 
CheckDotPattern Expr Term 
CheckProjection Range QName Type 
IsTypeCall Comparison Expr Sort 
IsType_ Expr 
InferVar Name 
InferDef QName 
CheckArguments Expr [NamedArg Expr] Type (Maybe Type) 
CheckMetaSolution Range MetaId Type Term 
CheckTargetType Range Type Type 
CheckDataDef Range QName [LamBinding] [Constructor] 
CheckRecDef Range QName [LamBinding] [Constructor] 
CheckConstructor QName Telescope Sort Constructor 
CheckConArgFitsIn QName Bool Type Sort 
CheckFunDefCall Range QName [Clause] Bool

Highlight (interactively) if and only if the boolean is True.

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 φ ⊢ f vs = rhs for now, but we store the simplifications of f vs[φ] and rhs[φ].

Fields

  • Range

    Clause range

  • QName

    Function name

  • Term

    (As-is) Function applied to the patterns in this clause

  • Term

    (Simplified) Function applied to the patterns in this clause

  • Term

    (Simplified) clause RHS

  • Type

    (Simplified) clause type

ScopeCheckExpr Expr 
ScopeCheckDeclaration NiceDeclaration 
ScopeCheckLHS QName Pattern 
NoHighlighting 
ModuleContents

Interaction command: show module contents.

SetRange Range

used by setCurrentRange

Instances

Instances details
Pretty Call Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

HasRange Call Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

getRange :: Call -> Range Source #

PrettyTCM Call Source # 
Instance details

Defined in Agda.TypeChecking.Pretty.Call

Methods

prettyTCM :: MonadPretty m => Call -> m Doc Source #

NFData Call Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: Call -> () #

Generic Call Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep Call 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep Call = D1 ('MetaData "Call" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (((((C1 ('MetaCons "CheckClause" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 SpineClause)) :+: C1 ('MetaCons "CheckLHS" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 SpineLHS))) :+: (C1 ('MetaCons "CheckPattern" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Pattern) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Telescope) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type))) :+: C1 ('MetaCons "CheckPatternLinearityType" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Name)))) :+: ((C1 ('MetaCons "CheckPatternLinearityValue" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Name)) :+: C1 ('MetaCons "CheckLetBinding" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 LetBinding))) :+: (C1 ('MetaCons "InferExpr" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr)) :+: (C1 ('MetaCons "CheckExprCall" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type))) :+: C1 ('MetaCons "CheckDotPattern" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)))))) :+: (((C1 ('MetaCons "CheckProjection" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type))) :+: C1 ('MetaCons "IsTypeCall" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Sort)))) :+: (C1 ('MetaCons "IsType_" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr)) :+: C1 ('MetaCons "InferVar" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Name)))) :+: ((C1 ('MetaCons "InferDef" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "CheckArguments" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [NamedArg Expr])) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe Type))))) :+: (C1 ('MetaCons "CheckMetaSolution" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 MetaId)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term))) :+: (C1 ('MetaCons "CheckTargetType" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type))) :+: C1 ('MetaCons "CheckDataDef" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [LamBinding]) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Constructor])))))))) :+: ((((C1 ('MetaCons "CheckRecDef" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [LamBinding]) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Constructor]))) :+: C1 ('MetaCons "CheckConstructor" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Telescope)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Sort) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Constructor)))) :+: (C1 ('MetaCons "CheckConArgFitsIn" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Sort))) :+: C1 ('MetaCons "CheckFunDefCall" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Clause]) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool))))) :+: ((C1 ('MetaCons "CheckPragma" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Pragma)) :+: C1 ('MetaCons "CheckPrimitive" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr)))) :+: (C1 ('MetaCons "CheckIsEmpty" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)) :+: (C1 ('MetaCons "CheckConfluence" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "CheckModuleParameters" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ModuleName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Telescope)))))) :+: (((C1 ('MetaCons "CheckWithFunctionType" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)) :+: C1 ('MetaCons "CheckSectionApplication" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Erased)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ModuleName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ModuleApplication)))) :+: (C1 ('MetaCons "CheckNamedWhere" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ModuleName)) :+: (C1 ('MetaCons "CheckIApplyConfluence" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term))) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)))) :+: C1 ('MetaCons "ScopeCheckExpr" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr))))) :+: ((C1 ('MetaCons "ScopeCheckDeclaration" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 NiceDeclaration)) :+: C1 ('MetaCons "ScopeCheckLHS" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Pattern))) :+: (C1 ('MetaCons "NoHighlighting" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "ModuleContents" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "SetRange" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range))))))))

Methods

from :: Call -> Rep Call x #

to :: Rep Call x -> Call #

type Rep Call Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep Call = D1 ('MetaData "Call" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (((((C1 ('MetaCons "CheckClause" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 SpineClause)) :+: C1 ('MetaCons "CheckLHS" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 SpineLHS))) :+: (C1 ('MetaCons "CheckPattern" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Pattern) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Telescope) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type))) :+: C1 ('MetaCons "CheckPatternLinearityType" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Name)))) :+: ((C1 ('MetaCons "CheckPatternLinearityValue" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Name)) :+: C1 ('MetaCons "CheckLetBinding" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 LetBinding))) :+: (C1 ('MetaCons "InferExpr" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr)) :+: (C1 ('MetaCons "CheckExprCall" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type))) :+: C1 ('MetaCons "CheckDotPattern" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)))))) :+: (((C1 ('MetaCons "CheckProjection" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type))) :+: C1 ('MetaCons "IsTypeCall" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Sort)))) :+: (C1 ('MetaCons "IsType_" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr)) :+: C1 ('MetaCons "InferVar" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Name)))) :+: ((C1 ('MetaCons "InferDef" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "CheckArguments" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [NamedArg Expr])) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe Type))))) :+: (C1 ('MetaCons "CheckMetaSolution" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 MetaId)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term))) :+: (C1 ('MetaCons "CheckTargetType" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type))) :+: C1 ('MetaCons "CheckDataDef" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [LamBinding]) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Constructor])))))))) :+: ((((C1 ('MetaCons "CheckRecDef" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [LamBinding]) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Constructor]))) :+: C1 ('MetaCons "CheckConstructor" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Telescope)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Sort) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Constructor)))) :+: (C1 ('MetaCons "CheckConArgFitsIn" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Sort))) :+: C1 ('MetaCons "CheckFunDefCall" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Clause]) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool))))) :+: ((C1 ('MetaCons "CheckPragma" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Pragma)) :+: C1 ('MetaCons "CheckPrimitive" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr)))) :+: (C1 ('MetaCons "CheckIsEmpty" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)) :+: (C1 ('MetaCons "CheckConfluence" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "CheckModuleParameters" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ModuleName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Telescope)))))) :+: (((C1 ('MetaCons "CheckWithFunctionType" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)) :+: C1 ('MetaCons "CheckSectionApplication" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Erased)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ModuleName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ModuleApplication)))) :+: (C1 ('MetaCons "CheckNamedWhere" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ModuleName)) :+: (C1 ('MetaCons "CheckIApplyConfluence" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term))) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)))) :+: C1 ('MetaCons "ScopeCheckExpr" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr))))) :+: ((C1 ('MetaCons "ScopeCheckDeclaration" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 NiceDeclaration)) :+: C1 ('MetaCons "ScopeCheckLHS" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Pattern))) :+: (C1 ('MetaCons "NoHighlighting" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "ModuleContents" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "SetRange" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range))))))))

data CallInfo Source #

Information about a call.

Constructors

CallInfo 

Fields

Instances

Instances details
Pretty CallInfo Source #

We only show the name of the callee.

Instance details

Defined in Agda.TypeChecking.Monad.Base

HasRange CallInfo Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

PrettyTCM CallInfo Source # 
Instance details

Defined in Agda.TypeChecking.Pretty.Call

NFData CallInfo Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: CallInfo -> () #

Generic CallInfo Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep CallInfo 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep CallInfo = D1 ('MetaData "CallInfo" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "CallInfo" 'PrefixI 'True) (S1 ('MetaSel ('Just "callInfoTarget") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Just "callInfoCall") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Closure Term))))

Methods

from :: CallInfo -> Rep CallInfo x #

to :: Rep CallInfo x -> CallInfo #

Show CallInfo Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep CallInfo Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep CallInfo = D1 ('MetaData "CallInfo" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "CallInfo" 'PrefixI 'True) (S1 ('MetaSel ('Just "callInfoTarget") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Just "callInfoCall") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Closure Term))))

data Candidate Source #

A candidate solution for an instance meta is a term with its type. It may be the case that the candidate is not fully applied yet or of the wrong type, hence the need for the type.

Instances

Instances details
HasOverlapMode Candidate Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Free Candidate Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

freeVars' :: IsVarSet a c => Candidate -> FreeM a c Source #

PrettyTCM Candidate Source # 
Instance details

Defined in Agda.TypeChecking.Pretty

Instantiate Candidate Source # 
Instance details

Defined in Agda.TypeChecking.Reduce

InstantiateFull Candidate Source # 
Instance details

Defined in Agda.TypeChecking.Reduce

Normalise Candidate Source # 
Instance details

Defined in Agda.TypeChecking.Reduce

Reduce Candidate Source # 
Instance details

Defined in Agda.TypeChecking.Reduce

Simplify Candidate Source # 
Instance details

Defined in Agda.TypeChecking.Reduce

Subst Candidate Source # 
Instance details

Defined in Agda.TypeChecking.Substitute

Associated Types

type SubstArg Candidate 
Instance details

Defined in Agda.TypeChecking.Substitute

NFData Candidate Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: Candidate -> () #

Generic Candidate Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep Candidate 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep Candidate = D1 ('MetaData "Candidate" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "Candidate" 'PrefixI 'True) ((S1 ('MetaSel ('Just "candidateKind") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 CandidateKind) :*: S1 ('MetaSel ('Just "candidateTerm") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)) :*: (S1 ('MetaSel ('Just "candidateType") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type) :*: S1 ('MetaSel ('Just "candidateOverlap") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 OverlapMode))))
Show Candidate Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Eq Candidate Source # 
Instance details

Defined in Agda.TypeChecking.Substitute

Ord Candidate Source # 
Instance details

Defined in Agda.TypeChecking.Substitute

type SubstArg Candidate Source # 
Instance details

Defined in Agda.TypeChecking.Substitute

type Rep Candidate Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep Candidate = D1 ('MetaData "Candidate" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "Candidate" 'PrefixI 'True) ((S1 ('MetaSel ('Just "candidateKind") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 CandidateKind) :*: S1 ('MetaSel ('Just "candidateTerm") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)) :*: (S1 ('MetaSel ('Just "candidateType") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type) :*: S1 ('MetaSel ('Just "candidateOverlap") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 OverlapMode))))

data CandidateKind Source #

Instances

Instances details
NFData CandidateKind Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: CandidateKind -> () #

Generic CandidateKind Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep CandidateKind 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep CandidateKind = D1 ('MetaData "CandidateKind" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "LocalCandidate" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "GlobalCandidate" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)))
Show CandidateKind Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Eq CandidateKind Source # 
Instance details

Defined in Agda.TypeChecking.Substitute

Ord CandidateKind Source # 
Instance details

Defined in Agda.TypeChecking.Substitute

type Rep CandidateKind Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep CandidateKind = D1 ('MetaData "CandidateKind" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "LocalCandidate" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "GlobalCandidate" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)))

data CannotQuote Source #

Extra information for error CannotQuote.

Constructors

CannotQuoteAmbiguous (List2 QName)

quote is applied to an ambiguous name.

CannotQuoteExpression Expr

quote is applied to an expression that is not an unambiguous defined name.

CannotQuoteHidden

quote is applied to a non-visible argument.

CannotQuoteNothing

quote is unapplied.

CannotQuotePattern (NamedArg Pattern)

quote is applied to a pattern that is not an unambiguous defined name.

Instances

Instances details
NFData CannotQuote Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: CannotQuote -> () #

Generic CannotQuote Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep CannotQuote 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep CannotQuote = D1 ('MetaData "CannotQuote" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) ((C1 ('MetaCons "CannotQuoteAmbiguous" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List2 QName))) :+: C1 ('MetaCons "CannotQuoteExpression" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr))) :+: (C1 ('MetaCons "CannotQuoteHidden" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "CannotQuoteNothing" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "CannotQuotePattern" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (NamedArg Pattern))))))
Show CannotQuote Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep CannotQuote Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep CannotQuote = D1 ('MetaData "CannotQuote" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) ((C1 ('MetaCons "CannotQuoteAmbiguous" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List2 QName))) :+: C1 ('MetaCons "CannotQuoteExpression" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr))) :+: (C1 ('MetaCons "CannotQuoteHidden" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "CannotQuoteNothing" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "CannotQuotePattern" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (NamedArg Pattern))))))

data CheckedArg Source #

Constructors

CheckedArg 

Fields

Instances

Instances details
Show CheckedArg Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

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.

newtype CheckpointId Source #

Constructors

CheckpointId Int 

Instances

Instances details
Pretty CheckpointId Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

HasFresh CheckpointId Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

PrettyTCM CheckpointId Source # 
Instance details

Defined in Agda.TypeChecking.Pretty

EmbPrj CheckpointId Source # 
Instance details

Defined in Agda.TypeChecking.Serialise.Instances.Internal

NFData CheckpointId Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: CheckpointId -> () #

Enum CheckpointId Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Num CheckpointId Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Integral CheckpointId Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Real CheckpointId Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Show CheckpointId Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Eq CheckpointId Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Ord CheckpointId Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

data Closure a Source #

Instances

Instances details
Functor Closure Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

fmap :: (a -> b) -> Closure a -> Closure b #

(<$) :: a -> Closure b -> Closure a #

Foldable Closure Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

fold :: Monoid m => Closure m -> m #

foldMap :: Monoid m => (a -> m) -> Closure a -> m #

foldMap' :: Monoid m => (a -> m) -> Closure a -> m #

foldr :: (a -> b -> b) -> b -> Closure a -> b #

foldr' :: (a -> b -> b) -> b -> Closure a -> b #

foldl :: (b -> a -> b) -> b -> Closure a -> b #

foldl' :: (b -> a -> b) -> b -> Closure a -> b #

foldr1 :: (a -> a -> a) -> Closure a -> a #

foldl1 :: (a -> a -> a) -> Closure a -> a #

toList :: Closure a -> [a] #

null :: Closure a -> Bool #

length :: Closure a -> Int #

elem :: Eq a => a -> Closure a -> Bool #

maximum :: Ord a => Closure a -> a #

minimum :: Ord a => Closure a -> a #

sum :: Num a => Closure a -> a #

product :: Num a => Closure a -> a #

LensIsAbstract (Closure a) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

HasRange a => HasRange (Closure a) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

getRange :: Closure a -> Range Source #

KillRange a => KillRange (Closure a) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

MentionsMeta a => MentionsMeta (Closure a) Source # 
Instance details

Defined in Agda.TypeChecking.MetaVars.Mention

LensTCEnv (Closure a) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

PrettyTCM a => PrettyTCM (Closure a) Source # 
Instance details

Defined in Agda.TypeChecking.Pretty

Methods

prettyTCM :: MonadPretty m => Closure a -> m Doc Source #

Instantiate a => Instantiate (Closure a) Source # 
Instance details

Defined in Agda.TypeChecking.Reduce

InstantiateFull a => InstantiateFull (Closure a) Source # 
Instance details

Defined in Agda.TypeChecking.Reduce

Normalise a => Normalise (Closure a) Source # 
Instance details

Defined in Agda.TypeChecking.Reduce

Reduce a => Reduce (Closure a) Source # 
Instance details

Defined in Agda.TypeChecking.Reduce

Simplify a => Simplify (Closure a) Source # 
Instance details

Defined in Agda.TypeChecking.Reduce

Methods

simplify' :: Closure a -> ReduceM (Closure a)

NFData a => NFData (Closure a) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: Closure a -> () #

Generic (Closure a) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep (Closure a) 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

from :: Closure a -> Rep (Closure a) x #

to :: Rep (Closure a) x -> Closure a #

Show a => Show (Closure a) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

showsPrec :: Int -> Closure a -> ShowS #

show :: Closure a -> String #

showList :: [Closure a] -> ShowS #

LensClosure (Closure a) a Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep (Closure a) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

data CompKit Source #

Constructors

CompKit 

Instances

Instances details
NamesIn CompKit Source # 
Instance details

Defined in Agda.Syntax.Internal.Names

Methods

namesAndMetasIn' :: Monoid m => (Either QName MetaId -> m) -> CompKit -> m Source #

KillRange CompKit Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

EmbPrj CompKit Source # 
Instance details

Defined in Agda.TypeChecking.Serialise.Instances.Internal

NFData CompKit Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: CompKit -> () #

Generic CompKit Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep CompKit 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep CompKit = D1 ('MetaData "CompKit" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "CompKit" 'PrefixI 'True) (S1 ('MetaSel ('Just "nameOfHComp") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe QName)) :*: S1 ('MetaSel ('Just "nameOfTransp") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe QName))))

Methods

from :: CompKit -> Rep CompKit x #

to :: Rep CompKit x -> CompKit #

Show CompKit Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Eq CompKit Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

(==) :: CompKit -> CompKit -> Bool #

(/=) :: CompKit -> CompKit -> Bool #

Ord CompKit Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep CompKit Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep CompKit = D1 ('MetaData "CompKit" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "CompKit" 'PrefixI 'True) (S1 ('MetaSel ('Just "nameOfHComp") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe QName)) :*: S1 ('MetaSel ('Just "nameOfTransp") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe QName))))

data CompareAs Source #

We can either compare two terms at a given type, or compare two types without knowing (or caring about) their sorts.

Constructors

AsTermsOf Type

Type should not be Size. But currently, we do not rely on this invariant.

AsSizes

Replaces AsTermsOf Size.

AsTypes 

Instances

Instances details
Pretty CompareAs Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

TermLike CompareAs Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

traverseTermM :: Monad m => (Term -> m Term) -> CompareAs -> m CompareAs Source #

foldTerm :: Monoid m => (Term -> m) -> CompareAs -> m Source #

AllMetas CompareAs Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

allMetas :: Monoid m => (MetaId -> m) -> CompareAs -> m Source #

Free CompareAs Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

freeVars' :: IsVarSet a c => CompareAs -> FreeM a c Source #

MentionsMeta CompareAs Source # 
Instance details

Defined in Agda.TypeChecking.MetaVars.Mention

IsSizeType CompareAs Source # 
Instance details

Defined in Agda.TypeChecking.Monad.SizedTypes

PrettyTCM CompareAs Source # 
Instance details

Defined in Agda.TypeChecking.Pretty

Instantiate CompareAs Source # 
Instance details

Defined in Agda.TypeChecking.Reduce

InstantiateFull CompareAs Source # 
Instance details

Defined in Agda.TypeChecking.Reduce

IsMeta CompareAs Source # 
Instance details

Defined in Agda.TypeChecking.Reduce

Normalise CompareAs Source # 
Instance details

Defined in Agda.TypeChecking.Reduce

Reduce CompareAs Source # 
Instance details

Defined in Agda.TypeChecking.Reduce

Simplify CompareAs Source # 
Instance details

Defined in Agda.TypeChecking.Reduce

Subst CompareAs Source # 
Instance details

Defined in Agda.TypeChecking.Substitute

Associated Types

type SubstArg CompareAs 
Instance details

Defined in Agda.TypeChecking.Substitute

NFData CompareAs Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: CompareAs -> () #

Generic CompareAs Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep CompareAs 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep CompareAs = D1 ('MetaData "CompareAs" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "AsTermsOf" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)) :+: (C1 ('MetaCons "AsSizes" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "AsTypes" 'PrefixI 'False) (U1 :: Type -> Type)))
Show CompareAs Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type SubstArg CompareAs Source # 
Instance details

Defined in Agda.TypeChecking.Substitute

type Rep CompareAs Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep CompareAs = D1 ('MetaData "CompareAs" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "AsTermsOf" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)) :+: (C1 ('MetaCons "AsSizes" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "AsTypes" 'PrefixI 'False) (U1 :: Type -> Type)))

data CompilerPragma Source #

The backends are responsible for parsing their own pragmas.

Instances

Instances details
HasRange CompilerPragma Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

KillRange CompiledRepresentation Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

EmbPrj CompilerPragma Source # 
Instance details

Defined in Agda.TypeChecking.Serialise.Instances.Compilers

NFData CompilerPragma Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: CompilerPragma -> () #

Generic CompilerPragma Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep CompilerPragma 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep CompilerPragma = D1 ('MetaData "CompilerPragma" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "CompilerPragma" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String)))
Show CompilerPragma Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Eq CompilerPragma Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep CompilerPragma Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep CompilerPragma = D1 ('MetaData "CompilerPragma" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "CompilerPragma" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String)))

data Constraint Source #

Constructors

ValueCmp Comparison CompareAs Term Term 
ValueCmpOnFace Comparison Term Type Term Term 
ElimCmp [Polarity] [IsForced] Type Term [Elim] [Elim] 
SortCmp Comparison Sort Sort 
LevelCmp Comparison Level Level 
HasBiggerSort Sort 
HasPTSRule (Dom Type) (Abs Sort) 
CheckDataSort QName Sort

Check that the sort Sort of data type QName admits data/record types. E.g., sorts IUniv, SizeUniv etc. do not admit such constructions. See checkDataSort.

CheckMetaInst MetaId 
CheckType Type 
UnBlock MetaId

Meta created for a term blocked by a postponed type checking problem or unsolved constraints. The MetaInstantiation for the meta (when unsolved) is either BlockedConst or PostponedTypeCheckingProblem.

IsEmpty Range Type

The range is the one of the absurd pattern.

CheckSizeLtSat Term

Check that the Term is either not a SIZELT or a non-empty SIZELT.

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

CheckLockedVars t ty lk lk_ty with t : ty, lk : lk_ty and t lk well-typed.

UsableAtModality WhyCheckModality (Maybe Sort) Modality Term

Is the term usable at the given modality? This check should run if the Sort is Nothing or isFibrant.

Instances

Instances details
TermLike Constraint Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

traverseTermM :: Monad m => (Term -> m Term) -> Constraint -> m Constraint Source #

foldTerm :: Monoid m => (Term -> m) -> Constraint -> m Source #

AllMetas Constraint Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

allMetas :: Monoid m => (MetaId -> m) -> Constraint -> m Source #

HasRange Constraint Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Reify Constraint Source # 
Instance details

Defined in Agda.Interaction.BasicOps

Associated Types

type ReifiesTo Constraint 
Instance details

Defined in Agda.Interaction.BasicOps

Free Constraint Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

freeVars' :: IsVarSet a c => Constraint -> FreeM a c Source #

MentionsMeta Constraint Source # 
Instance details

Defined in Agda.TypeChecking.MetaVars.Mention

PrettyTCM Constraint Source # 
Instance details

Defined in Agda.TypeChecking.Pretty.Constraint

Instantiate Constraint Source # 
Instance details

Defined in Agda.TypeChecking.Reduce

InstantiateFull Constraint Source # 
Instance details

Defined in Agda.TypeChecking.Reduce

Normalise Constraint Source # 
Instance details

Defined in Agda.TypeChecking.Reduce

Reduce Constraint Source # 
Instance details

Defined in Agda.TypeChecking.Reduce

Simplify Constraint Source # 
Instance details

Defined in Agda.TypeChecking.Reduce

Subst Constraint Source # 
Instance details

Defined in Agda.TypeChecking.Substitute

Associated Types

type SubstArg Constraint 
Instance details

Defined in Agda.TypeChecking.Substitute

NFData Constraint Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: Constraint -> () #

Generic Constraint Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep Constraint 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep Constraint = D1 ('MetaData "Constraint" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) ((((C1 ('MetaCons "ValueCmp" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 CompareAs)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term))) :+: C1 ('MetaCons "ValueCmpOnFace" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term))))) :+: (C1 ('MetaCons "ElimCmp" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Polarity]) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [IsForced]) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type))) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Elim]) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Elim])))) :+: C1 ('MetaCons "SortCmp" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Sort) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Sort))))) :+: ((C1 ('MetaCons "LevelCmp" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Level) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Level))) :+: C1 ('MetaCons "HasBiggerSort" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Sort))) :+: (C1 ('MetaCons "HasPTSRule" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Dom Type)) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Abs Sort))) :+: (C1 ('MetaCons "CheckDataSort" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Sort)) :+: C1 ('MetaCons "CheckMetaInst" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 MetaId)))))) :+: (((C1 ('MetaCons "CheckType" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)) :+: C1 ('MetaCons "UnBlock" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 MetaId))) :+: (C1 ('MetaCons "IsEmpty" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)) :+: (C1 ('MetaCons "CheckSizeLtSat" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)) :+: C1 ('MetaCons "FindInstance" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 MetaId) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe [Candidate])))))) :+: ((C1 ('MetaCons "ResolveInstanceHead" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "CheckFunDef" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 DefInfo) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Clause]) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 TCErr)))) :+: (C1 ('MetaCons "UnquoteTactic" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type))) :+: (C1 ('MetaCons "CheckLockedVars" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Arg Term)) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type))) :+: C1 ('MetaCons "UsableAtModality" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 WhyCheckModality) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe Sort))) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Modality) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term))))))))
Show Constraint Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type ReifiesTo Constraint Source # 
Instance details

Defined in Agda.Interaction.BasicOps

type SubstArg Constraint Source # 
Instance details

Defined in Agda.TypeChecking.Substitute

type Rep Constraint Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep Constraint = D1 ('MetaData "Constraint" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) ((((C1 ('MetaCons "ValueCmp" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 CompareAs)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term))) :+: C1 ('MetaCons "ValueCmpOnFace" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term))))) :+: (C1 ('MetaCons "ElimCmp" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Polarity]) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [IsForced]) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type))) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Elim]) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Elim])))) :+: C1 ('MetaCons "SortCmp" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Sort) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Sort))))) :+: ((C1 ('MetaCons "LevelCmp" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Level) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Level))) :+: C1 ('MetaCons "HasBiggerSort" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Sort))) :+: (C1 ('MetaCons "HasPTSRule" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Dom Type)) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Abs Sort))) :+: (C1 ('MetaCons "CheckDataSort" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Sort)) :+: C1 ('MetaCons "CheckMetaInst" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 MetaId)))))) :+: (((C1 ('MetaCons "CheckType" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)) :+: C1 ('MetaCons "UnBlock" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 MetaId))) :+: (C1 ('MetaCons "IsEmpty" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)) :+: (C1 ('MetaCons "CheckSizeLtSat" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)) :+: C1 ('MetaCons "FindInstance" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 MetaId) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe [Candidate])))))) :+: ((C1 ('MetaCons "ResolveInstanceHead" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "CheckFunDef" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 DefInfo) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Clause]) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 TCErr)))) :+: (C1 ('MetaCons "UnquoteTactic" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type))) :+: (C1 ('MetaCons "CheckLockedVars" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Arg Term)) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type))) :+: C1 ('MetaCons "UsableAtModality" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 WhyCheckModality) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe Sort))) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Modality) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term))))))))

data ConstructorData Source #

Constructors

ConstructorData 

Fields

  • _conPars :: Int

    Number of parameters.

  • _conArity :: Int

    Number of arguments (excluding parameters).

  • _conSrcCon :: ConHead

    Name of (original) constructor and fields. (This might be in a module instance.)

  • _conData :: QName

    Name of datatype or record type.

  • _conAbstr :: IsAbstract
     
  • _conComp :: CompKit

    Cubical composition.

  • _conProj :: Maybe [QName]

    Projections. Nothing if not yet computed.

  • _conForced :: [IsForced]

    Which arguments are forced (i.e. determined by the type of the constructor)? Either this list is empty (if the forcing analysis isn't run), or its length is conArity.

  • _conErased :: Maybe [Bool]

    Which arguments are erased at runtime (computed during compilation to treeless)? True means erased, False means retained. Nothing if no erasure analysis has been performed yet. The length of the list is conArity.

  • _conErasure :: !Bool

    Was --erasure in effect when the constructor was defined? (This can affect the constructor's type.)

  • _conInline :: !Bool

    Shall we translate the constructor on the root of the rhs into copattern matching on the lhs? Activated by INLINE pragma.

Instances

Instances details
Pretty ConstructorData Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

NFData ConstructorData Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: ConstructorData -> () #

Generic ConstructorData Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Show ConstructorData Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep ConstructorData Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

data ConstructorDisambiguationData Source #

Information we have constructored in the middle of disambiguating a constructor.

Constructors

ConstructorDisambiguationData 

Fields

Instances

Instances details
NFData ConstructorDisambiguationData Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Generic ConstructorDisambiguationData Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep ConstructorDisambiguationData 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep ConstructorDisambiguationData = D1 ('MetaData "ConstructorDisambiguationData" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "ConstructorDisambiguationData" 'PrefixI 'True) ((S1 ('MetaSel ('Just "bcdConName") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Just "bcdCandidates") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 (QName, Type, ConHead)))) :*: (S1 ('MetaSel ('Just "bcdArguments") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Args) :*: S1 ('MetaSel ('Just "bcdType") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type))))
type Rep ConstructorDisambiguationData Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep ConstructorDisambiguationData = D1 ('MetaData "ConstructorDisambiguationData" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "ConstructorDisambiguationData" 'PrefixI 'True) ((S1 ('MetaSel ('Just "bcdConName") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Just "bcdCandidates") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 (QName, Type, ConHead)))) :*: (S1 ('MetaSel ('Just "bcdArguments") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Args) :*: S1 ('MetaSel ('Just "bcdType") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type))))

type CurrentTypeCheckLog = [(TypeCheckAction, PostScopeState)] Source #

Like CachedTypeCheckLog, but storing the log for an ongoing type checking of a module. Stored in reverse order (last performed action first).

data DataOrRecSigData Source #

Constructors

DataOrRecSigData 

Fields

Instances

Instances details
Pretty DataOrRecSigData Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

NFData DataOrRecSigData Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: DataOrRecSigData -> () #

Generic DataOrRecSigData Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep DataOrRecSigData 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep DataOrRecSigData = D1 ('MetaData "DataOrRecSigData" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "DataOrRecSigData" 'PrefixI 'True) (S1 ('MetaSel ('Just "_datarecPars") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Int)))
Show DataOrRecSigData Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep DataOrRecSigData Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep DataOrRecSigData = D1 ('MetaData "DataOrRecSigData" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "DataOrRecSigData" 'PrefixI 'True) (S1 ('MetaSel ('Just "_datarecPars") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Int)))

data DatatypeData Source #

Constructors

DatatypeData 

Fields

Instances

Instances details
Pretty DatatypeData Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

NFData DatatypeData Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: DatatypeData -> () #

Generic DatatypeData Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Show DatatypeData Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep DatatypeData Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

data Definition Source #

Constructors

Defn 

Fields

Instances

Instances details
LensArgInfo Definition Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

LensModalPolarity Definition Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

LensModality Definition Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

LensQuantity Definition Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

LensRelevance Definition Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Pretty Definition Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

NamesIn Definition Source # 
Instance details

Defined in Agda.Syntax.Internal.Names

KillRange Definition Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

KillRange Definitions Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

InstantiateFull Definition Source # 
Instance details

Defined in Agda.TypeChecking.Reduce

EmbPrj Definition Source # 
Instance details

Defined in Agda.TypeChecking.Serialise.Instances.Internal

Abstract Definition Source # 
Instance details

Defined in Agda.TypeChecking.Substitute

Apply Definition Source # 
Instance details

Defined in Agda.TypeChecking.Substitute

NFData Definition Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: Definition -> () #

Generic Definition Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep Definition 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep Definition = D1 ('MetaData "Definition" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "Defn" 'PrefixI 'True) ((((S1 ('MetaSel ('Just "defArgInfo") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ArgInfo) :*: S1 ('MetaSel ('Just "defName") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :*: (S1 ('MetaSel ('Just "defType") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type) :*: S1 ('MetaSel ('Just "defPolarity") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Polarity]))) :*: ((S1 ('MetaSel ('Just "defArgOccurrences") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Occurrence]) :*: S1 ('MetaSel ('Just "defGeneralizedParams") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Maybe Name])) :*: (S1 ('MetaSel ('Just "defDisplay") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [LocalDisplayForm]) :*: (S1 ('MetaSel ('Just "defMutual") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 MutualId) :*: S1 ('MetaSel ('Just "defCompiledRep") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 CompiledRepresentation))))) :*: (((S1 ('MetaSel ('Just "defInstance") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe InstanceInfo)) :*: S1 ('MetaSel ('Just "defCopy") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool)) :*: (S1 ('MetaSel ('Just "defMatchable") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Set QName)) :*: S1 ('MetaSel ('Just "defNoCompilation") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool))) :*: ((S1 ('MetaSel ('Just "defInjective") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool) :*: S1 ('MetaSel ('Just "defCopatternLHS") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool)) :*: (S1 ('MetaSel ('Just "defBlocked") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Blocked_) :*: (S1 ('MetaSel ('Just "defLanguage") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Language) :*: S1 ('MetaSel ('Just "theDef") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Defn)))))))
Show Definition Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep Definition Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep Definition = D1 ('MetaData "Definition" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "Defn" 'PrefixI 'True) ((((S1 ('MetaSel ('Just "defArgInfo") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ArgInfo) :*: S1 ('MetaSel ('Just "defName") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :*: (S1 ('MetaSel ('Just "defType") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type) :*: S1 ('MetaSel ('Just "defPolarity") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Polarity]))) :*: ((S1 ('MetaSel ('Just "defArgOccurrences") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Occurrence]) :*: S1 ('MetaSel ('Just "defGeneralizedParams") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Maybe Name])) :*: (S1 ('MetaSel ('Just "defDisplay") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [LocalDisplayForm]) :*: (S1 ('MetaSel ('Just "defMutual") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 MutualId) :*: S1 ('MetaSel ('Just "defCompiledRep") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 CompiledRepresentation))))) :*: (((S1 ('MetaSel ('Just "defInstance") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe InstanceInfo)) :*: S1 ('MetaSel ('Just "defCopy") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool)) :*: (S1 ('MetaSel ('Just "defMatchable") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Set QName)) :*: S1 ('MetaSel ('Just "defNoCompilation") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool))) :*: ((S1 ('MetaSel ('Just "defInjective") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool) :*: S1 ('MetaSel ('Just "defCopatternLHS") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool)) :*: (S1 ('MetaSel ('Just "defBlocked") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Blocked_) :*: (S1 ('MetaSel ('Just "defLanguage") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Language) :*: S1 ('MetaSel ('Just "theDef") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Defn)))))))

data Defn Source #

Constructors

AxiomDefn AxiomData

Postulate.

DataOrRecSigDefn DataOrRecSigData

Data or record type signature that doesn't yet have a definition.

GeneralizableVar

Generalizable variable (introduced in variable block).

Fields

AbstractDefn Defn

Returned by getConstInfo if definition is abstract.

FunctionDefn FunctionData 
DatatypeDefn DatatypeData 
RecordDefn RecordData 
ConstructorDefn ConstructorData 
PrimitiveDefn PrimitiveData

Primitive or builtin functions.

PrimitiveSortDefn PrimitiveSortData 

Instances

Instances details
Pretty Defn Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

NamesIn Defn Source # 
Instance details

Defined in Agda.Syntax.Internal.Names

Methods

namesAndMetasIn' :: Monoid m => (Either QName MetaId -> m) -> Defn -> m Source #

KillRange Defn Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Occurs Defn Source # 
Instance details

Defined in Agda.TypeChecking.MetaVars.Occurs

InstantiateFull Defn Source # 
Instance details

Defined in Agda.TypeChecking.Reduce

EmbPrj Defn Source # 
Instance details

Defined in Agda.TypeChecking.Serialise.Instances.Internal

Abstract Defn Source # 
Instance details

Defined in Agda.TypeChecking.Substitute

Apply Defn Source # 
Instance details

Defined in Agda.TypeChecking.Substitute

Methods

apply :: Defn -> Args -> Defn Source #

applyE :: Defn -> Elims -> Defn Source #

NFData Defn Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: Defn -> () #

Generic Defn Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep Defn 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep Defn = D1 ('MetaData "Defn" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (((C1 ('MetaCons "AxiomDefn" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 AxiomData)) :+: C1 ('MetaCons "DataOrRecSigDefn" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 DataOrRecSigData))) :+: (C1 ('MetaCons "GeneralizableVar" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 NumGeneralizableArgs)) :+: (C1 ('MetaCons "AbstractDefn" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Defn)) :+: C1 ('MetaCons "FunctionDefn" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 FunctionData))))) :+: ((C1 ('MetaCons "DatatypeDefn" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 DatatypeData)) :+: C1 ('MetaCons "RecordDefn" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 RecordData))) :+: (C1 ('MetaCons "ConstructorDefn" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ConstructorData)) :+: (C1 ('MetaCons "PrimitiveDefn" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 PrimitiveData)) :+: C1 ('MetaCons "PrimitiveSortDefn" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 PrimitiveSortData))))))

Methods

from :: Defn -> Rep Defn x #

to :: Rep Defn x -> Defn #

Show Defn Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

showsPrec :: Int -> Defn -> ShowS #

show :: Defn -> String #

showList :: [Defn] -> ShowS #

type Rep Defn Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep Defn = D1 ('MetaData "Defn" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (((C1 ('MetaCons "AxiomDefn" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 AxiomData)) :+: C1 ('MetaCons "DataOrRecSigDefn" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 DataOrRecSigData))) :+: (C1 ('MetaCons "GeneralizableVar" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 NumGeneralizableArgs)) :+: (C1 ('MetaCons "AbstractDefn" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Defn)) :+: C1 ('MetaCons "FunctionDefn" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 FunctionData))))) :+: ((C1 ('MetaCons "DatatypeDefn" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 DatatypeData)) :+: C1 ('MetaCons "RecordDefn" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 RecordData))) :+: (C1 ('MetaCons "ConstructorDefn" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ConstructorData)) :+: (C1 ('MetaCons "PrimitiveDefn" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 PrimitiveData)) :+: C1 ('MetaCons "PrimitiveSortDefn" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 PrimitiveSortData))))))

data DisambiguatedName Source #

Name disambiguation for the sake of highlighting.

Instances

Instances details
NFData DisambiguatedName Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: DisambiguatedName -> () #

Generic DisambiguatedName Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep DisambiguatedName 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep DisambiguatedName = D1 ('MetaData "DisambiguatedName" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "DisambiguatedName" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 NameKind) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)))
type Rep DisambiguatedName Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep DisambiguatedName = D1 ('MetaData "DisambiguatedName" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "DisambiguatedName" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 NameKind) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)))

data DisplayForm Source #

A DisplayForm is in essence a rewrite rule q ts --> dt for a defined symbol (could be a constructor as well) q. The right hand side is a DisplayTerm which is used to reify to a more readable Syntax.

The patterns ts are just terms, but the first dfPatternVars variables are pattern variables that matches any term.

Constructors

Display 

Fields

  • dfPatternVars :: Nat

    Number n of pattern variables in dfPats.

  • dfPats :: Elims

    Left hand side patterns, the n first free variables are pattern variables, any variables above n are fixed and only match that particular variable. This happens when you have display forms inside parameterised modules that match on the module parameters. The ArgInfo is ignored in these patterns.

  • dfRHS :: DisplayTerm

    Right hand side.

Instances

Instances details
Pretty DisplayForm Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

NamesIn DisplayForm Source # 
Instance details

Defined in Agda.Syntax.Internal.Names

KillRange DisplayForm Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Free DisplayForm Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

freeVars' :: IsVarSet a c => DisplayForm -> FreeM a c Source #

ChaseDisplayForms DisplayForm Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Signature

InstantiateFull DisplayForm Source # 
Instance details

Defined in Agda.TypeChecking.Reduce

Normalise DisplayForm Source # 
Instance details

Defined in Agda.TypeChecking.Reduce

Simplify DisplayForm Source # 
Instance details

Defined in Agda.TypeChecking.Reduce

EmbPrj DisplayForm Source # 
Instance details

Defined in Agda.TypeChecking.Serialise.Instances.Internal

Subst DisplayForm Source # 
Instance details

Defined in Agda.TypeChecking.Substitute

Associated Types

type SubstArg DisplayForm 
Instance details

Defined in Agda.TypeChecking.Substitute

NFData DisplayForm Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: DisplayForm -> () #

Generic DisplayForm Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep DisplayForm 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep DisplayForm = D1 ('MetaData "DisplayForm" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "Display" 'PrefixI 'True) (S1 ('MetaSel ('Just "dfPatternVars") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Nat) :*: (S1 ('MetaSel ('Just "dfPats") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Elims) :*: S1 ('MetaSel ('Just "dfRHS") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 DisplayTerm))))
Show DisplayForm Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type SubstArg DisplayForm Source # 
Instance details

Defined in Agda.TypeChecking.Substitute

type Rep DisplayForm Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep DisplayForm = D1 ('MetaData "DisplayForm" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "Display" 'PrefixI 'True) (S1 ('MetaSel ('Just "dfPatternVars") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Nat) :*: (S1 ('MetaSel ('Just "dfPats") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Elims) :*: S1 ('MetaSel ('Just "dfRHS") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 DisplayTerm))))

data DisplayTerm Source #

A structured presentation of a Term for reification into Syntax.

Constructors

DWithApp DisplayTerm (List1 DisplayTerm) Elims

(f vs | ws) es. The first DisplayTerm is the parent function f with its args vs. The list of DisplayTerms are the with expressions ws. The Elims are additional arguments es (possible in case the with-application is of function type) or projections (if it is of record type).

DCon ConHead ConInfo [Arg DisplayTerm]

c vs.

DDef QName [Elim' DisplayTerm]

d vs.

DDot' Term Elims

.(v es). See DTerm'.

DTerm' Term Elims

v es. This is a frozen elimination that is not always safe to run, because display forms may be ill-typed. (See issue #6476.)

Instances

Instances details
Pretty DisplayTerm Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

NamesIn DisplayTerm Source # 
Instance details

Defined in Agda.Syntax.Internal.Names

KillRange DisplayTerm Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Reify DisplayTerm Source # 
Instance details

Defined in Agda.Syntax.Translation.InternalToAbstract

Associated Types

type ReifiesTo DisplayTerm 
Instance details

Defined in Agda.Syntax.Translation.InternalToAbstract

Free DisplayTerm Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

freeVars' :: IsVarSet a c => DisplayTerm -> FreeM a c Source #

ChaseDisplayForms DisplayTerm Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Signature

PrettyTCM DisplayTerm Source # 
Instance details

Defined in Agda.TypeChecking.Pretty

InstantiateFull DisplayTerm Source # 
Instance details

Defined in Agda.TypeChecking.Reduce

EmbPrj DisplayTerm Source # 
Instance details

Defined in Agda.TypeChecking.Serialise.Instances.Internal

Apply DisplayTerm Source # 
Instance details

Defined in Agda.TypeChecking.Substitute

Subst DisplayTerm Source # 
Instance details

Defined in Agda.TypeChecking.Substitute

Associated Types

type SubstArg DisplayTerm 
Instance details

Defined in Agda.TypeChecking.Substitute

NFData DisplayTerm Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: DisplayTerm -> () #

Generic DisplayTerm Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep DisplayTerm 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep DisplayTerm = D1 ('MetaData "DisplayTerm" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) ((C1 ('MetaCons "DWithApp" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 DisplayTerm) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 DisplayTerm)) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Elims))) :+: C1 ('MetaCons "DCon" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ConHead) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ConInfo) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Arg DisplayTerm])))) :+: (C1 ('MetaCons "DDef" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Elim' DisplayTerm])) :+: (C1 ('MetaCons "DDot'" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Elims)) :+: C1 ('MetaCons "DTerm'" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Elims)))))
Show DisplayTerm Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

PrettyTCM (Elim' DisplayTerm) Source # 
Instance details

Defined in Agda.TypeChecking.Pretty

type ReifiesTo DisplayTerm Source # 
Instance details

Defined in Agda.Syntax.Translation.InternalToAbstract

type SubstArg DisplayTerm Source # 
Instance details

Defined in Agda.TypeChecking.Substitute

type Rep DisplayTerm Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep DisplayTerm = D1 ('MetaData "DisplayTerm" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) ((C1 ('MetaCons "DWithApp" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 DisplayTerm) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 DisplayTerm)) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Elims))) :+: C1 ('MetaCons "DCon" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ConHead) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ConInfo) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Arg DisplayTerm])))) :+: (C1 ('MetaCons "DDef" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Elim' DisplayTerm])) :+: (C1 ('MetaCons "DDot'" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Elims)) :+: C1 ('MetaCons "DTerm'" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Elims)))))

data DoGeneralize Source #

Constructors

YesGeneralizeVar

Generalize because it is a generalizable variable.

YesGeneralizeMeta

Generalize because it is a metavariable and we're currently checking the type of a generalizable variable (this should get the default modality).

NoGeneralize

Don't generalize.

Instances

Instances details
KillRange DoGeneralize Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

EmbPrj DoGeneralize Source # 
Instance details

Defined in Agda.TypeChecking.Serialise.Instances.Internal

NFData DoGeneralize Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: DoGeneralize -> () #

Generic DoGeneralize Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep DoGeneralize 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep DoGeneralize = D1 ('MetaData "DoGeneralize" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "YesGeneralizeVar" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "YesGeneralizeMeta" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "NoGeneralize" 'PrefixI 'False) (U1 :: Type -> Type)))
Show DoGeneralize Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Eq DoGeneralize Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Ord DoGeneralize Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep DoGeneralize Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep DoGeneralize = D1 ('MetaData "DoGeneralize" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "YesGeneralizeVar" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "YesGeneralizeMeta" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "NoGeneralize" 'PrefixI 'False) (U1 :: Type -> Type)))

data EtaEquality Source #

Should a record type admit eta-equality?

Constructors

Specified

User specifed 'eta-equality' or 'no-eta-equality'.

Inferred

Positivity checker inferred whether eta is safe.

Instances

Instances details
CopatternMatchingAllowed EtaEquality Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

PatternMatchingAllowed EtaEquality Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

KillRange EtaEquality Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

EmbPrj EtaEquality Source # 
Instance details

Defined in Agda.TypeChecking.Serialise.Instances.Internal

NFData EtaEquality Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: EtaEquality -> () #

Generic EtaEquality Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep EtaEquality 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep EtaEquality = D1 ('MetaData "EtaEquality" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "Specified" 'PrefixI 'True) (S1 ('MetaSel ('Just "theEtaEquality") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 HasEta)) :+: C1 ('MetaCons "Inferred" 'PrefixI 'True) (S1 ('MetaSel ('Just "theEtaEquality") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 HasEta)))
Show EtaEquality Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Eq EtaEquality Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep EtaEquality Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep EtaEquality = D1 ('MetaData "EtaEquality" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "Specified" 'PrefixI 'True) (S1 ('MetaSel ('Just "theEtaEquality") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 HasEta)) :+: C1 ('MetaCons "Inferred" 'PrefixI 'True) (S1 ('MetaSel ('Just "theEtaEquality") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 HasEta)))

data ExecError Source #

Error when trying to call an external executable during reflection.

Constructors

ExeNotTrusted ExeName ExeMap

The given executable is not listed as trusted.

ExeNotFound ExeName FilePath

The given executable could not be found under the given path.

ExeNotExecutable ExeName FilePath

The given file path does not have executable permissions.

Instances

Instances details
PrettyTCM ExecError Source # 
Instance details

Defined in Agda.TypeChecking.Errors

NFData ExecError Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: ExecError -> () #

Generic ExecError Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Show ExecError Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep ExecError Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

data ExpandHidden Source #

Constructors

ExpandLast

Add implicit arguments in the end until type is no longer hidden Pi.

DontExpandLast

Do not append implicit arguments.

ReallyDontExpandLast

Makes doExpandLast have no effect. Used to avoid implicit insertion of arguments to metavariables.

Instances

Instances details
NFData ExpandHidden Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: ExpandHidden -> () #

Generic ExpandHidden Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep ExpandHidden 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep ExpandHidden = D1 ('MetaData "ExpandHidden" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "ExpandLast" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "DontExpandLast" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "ReallyDontExpandLast" 'PrefixI 'False) (U1 :: Type -> Type)))
Eq ExpandHidden Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep ExpandHidden Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep ExpandHidden = D1 ('MetaData "ExpandHidden" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "ExpandLast" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "DontExpandLast" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "ReallyDontExpandLast" 'PrefixI 'False) (U1 :: Type -> Type)))

data ExtLamInfo Source #

Additional information for extended lambdas.

Constructors

ExtLamInfo 

Fields

  • extLamModule :: ModuleName

    For complicated reasons the scope checker decides the QName of a pattern lambda, and thus its module. We really need to decide the module during type checking though, since if the lambda appears in a refined context the module picked by the scope checker has very much the wrong parameters.

  • extLamAbsurd :: Bool

    Was this definition created from an absurd lambda λ ()?

  • extLamSys :: !(Maybe System)
     

Instances

Instances details
NamesIn ExtLamInfo Source # 
Instance details

Defined in Agda.Syntax.Internal.Names

KillRange ExtLamInfo Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

InstantiateFull ExtLamInfo Source # 
Instance details

Defined in Agda.TypeChecking.Reduce

EmbPrj ExtLamInfo Source # 
Instance details

Defined in Agda.TypeChecking.Serialise.Instances.Internal

Apply ExtLamInfo Source # 
Instance details

Defined in Agda.TypeChecking.Substitute

NFData ExtLamInfo Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: ExtLamInfo -> () #

Generic ExtLamInfo Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep ExtLamInfo 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep ExtLamInfo = D1 ('MetaData "ExtLamInfo" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "ExtLamInfo" 'PrefixI 'True) (S1 ('MetaSel ('Just "extLamModule") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ModuleName) :*: (S1 ('MetaSel ('Just "extLamAbsurd") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool) :*: S1 ('MetaSel ('Just "extLamSys") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (Maybe System)))))
Show ExtLamInfo Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep ExtLamInfo Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep ExtLamInfo = D1 ('MetaData "ExtLamInfo" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "ExtLamInfo" 'PrefixI 'True) (S1 ('MetaSel ('Just "extLamModule") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ModuleName) :*: (S1 ('MetaSel ('Just "extLamAbsurd") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool) :*: S1 ('MetaSel ('Just "extLamSys") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (Maybe System)))))

type Face = [(Term, Bool)] Source #

newtype Fields Source #

Constructors

Fields [(Name, Type)] 

Instances

Instances details
Null Fields Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

data ForeignCode Source #

Constructors

ForeignCode Range String 

Instances

Instances details
EmbPrj ForeignCode Source # 
Instance details

Defined in Agda.TypeChecking.Serialise.Instances.Compilers

NFData ForeignCode Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: ForeignCode -> () #

Generic ForeignCode Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep ForeignCode 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep ForeignCode = D1 ('MetaData "ForeignCode" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "ForeignCode" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String)))
Show ForeignCode Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep ForeignCode Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep ForeignCode = D1 ('MetaData "ForeignCode" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "ForeignCode" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String)))

newtype ForeignCodeStack Source #

Foreign code fragments are stored in reversed order to support efficient appending: head points to the latest pragma in module.

Instances

Instances details
EmbPrj ForeignCodeStack Source # 
Instance details

Defined in Agda.TypeChecking.Serialise.Instances.Compilers

NFData ForeignCodeStack Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: ForeignCodeStack -> () #

Generic ForeignCodeStack Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep ForeignCodeStack 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep ForeignCodeStack = D1 ('MetaData "ForeignCodeStack" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'True) (C1 ('MetaCons "ForeignCodeStack" 'PrefixI 'True) (S1 ('MetaSel ('Just "getForeignCodeStack") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [ForeignCode])))
Show ForeignCodeStack Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep ForeignCodeStack Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep ForeignCodeStack = D1 ('MetaData "ForeignCodeStack" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'True) (C1 ('MetaCons "ForeignCodeStack" 'PrefixI 'True) (S1 ('MetaSel ('Just "getForeignCodeStack") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [ForeignCode])))

class FreshName a where Source #

Create a fresh name from a.

Methods

freshName_ :: MonadFresh NameId m => a -> m Name Source #

Instances

Instances details
FreshName Name Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

FreshName Range Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

FreshName String Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

FreshName () Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

freshName_ :: MonadFresh NameId m => () -> m Name Source #

FreshName (Range, String) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

data Frozen Source #

Frozen meta variable cannot be instantiated by unification. This serves to prevent the completion of a definition by its use outside of the current block. (See issues 118, 288, 399).

Constructors

Frozen

Do not instantiate.

Instantiable 

Instances

Instances details
NFData Frozen Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: Frozen -> () #

Generic Frozen Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep Frozen 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep Frozen = D1 ('MetaData "Frozen" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "Frozen" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "Instantiable" 'PrefixI 'False) (U1 :: Type -> Type))

Methods

from :: Frozen -> Rep Frozen x #

to :: Rep Frozen x -> Frozen #

Show Frozen Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Eq Frozen Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

(==) :: Frozen -> Frozen -> Bool #

(/=) :: Frozen -> Frozen -> Bool #

type Rep Frozen Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep Frozen = D1 ('MetaData "Frozen" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "Frozen" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "Instantiable" 'PrefixI 'False) (U1 :: Type -> Type))

data FunctionData Source #

Constructors

FunctionData 

Fields

Instances

Instances details
Pretty FunctionData Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

NFData FunctionData Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: FunctionData -> () #

Generic FunctionData Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep FunctionData 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep FunctionData = D1 ('MetaData "FunctionData" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "FunctionData" 'PrefixI 'True) (((S1 ('MetaSel ('Just "_funClauses") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Clause]) :*: (S1 ('MetaSel ('Just "_funCompiled") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe CompiledClauses)) :*: S1 ('MetaSel ('Just "_funSplitTree") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe SplitTree)))) :*: ((S1 ('MetaSel ('Just "_funTreeless") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe Compiled)) :*: S1 ('MetaSel ('Just "_funCovering") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Clause])) :*: (S1 ('MetaSel ('Just "_funInv") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 FunctionInverse) :*: S1 ('MetaSel ('Just "_funMutual") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe [QName]))))) :*: ((S1 ('MetaSel ('Just "_funProjection") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Either ProjectionLikenessMissing Projection)) :*: (S1 ('MetaSel ('Just "_funFlags") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (SmallSet FunctionFlag)) :*: S1 ('MetaSel ('Just "_funTerminates") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe Bool)))) :*: ((S1 ('MetaSel ('Just "_funExtLam") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe ExtLamInfo)) :*: S1 ('MetaSel ('Just "_funWith") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe QName))) :*: (S1 ('MetaSel ('Just "_funIsKanOp") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe QName)) :*: S1 ('MetaSel ('Just "_funOpaque") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 IsOpaque))))))
Show FunctionData Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep FunctionData Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep FunctionData = D1 ('MetaData "FunctionData" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "FunctionData" 'PrefixI 'True) (((S1 ('MetaSel ('Just "_funClauses") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Clause]) :*: (S1 ('MetaSel ('Just "_funCompiled") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe CompiledClauses)) :*: S1 ('MetaSel ('Just "_funSplitTree") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe SplitTree)))) :*: ((S1 ('MetaSel ('Just "_funTreeless") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe Compiled)) :*: S1 ('MetaSel ('Just "_funCovering") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Clause])) :*: (S1 ('MetaSel ('Just "_funInv") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 FunctionInverse) :*: S1 ('MetaSel ('Just "_funMutual") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe [QName]))))) :*: ((S1 ('MetaSel ('Just "_funProjection") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Either ProjectionLikenessMissing Projection)) :*: (S1 ('MetaSel ('Just "_funFlags") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (SmallSet FunctionFlag)) :*: S1 ('MetaSel ('Just "_funTerminates") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe Bool)))) :*: ((S1 ('MetaSel ('Just "_funExtLam") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe ExtLamInfo)) :*: S1 ('MetaSel ('Just "_funWith") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe QName))) :*: (S1 ('MetaSel ('Just "_funIsKanOp") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe QName)) :*: S1 ('MetaSel ('Just "_funOpaque") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 IsOpaque))))))

data FunctionFlag Source #

Constructors

FunStatic

Should calls to this function be normalised at compile-time?

FunInline

Should calls to this function be inlined by the compiler?

FunMacro

Is this function a macro?

FunFirstOrder

Is this function INJECTIVE_FOR_INFERENCE? Indicates whether the first-order shortcut should be applied to the definition.

FunErasure

Was --erasure in effect when the function was defined? (This can affect the type of a projection.)

FunAbstract

Is the function abstract?

FunProj

Is this function a descendant of a field (typically, a projection)?

Instances

Instances details
KillRange FunctionFlag Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

EmbPrj FunctionFlag Source # 
Instance details

Defined in Agda.TypeChecking.Serialise.Instances.Internal

SmallSetElement FunctionFlag Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

NFData FunctionFlag Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: FunctionFlag -> () #

Bounded FunctionFlag Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Enum FunctionFlag Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Generic FunctionFlag Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep FunctionFlag 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep FunctionFlag = D1 ('MetaData "FunctionFlag" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) ((C1 ('MetaCons "FunStatic" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "FunInline" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "FunMacro" 'PrefixI 'False) (U1 :: Type -> Type))) :+: ((C1 ('MetaCons "FunFirstOrder" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "FunErasure" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "FunAbstract" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "FunProj" 'PrefixI 'False) (U1 :: Type -> Type))))
Ix FunctionFlag Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Show FunctionFlag Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Eq FunctionFlag Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Ord FunctionFlag Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

KillRange (SmallSet FunctionFlag) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep FunctionFlag Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep FunctionFlag = D1 ('MetaData "FunctionFlag" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) ((C1 ('MetaCons "FunStatic" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "FunInline" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "FunMacro" 'PrefixI 'False) (U1 :: Type -> Type))) :+: ((C1 ('MetaCons "FunFirstOrder" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "FunErasure" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "FunAbstract" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "FunProj" 'PrefixI 'False) (U1 :: Type -> Type))))

data FunctionInverse' c Source #

Constructors

NotInjective 
Inverse (InversionMap c) 

Instances

Instances details
DropArgs FunctionInverse Source # 
Instance details

Defined in Agda.TypeChecking.DropArgs

InstantiateFull FunctionInverse Source # 
Instance details

Defined in Agda.TypeChecking.Reduce

Abstract FunctionInverse Source # 
Instance details

Defined in Agda.TypeChecking.Substitute

Apply FunctionInverse Source # 
Instance details

Defined in Agda.TypeChecking.Substitute

Functor FunctionInverse' Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

fmap :: (a -> b) -> FunctionInverse' a -> FunctionInverse' b #

(<$) :: a -> FunctionInverse' b -> FunctionInverse' a #

Pretty c => Pretty (FunctionInverse' c) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

NamesIn a => NamesIn (FunctionInverse' a) Source # 
Instance details

Defined in Agda.Syntax.Internal.Names

KillRange c => KillRange (FunctionInverse' c) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

EmbPrj a => EmbPrj (FunctionInverse' a) Source # 
Instance details

Defined in Agda.TypeChecking.Serialise.Instances.Internal

NFData c => NFData (FunctionInverse' c) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: FunctionInverse' c -> () #

Generic (FunctionInverse' c) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep (FunctionInverse' c) 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep (FunctionInverse' c) = D1 ('MetaData "FunctionInverse'" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "NotInjective" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "Inverse" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (InversionMap c))))
Show c => Show (FunctionInverse' c) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep (FunctionInverse' c) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep (FunctionInverse' c) = D1 ('MetaData "FunctionInverse'" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "NotInjective" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "Inverse" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (InversionMap c))))

data GHCBackendError Source #

Errors raised by the GHC backend.

Constructors

ConstructorCountMismatch QName [QName] [String]

The number of Haskell constructors (String list) does not match the number of constructors of the given data type.

NotAHaskellType Term WhyNotAHaskellType

GHC backend fails to represent given Agda type in Haskell.

WrongTypeOfMain QName Type

The type of main should be IO _ (QName) but is instead Type.

Instances

Instances details
PrettyTCM GHCBackendError Source # 
Instance details

Defined in Agda.TypeChecking.Errors

NFData GHCBackendError Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: GHCBackendError -> () #

Generic GHCBackendError Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Show GHCBackendError Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep GHCBackendError Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

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

Instances details
NFData GeneralizedValue Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: GeneralizedValue -> () #

Generic GeneralizedValue Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep GeneralizedValue 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep GeneralizedValue = D1 ('MetaData "GeneralizedValue" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "GeneralizedValue" 'PrefixI 'True) (S1 ('MetaSel ('Just "genvalCheckpoint") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 CheckpointId) :*: (S1 ('MetaSel ('Just "genvalTerm") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Just "genvalType") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type))))
Show GeneralizedValue Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep GeneralizedValue Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep GeneralizedValue = D1 ('MetaData "GeneralizedValue" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "GeneralizedValue" 'PrefixI 'True) (S1 ('MetaSel ('Just "genvalCheckpoint") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 CheckpointId) :*: (S1 ('MetaSel ('Just "genvalTerm") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Just "genvalType") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type))))

class Enum i => HasFresh i where Source #

Minimal complete definition

freshLens

newtype IPBoundary' t Source #

Constructors

IPBoundary 

Fields

Instances

Instances details
Functor IPBoundary' Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

fmap :: (a -> b) -> IPBoundary' a -> IPBoundary' b #

(<$) :: a -> IPBoundary' b -> IPBoundary' a #

Foldable IPBoundary' Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

fold :: Monoid m => IPBoundary' m -> m #

foldMap :: Monoid m => (a -> m) -> IPBoundary' a -> m #

foldMap' :: Monoid m => (a -> m) -> IPBoundary' a -> m #

foldr :: (a -> b -> b) -> b -> IPBoundary' a -> b #

foldr' :: (a -> b -> b) -> b -> IPBoundary' a -> b #

foldl :: (b -> a -> b) -> b -> IPBoundary' a -> b #

foldl' :: (b -> a -> b) -> b -> IPBoundary' a -> b #

foldr1 :: (a -> a -> a) -> IPBoundary' a -> a #

foldl1 :: (a -> a -> a) -> IPBoundary' a -> a #

toList :: IPBoundary' a -> [a] #

null :: IPBoundary' a -> Bool #

length :: IPBoundary' a -> Int #

elem :: Eq a => a -> IPBoundary' a -> Bool #

maximum :: Ord a => IPBoundary' a -> a #

minimum :: Ord a => IPBoundary' a -> a #

sum :: Num a => IPBoundary' a -> a #

product :: Num a => IPBoundary' a -> a #

Traversable IPBoundary' Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

traverse :: Applicative f => (a -> f b) -> IPBoundary' a -> f (IPBoundary' b) #

sequenceA :: Applicative f => IPBoundary' (f a) -> f (IPBoundary' a) #

mapM :: Monad m => (a -> m b) -> IPBoundary' a -> m (IPBoundary' b) #

sequence :: Monad m => IPBoundary' (m a) -> m (IPBoundary' a) #

ToConcrete a => ToConcrete (IPBoundary' a) Source # 
Instance details

Defined in Agda.Interaction.BasicOps

Associated Types

type ConOfAbs (IPBoundary' a) 
Instance details

Defined in Agda.Interaction.BasicOps

Methods

toConcrete :: MonadToConcrete m => IPBoundary' a -> m (ConOfAbs (IPBoundary' a)) Source #

bindToConcrete :: MonadToConcrete m => IPBoundary' a -> (ConOfAbs (IPBoundary' a) -> m b) -> m b Source #

Reify a => Reify (IPBoundary' a) Source # 
Instance details

Defined in Agda.Interaction.BasicOps

Associated Types

type ReifiesTo (IPBoundary' a) 
Instance details

Defined in Agda.Interaction.BasicOps

Instantiate t => Instantiate (IPBoundary' t) Source # 
Instance details

Defined in Agda.TypeChecking.Reduce

InstantiateFull t => InstantiateFull (IPBoundary' t) Source # 
Instance details

Defined in Agda.TypeChecking.Reduce

Normalise t => Normalise (IPBoundary' t) Source # 
Instance details

Defined in Agda.TypeChecking.Reduce

Reduce t => Reduce (IPBoundary' t) Source # 
Instance details

Defined in Agda.TypeChecking.Reduce

Simplify t => Simplify (IPBoundary' t) Source # 
Instance details

Defined in Agda.TypeChecking.Reduce

NFData t => NFData (IPBoundary' t) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: IPBoundary' t -> () #

Generic (IPBoundary' t) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep (IPBoundary' t) 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep (IPBoundary' t) = D1 ('MetaData "IPBoundary'" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'True) (C1 ('MetaCons "IPBoundary" 'PrefixI 'True) (S1 ('MetaSel ('Just "getBoundary") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Map (IntMap Bool) t))))

Methods

from :: IPBoundary' t -> Rep (IPBoundary' t) x #

to :: Rep (IPBoundary' t) x -> IPBoundary' t #

Show t => Show (IPBoundary' t) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type ConOfAbs (IPBoundary' a) Source # 
Instance details

Defined in Agda.Interaction.BasicOps

type ReifiesTo (IPBoundary' a) Source # 
Instance details

Defined in Agda.Interaction.BasicOps

type Rep (IPBoundary' t) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep (IPBoundary' t) = D1 ('MetaData "IPBoundary'" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'True) (C1 ('MetaCons "IPBoundary" 'PrefixI 'True) (S1 ('MetaSel ('Just "getBoundary") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Map (IntMap Bool) t))))

data IPClause Source #

Which clause is an interaction point located in?

Constructors

IPClause 

Fields

IPNoClause

The interaction point is not in the rhs of a clause.

Instances

Instances details
NFData IPClause Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: IPClause -> () #

Generic IPClause Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep IPClause 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

from :: IPClause -> Rep IPClause x #

to :: Rep IPClause x -> IPClause #

Eq IPClause Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep IPClause Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

data IllegalRewriteRuleReason Source #

Instances

Instances details
EmbPrj IllegalRewriteRuleReason Source # 
Instance details

Defined in Agda.TypeChecking.Serialise.Instances.Errors

NFData IllegalRewriteRuleReason Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Generic IllegalRewriteRuleReason Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep IllegalRewriteRuleReason 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep IllegalRewriteRuleReason = D1 ('MetaData "IllegalRewriteRuleReason" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (((C1 ('MetaCons "LHSNotDefinitionOrConstructor" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "VariablesNotBoundByLHS" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 IntSet)) :+: C1 ('MetaCons "VariablesBoundMoreThanOnce" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 IntSet)))) :+: ((C1 ('MetaCons "LHSReduces" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)) :+: C1 ('MetaCons "HeadSymbolIsProjectionLikeFunction" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName))) :+: (C1 ('MetaCons "HeadSymbolIsTypeConstructor" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "HeadSymbolContainsMetas" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName))))) :+: (((C1 ('MetaCons "ConstructorParametersNotGeneral" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ConHead) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Args)) :+: C1 ('MetaCons "ContainsUnsolvedMetaVariables" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Set1 MetaId)))) :+: (C1 ('MetaCons "BlockedOnProblems" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Set1 ProblemId))) :+: C1 ('MetaCons "RequiresDefinitions" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Set1 QName))))) :+: ((C1 ('MetaCons "DoesNotTargetRewriteRelation" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BeforeFunctionDefinition" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "BeforeMutualFunctionDefinition" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "DuplicateRewriteRule" 'PrefixI 'False) (U1 :: Type -> Type)))))
Show IllegalRewriteRuleReason Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep IllegalRewriteRuleReason Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep IllegalRewriteRuleReason = D1 ('MetaData "IllegalRewriteRuleReason" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (((C1 ('MetaCons "LHSNotDefinitionOrConstructor" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "VariablesNotBoundByLHS" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 IntSet)) :+: C1 ('MetaCons "VariablesBoundMoreThanOnce" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 IntSet)))) :+: ((C1 ('MetaCons "LHSReduces" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)) :+: C1 ('MetaCons "HeadSymbolIsProjectionLikeFunction" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName))) :+: (C1 ('MetaCons "HeadSymbolIsTypeConstructor" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "HeadSymbolContainsMetas" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName))))) :+: (((C1 ('MetaCons "ConstructorParametersNotGeneral" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ConHead) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Args)) :+: C1 ('MetaCons "ContainsUnsolvedMetaVariables" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Set1 MetaId)))) :+: (C1 ('MetaCons "BlockedOnProblems" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Set1 ProblemId))) :+: C1 ('MetaCons "RequiresDefinitions" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Set1 QName))))) :+: ((C1 ('MetaCons "DoesNotTargetRewriteRelation" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "BeforeFunctionDefinition" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "BeforeMutualFunctionDefinition" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "DuplicateRewriteRule" 'PrefixI 'False) (U1 :: Type -> Type)))))

data IncorrectTypeForRewriteRelationReason Source #

Instances

Instances details
NFData IncorrectTypeForRewriteRelationReason Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Generic IncorrectTypeForRewriteRelationReason Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep IncorrectTypeForRewriteRelationReason 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep IncorrectTypeForRewriteRelationReason = D1 ('MetaData "IncorrectTypeForRewriteRelationReason" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "ShouldAcceptAtLeastTwoArguments" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "FinalTwoArgumentsNotVisible" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "TypeDoesNotEndInSort" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Telescope))))
Show IncorrectTypeForRewriteRelationReason Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep IncorrectTypeForRewriteRelationReason Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep IncorrectTypeForRewriteRelationReason = D1 ('MetaData "IncorrectTypeForRewriteRelationReason" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "ShouldAcceptAtLeastTwoArguments" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "FinalTwoArgumentsNotVisible" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "TypeDoesNotEndInSort" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Telescope))))

data InductionAndEta Source #

Instances

Instances details
NFData InductionAndEta Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: InductionAndEta -> () #

Generic InductionAndEta Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep InductionAndEta 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep InductionAndEta = D1 ('MetaData "InductionAndEta" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "InductionAndEta" 'PrefixI 'True) (S1 ('MetaSel ('Just "recordInduction") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe Induction)) :*: S1 ('MetaSel ('Just "recordEtaEquality") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 EtaEquality)))
Show InductionAndEta Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep InductionAndEta Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep InductionAndEta = D1 ('MetaData "InductionAndEta" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "InductionAndEta" 'PrefixI 'True) (S1 ('MetaSel ('Just "recordInduction") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe Induction)) :*: S1 ('MetaSel ('Just "recordEtaEquality") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 EtaEquality)))

data InstanceInfo Source #

Information about an instance definition.

Constructors

InstanceInfo 

Fields

Instances

Instances details
KillRange InstanceInfo Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

EmbPrj InstanceInfo Source # 
Instance details

Defined in Agda.TypeChecking.Serialise.Instances.Internal

NFData InstanceInfo Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: InstanceInfo -> () #

Generic InstanceInfo Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep InstanceInfo 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep InstanceInfo = D1 ('MetaData "InstanceInfo" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "InstanceInfo" 'PrefixI 'True) (S1 ('MetaSel ('Just "instanceClass") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Just "instanceOverlap") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 OverlapMode)))
Show InstanceInfo Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep InstanceInfo Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep InstanceInfo = D1 ('MetaData "InstanceInfo" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "InstanceInfo" 'PrefixI 'True) (S1 ('MetaSel ('Just "instanceClass") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Just "instanceOverlap") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 OverlapMode)))

data InstanceTable Source #

Records information about the instances in the signature. Does not deal with local instances.

Constructors

InstanceTable 

Fields

  • _itableTree :: DiscrimTree QName

    The actual discrimination tree for looking up instances with

  • _itableCounts :: Map QName Int

    For profiling, we store the number of instances on a per-class basis. This lets us compare the result from the discrimination tree with all the instances in scope, thus informing us how many validity checks were skipped.

Instances

Instances details
KillRange InstanceTable Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

EmbPrj InstanceTable Source # 
Instance details

Defined in Agda.TypeChecking.Serialise.Instances.Internal

NFData InstanceTable Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: InstanceTable -> () #

Monoid InstanceTable Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Semigroup InstanceTable Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Generic InstanceTable Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep InstanceTable 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep InstanceTable = D1 ('MetaData "InstanceTable" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "InstanceTable" 'PrefixI 'True) (S1 ('MetaSel ('Just "_itableTree") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (DiscrimTree QName)) :*: S1 ('MetaSel ('Just "_itableCounts") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Map QName Int))))
Show InstanceTable Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep InstanceTable Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep InstanceTable = D1 ('MetaData "InstanceTable" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "InstanceTable" 'PrefixI 'True) (S1 ('MetaSel ('Just "_itableTree") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (DiscrimTree QName)) :*: S1 ('MetaSel ('Just "_itableCounts") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Map QName Int))))

data Instantiation Source #

Meta-variable instantiations.

Constructors

Instantiation 

Fields

Instances

Instances details
InstantiateFull Instantiation Source # 
Instance details

Defined in Agda.TypeChecking.Reduce

EmbPrj Instantiation Source # 
Instance details

Defined in Agda.TypeChecking.Serialise.Instances.Internal

NFData Instantiation Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: Instantiation -> () #

Generic Instantiation Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep Instantiation 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep Instantiation = D1 ('MetaData "Instantiation" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "Instantiation" 'PrefixI 'True) (S1 ('MetaSel ('Just "instTel") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Arg String]) :*: S1 ('MetaSel ('Just "instBody") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)))
Show Instantiation Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep Instantiation Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep Instantiation = D1 ('MetaData "Instantiation" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "Instantiation" 'PrefixI 'True) (S1 ('MetaSel ('Just "instTel") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Arg String]) :*: S1 ('MetaSel ('Just "instBody") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)))

data InteractionError Source #

Errors raised in --interaction mode.

Constructors

CannotGive Expr

Failure of the give interactive tactic.

CannotRefine String

Failure of the refine interactive tactic.

CaseSplitError Doc

Failure of the makeCase interactive tactic.

ExpectedIdentifier Expr

Expected the given expression to be an identifier.

ExpectedApplication

Expected an argument of the form f e1 e2 .. en.

NoActionForInteractionPoint InteractionId

Interaction point has not been reached during type checking.

NoSuchInteractionPoint InteractionId

InteractionId does not resolve to an InteractionPoint.

UnexpectedWhere

where not allowed in hole.

Instances

Instances details
PrettyTCM InteractionError Source # 
Instance details

Defined in Agda.TypeChecking.Errors

NFData InteractionError Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: InteractionError -> () #

Generic InteractionError Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep InteractionError 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep InteractionError = D1 ('MetaData "InteractionError" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (((C1 ('MetaCons "CannotGive" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr)) :+: C1 ('MetaCons "CannotRefine" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String))) :+: (C1 ('MetaCons "CaseSplitError" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Doc)) :+: C1 ('MetaCons "ExpectedIdentifier" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr)))) :+: ((C1 ('MetaCons "ExpectedApplication" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "NoActionForInteractionPoint" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 InteractionId))) :+: (C1 ('MetaCons "NoSuchInteractionPoint" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 InteractionId)) :+: C1 ('MetaCons "UnexpectedWhere" 'PrefixI 'False) (U1 :: Type -> Type))))
Show InteractionError Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep InteractionError Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep InteractionError = D1 ('MetaData "InteractionError" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (((C1 ('MetaCons "CannotGive" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr)) :+: C1 ('MetaCons "CannotRefine" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String))) :+: (C1 ('MetaCons "CaseSplitError" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Doc)) :+: C1 ('MetaCons "ExpectedIdentifier" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr)))) :+: ((C1 ('MetaCons "ExpectedApplication" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "NoActionForInteractionPoint" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 InteractionId))) :+: (C1 ('MetaCons "NoSuchInteractionPoint" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 InteractionId)) :+: C1 ('MetaCons "UnexpectedWhere" 'PrefixI 'False) (U1 :: Type -> Type))))

type InteractionOutputCallback = Response_boot TCErr TCWarning WarningsAndNonFatalErrors -> TCM () Source #

Callback fuction to call when there is a response to give to the interactive frontend.

Note that the response is given in pieces and incrementally, so the user can have timely response even during long computations.

Typical InteractionOutputCallback functions:

  • Convert the response into a String representation and print it on standard output (suitable for inter-process communication).
  • Put the response into a mutable variable stored in the closure of the InteractionOutputCallback function. (suitable for intra-process communication).

data InteractionPoint Source #

Interaction points are created by the scope checker who sets the range. The meta variable is created by the type checker and then hooked up to the interaction point.

Constructors

InteractionPoint 

Fields

Instances

Instances details
HasTag InteractionPoint Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Tag InteractionPoint 
Instance details

Defined in Agda.TypeChecking.Monad.Base

NFData InteractionPoint Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: InteractionPoint -> () #

NFData InteractionPoints Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: InteractionPoints -> () #

Generic InteractionPoint Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep InteractionPoint 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Eq InteractionPoint Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Tag InteractionPoint Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep InteractionPoint Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

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)

data Interface Source #

Constructors

Interface 

Fields

Instances

Instances details
Pretty Interface Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

InstantiateFull Interface Source # 
Instance details

Defined in Agda.TypeChecking.Reduce

EmbPrj Interface Source # 
Instance details

Defined in Agda.TypeChecking.Serialise.Instances

NFData Interface Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: Interface -> () #

Generic Interface Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep Interface 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep Interface = D1 ('MetaData "Interface" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "Interface" 'PrefixI 'True) ((((S1 ('MetaSel ('Just "iSourceHash") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Hash) :*: (S1 ('MetaSel ('Just "iSource") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Text) :*: S1 ('MetaSel ('Just "iFileType") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 FileType))) :*: (S1 ('MetaSel ('Just "iImportedModules") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [(TopLevelModuleName, Hash)]) :*: (S1 ('MetaSel ('Just "iModuleName") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ModuleName) :*: S1 ('MetaSel ('Just "iTopLevelModuleName") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 TopLevelModuleName)))) :*: ((S1 ('MetaSel ('Just "iScope") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Map ModuleName Scope)) :*: (S1 ('MetaSel ('Just "iInsideScope") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ScopeInfo) :*: S1 ('MetaSel ('Just "iSignature") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Signature))) :*: (S1 ('MetaSel ('Just "iMetaBindings") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 RemoteMetaStore) :*: (S1 ('MetaSel ('Just "iDisplayForms") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 DisplayForms) :*: S1 ('MetaSel ('Just "iUserWarnings") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 UserWarnings))))) :*: (((S1 ('MetaSel ('Just "iImportWarning") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe Text)) :*: (S1 ('MetaSel ('Just "iBuiltin") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (BuiltinThings' (PrimitiveId, QName))) :*: S1 ('MetaSel ('Just "iForeignCode") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Map BackendName ForeignCodeStack)))) :*: (S1 ('MetaSel ('Just "iHighlighting") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 HighlightingInfo) :*: (S1 ('MetaSel ('Just "iDefaultPragmaOptions") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [OptionsPragma]) :*: S1 ('MetaSel ('Just "iFilePragmaOptions") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [OptionsPragma])))) :*: ((S1 ('MetaSel ('Just "iOptionsUsed") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 PragmaOptions) :*: (S1 ('MetaSel ('Just "iPatternSyns") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 PatternSynDefns) :*: S1 ('MetaSel ('Just "iWarnings") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Set TCWarning)))) :*: (S1 ('MetaSel ('Just "iPartialDefs") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Set QName)) :*: (S1 ('MetaSel ('Just "iOpaqueBlocks") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Map OpaqueId OpaqueBlock)) :*: S1 ('MetaSel ('Just "iOpaqueNames") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Map QName OpaqueId))))))))
Show Interface Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep Interface Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep Interface = D1 ('MetaData "Interface" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "Interface" 'PrefixI 'True) ((((S1 ('MetaSel ('Just "iSourceHash") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Hash) :*: (S1 ('MetaSel ('Just "iSource") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Text) :*: S1 ('MetaSel ('Just "iFileType") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 FileType))) :*: (S1 ('MetaSel ('Just "iImportedModules") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [(TopLevelModuleName, Hash)]) :*: (S1 ('MetaSel ('Just "iModuleName") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ModuleName) :*: S1 ('MetaSel ('Just "iTopLevelModuleName") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 TopLevelModuleName)))) :*: ((S1 ('MetaSel ('Just "iScope") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Map ModuleName Scope)) :*: (S1 ('MetaSel ('Just "iInsideScope") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ScopeInfo) :*: S1 ('MetaSel ('Just "iSignature") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Signature))) :*: (S1 ('MetaSel ('Just "iMetaBindings") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 RemoteMetaStore) :*: (S1 ('MetaSel ('Just "iDisplayForms") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 DisplayForms) :*: S1 ('MetaSel ('Just "iUserWarnings") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 UserWarnings))))) :*: (((S1 ('MetaSel ('Just "iImportWarning") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe Text)) :*: (S1 ('MetaSel ('Just "iBuiltin") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (BuiltinThings' (PrimitiveId, QName))) :*: S1 ('MetaSel ('Just "iForeignCode") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Map BackendName ForeignCodeStack)))) :*: (S1 ('MetaSel ('Just "iHighlighting") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 HighlightingInfo) :*: (S1 ('MetaSel ('Just "iDefaultPragmaOptions") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [OptionsPragma]) :*: S1 ('MetaSel ('Just "iFilePragmaOptions") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [OptionsPragma])))) :*: ((S1 ('MetaSel ('Just "iOptionsUsed") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 PragmaOptions) :*: (S1 ('MetaSel ('Just "iPatternSyns") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 PatternSynDefns) :*: S1 ('MetaSel ('Just "iWarnings") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Set TCWarning)))) :*: (S1 ('MetaSel ('Just "iPartialDefs") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Set QName)) :*: (S1 ('MetaSel ('Just "iOpaqueBlocks") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Map OpaqueId OpaqueBlock)) :*: S1 ('MetaSel ('Just "iOpaqueNames") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Map QName OpaqueId))))))))

data InvalidFileNameReason Source #

Extra information for InvalidFileName error.

Instances

Instances details
NFData InvalidFileNameReason Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: InvalidFileNameReason -> () #

Generic InvalidFileNameReason Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep InvalidFileNameReason 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep InvalidFileNameReason = D1 ('MetaData "InvalidFileNameReason" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "DoesNotCorrespondToValidModuleName" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "RootNameModuleNotAQualifiedModuleName" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Text)))
Show InvalidFileNameReason Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep InvalidFileNameReason Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep InvalidFileNameReason = D1 ('MetaData "InvalidFileNameReason" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "DoesNotCorrespondToValidModuleName" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "RootNameModuleNotAQualifiedModuleName" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Text)))

data IsAmbiguous Source #

Boolean flag whether a name is ambiguous.

Instances

Instances details
EmbPrj IsAmbiguous Source # 
Instance details

Defined in Agda.TypeChecking.Serialise.Instances.Errors

NFData IsAmbiguous Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: IsAmbiguous -> () #

Generic IsAmbiguous Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep IsAmbiguous 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep IsAmbiguous = D1 ('MetaData "IsAmbiguous" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "YesAmbiguous" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 AmbiguousQName)) :+: C1 ('MetaCons "NotAmbiguous" 'PrefixI 'False) (U1 :: Type -> Type))
Show IsAmbiguous Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep IsAmbiguous Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep IsAmbiguous = D1 ('MetaData "IsAmbiguous" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "YesAmbiguous" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 AmbiguousQName)) :+: C1 ('MetaCons "NotAmbiguous" 'PrefixI 'False) (U1 :: Type -> Type))

data IsForced Source #

Information about whether an argument is forced by the type of a function.

Constructors

Forced 
NotForced 

Instances

Instances details
KillRange IsForced Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

PrettyTCM IsForced Source # 
Instance details

Defined in Agda.TypeChecking.Pretty

ChooseFlex IsForced Source # 
Instance details

Defined in Agda.TypeChecking.Rules.LHS.Problem

EmbPrj IsForced Source # 
Instance details

Defined in Agda.TypeChecking.Serialise.Instances.Internal

NFData IsForced Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: IsForced -> () #

Generic IsForced Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep IsForced 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep IsForced = D1 ('MetaData "IsForced" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "Forced" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "NotForced" 'PrefixI 'False) (U1 :: Type -> Type))

Methods

from :: IsForced -> Rep IsForced x #

to :: Rep IsForced x -> IsForced #

Show IsForced Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Eq IsForced Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep IsForced Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep IsForced = D1 ('MetaData "IsForced" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "Forced" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "NotForced" 'PrefixI 'False) (U1 :: Type -> Type))

data IsReduced Source #

Three cases: 1. not reduced, 2. reduced, but blocked, 3. reduced, not blocked.

Constructors

NotReduced 
Reduced (Blocked ()) 

data JSBackendError Source #

Errors raised by the JS backend.

Constructors

BadCompilePragma 

Instances

Instances details
PrettyTCM JSBackendError Source # 
Instance details

Defined in Agda.TypeChecking.Errors

NFData JSBackendError Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: JSBackendError -> () #

Generic JSBackendError Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep JSBackendError 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep JSBackendError = D1 ('MetaData "JSBackendError" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "BadCompilePragma" 'PrefixI 'False) (U1 :: Type -> Type))
Show JSBackendError Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep JSBackendError Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep JSBackendError = D1 ('MetaData "JSBackendError" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "BadCompilePragma" 'PrefixI 'False) (U1 :: Type -> Type))

data Judgement a Source #

Parametrized since it is used without MetaId when creating a new meta.

Constructors

HasType 

Fields

IsSort 

Fields

Instances

Instances details
Pretty a => Pretty (Judgement a) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

PrettyTCM a => PrettyTCM (Judgement a) Source # 
Instance details

Defined in Agda.TypeChecking.Pretty

Methods

prettyTCM :: MonadPretty m => Judgement a -> m Doc Source #

InstantiateFull (Judgement MetaId) Source # 
Instance details

Defined in Agda.TypeChecking.Reduce

EmbPrj a => EmbPrj (Judgement a) Source # 
Instance details

Defined in Agda.TypeChecking.Serialise.Instances.Internal

NFData a => NFData (Judgement a) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: Judgement a -> () #

Generic (Judgement a) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep (Judgement a) 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep (Judgement a) = D1 ('MetaData "Judgement" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "HasType" 'PrefixI 'True) (S1 ('MetaSel ('Just "jMetaId") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 a) :*: (S1 ('MetaSel ('Just "jComparison") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: S1 ('MetaSel ('Just "jMetaType") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type))) :+: C1 ('MetaCons "IsSort" 'PrefixI 'True) (S1 ('MetaSel ('Just "jMetaId") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 a) :*: S1 ('MetaSel ('Just "jMetaType") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)))

Methods

from :: Judgement a -> Rep (Judgement a) x #

to :: Rep (Judgement a) x -> Judgement a #

Show a => Show (Judgement a) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep (Judgement a) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep (Judgement a) = D1 ('MetaData "Judgement" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "HasType" 'PrefixI 'True) (S1 ('MetaSel ('Just "jMetaId") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 a) :*: (S1 ('MetaSel ('Just "jComparison") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: S1 ('MetaSel ('Just "jMetaType") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type))) :+: C1 ('MetaCons "IsSort" 'PrefixI 'True) (S1 ('MetaSel ('Just "jMetaId") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 a) :*: S1 ('MetaSel ('Just "jMetaType") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)))

data LHSOrPatSyn Source #

Distinguish error message when parsing lhs or pattern synonym, resp.

Constructors

IsLHS 
IsPatSyn 

Instances

Instances details
EmbPrj LHSOrPatSyn Source # 
Instance details

Defined in Agda.TypeChecking.Serialise.Instances.Errors

NFData LHSOrPatSyn Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: LHSOrPatSyn -> () #

Bounded LHSOrPatSyn Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Enum LHSOrPatSyn Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Generic LHSOrPatSyn Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep LHSOrPatSyn 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep LHSOrPatSyn = D1 ('MetaData "LHSOrPatSyn" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "IsLHS" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "IsPatSyn" 'PrefixI 'False) (U1 :: Type -> Type))
Show LHSOrPatSyn Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Eq LHSOrPatSyn Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep LHSOrPatSyn Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep LHSOrPatSyn = D1 ('MetaData "LHSOrPatSyn" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "IsLHS" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "IsPatSyn" 'PrefixI 'False) (U1 :: Type -> Type))

class LensClosure b a | b -> a where Source #

Methods

lensClosure :: Lens' b (Closure a) Source #

class LensTCEnv a where Source #

Instances

Instances details
LensTCEnv TCEnv Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

LensTCEnv (Closure a) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

data LetBinding Source #

Constructors

LetBinding 

Instances

Instances details
InstantiateFull LetBinding Source # 
Instance details

Defined in Agda.TypeChecking.Reduce

Subst LetBinding Source # 
Instance details

Defined in Agda.TypeChecking.Substitute

Associated Types

type SubstArg LetBinding 
Instance details

Defined in Agda.TypeChecking.Substitute

NFData LetBinding Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: LetBinding -> () #

Generic LetBinding Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep LetBinding 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep LetBinding = D1 ('MetaData "LetBinding" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "LetBinding" 'PrefixI 'True) (S1 ('MetaSel ('Just "letOrigin") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Origin) :*: (S1 ('MetaSel ('Just "letTerm") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Just "letType") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Dom Type)))))
Show LetBinding Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type SubstArg LetBinding Source # 
Instance details

Defined in Agda.TypeChecking.Substitute

type Rep LetBinding Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep LetBinding = D1 ('MetaData "LetBinding" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "LetBinding" 'PrefixI 'True) (S1 ('MetaSel ('Just "letOrigin") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Origin) :*: (S1 ('MetaSel ('Just "letTerm") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Just "letType") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Dom Type)))))

data Listener Source #

Instances

Instances details
NFData Listener Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: Listener -> () #

Generic Listener Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep Listener 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

from :: Listener -> Rep Listener x #

to :: Rep Listener x -> Listener #

Eq Listener Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Ord Listener Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep Listener Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

data LoadedFileCache Source #

Instances

Instances details
NFData LoadedFileCache Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: LoadedFileCache -> () #

Generic LoadedFileCache Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep LoadedFileCache 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep LoadedFileCache = D1 ('MetaData "LoadedFileCache" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "LoadedFileCache" 'PrefixI 'True) (S1 ('MetaSel ('Just "lfcCached") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 CachedTypeCheckLog) :*: S1 ('MetaSel ('Just "lfcCurrent") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 CurrentTypeCheckLog)))
type Rep LoadedFileCache Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep LoadedFileCache = D1 ('MetaData "LoadedFileCache" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "LoadedFileCache" 'PrefixI 'True) (S1 ('MetaSel ('Just "lfcCached") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 CachedTypeCheckLog) :*: S1 ('MetaSel ('Just "lfcCurrent") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 CurrentTypeCheckLog)))

type LocalMetaStore = Map MetaId MetaVariable Source #

Used for meta-variables from the current module.

data MaybeReduced a Source #

Constructors

MaybeRed 

Instances

Instances details
Functor MaybeReduced Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

fmap :: (a -> b) -> MaybeReduced a -> MaybeReduced b #

(<$) :: a -> MaybeReduced b -> MaybeReduced a #

IsProjElim e => IsProjElim (MaybeReduced e) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

PrettyTCM a => PrettyTCM (MaybeReduced a) Source # 
Instance details

Defined in Agda.TypeChecking.Pretty

data MetaInfo Source #

MetaInfo is cloned from one meta to the next during pruning.

Constructors

MetaInfo 

Fields

Instances

Instances details
LensIsAbstract MetaInfo Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

LensModalPolarity MetaInfo Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

LensModality MetaInfo Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

LensQuantity MetaInfo Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

LensRelevance MetaInfo Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

HasRange MetaInfo Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

SetRange MetaInfo Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

NFData MetaInfo Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: MetaInfo -> () #

Generic MetaInfo Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep MetaInfo 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

from :: MetaInfo -> Rep MetaInfo x #

to :: Rep MetaInfo x -> MetaInfo #

LensClosure MetaInfo Range Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep MetaInfo Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

data MetaInstantiation Source #

Solution status of meta.

Constructors

InstV Instantiation

Solved by Instantiation.

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

Instances details
Pretty MetaInstantiation Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

NFData MetaInstantiation Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: MetaInstantiation -> () #

Generic MetaInstantiation Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep MetaInstantiation Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

newtype MetaPriority Source #

Meta variable priority: When we have an equation between meta-variables, which one should be instantiated?

Higher value means higher priority to be instantiated.

Constructors

MetaPriority Int 

data MetaVariable Source #

Information about local meta-variables.

Constructors

MetaVar 

Fields

Instances

Instances details
LensModality MetaVariable Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

LensQuantity MetaVariable Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

LensRelevance MetaVariable Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

HasRange MetaVariable Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

SetRange MetaVariable Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

NFData MetaVariable Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: MetaVariable -> () #

Generic MetaVariable Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

LensClosure MetaVariable Range Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep MetaVariable Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

data MissingTypeSignatureInfo Source #

Extra information for MissingTypeSignature error.

Constructors

MissingDataSignature Name

The data definition for Name lacks a data signature.

MissingRecordSignature Name

The record definition for Name lacks a record signature.

MissingFunctionSignature LHS

The function lhs misses a type signature.

Instances

Instances details
PrettyTCM MissingTypeSignatureInfo Source # 
Instance details

Defined in Agda.TypeChecking.Errors

NFData MissingTypeSignatureInfo Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Generic MissingTypeSignatureInfo Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep MissingTypeSignatureInfo 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep MissingTypeSignatureInfo = D1 ('MetaData "MissingTypeSignatureInfo" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "MissingDataSignature" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Name)) :+: (C1 ('MetaCons "MissingRecordSignature" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Name)) :+: C1 ('MetaCons "MissingFunctionSignature" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 LHS))))
Show MissingTypeSignatureInfo Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep MissingTypeSignatureInfo Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep MissingTypeSignatureInfo = D1 ('MetaData "MissingTypeSignatureInfo" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "MissingDataSignature" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Name)) :+: (C1 ('MetaCons "MissingRecordSignature" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Name)) :+: C1 ('MetaCons "MissingFunctionSignature" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 LHS))))

data ModuleCheckMode Source #

Distinguishes between type-checked and scope-checked interfaces when stored in the map of VisitedModules.

Instances

Instances details
NFData ModuleCheckMode Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: ModuleCheckMode -> () #

Bounded ModuleCheckMode Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Enum ModuleCheckMode Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Generic ModuleCheckMode Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep ModuleCheckMode 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep ModuleCheckMode = D1 ('MetaData "ModuleCheckMode" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "ModuleScopeChecked" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "ModuleTypeChecked" 'PrefixI 'False) (U1 :: Type -> Type))
Show ModuleCheckMode Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Eq ModuleCheckMode Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Ord ModuleCheckMode Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep ModuleCheckMode Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep ModuleCheckMode = D1 ('MetaData "ModuleCheckMode" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "ModuleScopeChecked" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "ModuleTypeChecked" 'PrefixI 'False) (U1 :: Type -> Type))

data ModuleInfo Source #

Constructors

ModuleInfo 

Fields

Instances

Instances details
NFData ModuleInfo Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: ModuleInfo -> () #

Generic ModuleInfo Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep ModuleInfo 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep ModuleInfo = D1 ('MetaData "ModuleInfo" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "ModuleInfo" 'PrefixI 'True) ((S1 ('MetaSel ('Just "miInterface") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Interface) :*: S1 ('MetaSel ('Just "miWarnings") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Set TCWarning))) :*: (S1 ('MetaSel ('Just "miPrimitive") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool) :*: S1 ('MetaSel ('Just "miMode") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ModuleCheckMode))))
type Rep ModuleInfo Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep ModuleInfo = D1 ('MetaData "ModuleInfo" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "ModuleInfo" 'PrefixI 'True) ((S1 ('MetaSel ('Just "miInterface") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Interface) :*: S1 ('MetaSel ('Just "miWarnings") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Set TCWarning))) :*: (S1 ('MetaSel ('Just "miPrimitive") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool) :*: S1 ('MetaSel ('Just "miMode") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ModuleCheckMode))))

class Monad m => MonadBlock (m :: Type -> Type) where Source #

Minimal complete definition

catchPatternErr

Methods

patternViolation :: Blocker -> m a Source #

`patternViolation b` aborts the current computation

default patternViolation :: forall (t :: (Type -> Type) -> Type -> Type) (n :: Type -> Type) a. (MonadTrans t, MonadBlock n, m ~ t n) => Blocker -> m a Source #

catchPatternErr :: (Blocker -> m a) -> m a -> m a Source #

`catchPatternErr handle m` runs m, handling pattern violations with handle (doesn't roll back the state)

Instances

Instances details
MonadBlock TCM Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

MonadBlock NLM Source # 
Instance details

Defined in Agda.TypeChecking.Rewriting.NonLinMatch

Monad m => MonadBlock (PureConversionT m) Source # 
Instance details

Defined in Agda.TypeChecking.Conversion.Pure

Monad m => MonadBlock (BlockT m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

MonadBlock m => MonadBlock (MaybeT m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Monad m => MonadBlock (ExceptT TCErr m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

MonadBlock m => MonadBlock (ReaderT e m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

patternViolation :: Blocker -> ReaderT e m a Source #

catchPatternErr :: (Blocker -> ReaderT e m a) -> ReaderT e m a -> ReaderT e m a Source #

class Monad m => MonadFresh i (m :: Type -> Type) where Source #

Minimal complete definition

Nothing

Methods

fresh :: m i Source #

default fresh :: forall (t :: (Type -> Type) -> Type -> Type) (n :: Type -> Type). (MonadTrans t, MonadFresh i n, t n ~ m) => m i Source #

Instances

Instances details
HasFresh i => MonadFresh i TCM Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

fresh :: TCM i Source #

Monad m => MonadFresh NameId (PureConversionT m) Source # 
Instance details

Defined in Agda.TypeChecking.Conversion.Pure

Monad m => MonadFresh ProblemId (PureConversionT m) Source # 
Instance details

Defined in Agda.TypeChecking.Conversion.Pure

Monad m => MonadFresh Int (PureConversionT m) Source # 
Instance details

Defined in Agda.TypeChecking.Conversion.Pure

MonadFresh i m => MonadFresh i (ListT m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

fresh :: ListT m i Source #

MonadFresh i m => MonadFresh i (MaybeT m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

fresh :: MaybeT m i Source #

MonadFresh i m => MonadFresh i (ExceptT e m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

fresh :: ExceptT e m i Source #

MonadFresh i m => MonadFresh i (IdentityT m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

fresh :: IdentityT m i Source #

MonadFresh i m => MonadFresh i (ReaderT r m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

fresh :: ReaderT r m i Source #

MonadFresh i m => MonadFresh i (StateT s m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

fresh :: StateT s m i Source #

(MonadFresh i m, Monoid w) => MonadFresh i (WriterT w m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

fresh :: WriterT w m i Source #

class (Applicative m, MonadTCEnv m, ReadTCState m, HasOptions m) => MonadReduce (m :: Type -> Type) where Source #

Minimal complete definition

Nothing

Methods

liftReduce :: ReduceM a -> m a Source #

default liftReduce :: forall (t :: (Type -> Type) -> Type -> Type) (n :: Type -> Type) a. (MonadTrans t, MonadReduce n, t n ~ m) => ReduceM a -> m a Source #

Instances

Instances details
MonadReduce TerM Source # 
Instance details

Defined in Agda.Termination.Monad

Methods

liftReduce :: ReduceM a -> TerM a Source #

MonadReduce ReduceM Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

liftReduce :: ReduceM a -> ReduceM a Source #

MonadReduce NLM Source # 
Instance details

Defined in Agda.TypeChecking.Rewriting.NonLinMatch

Methods

liftReduce :: ReduceM a -> NLM a Source #

MonadReduce m => MonadReduce (PureConversionT m) Source # 
Instance details

Defined in Agda.TypeChecking.Conversion.Pure

MonadReduce m => MonadReduce (BlockT m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

liftReduce :: ReduceM a -> BlockT m a Source #

MonadIO m => MonadReduce (TCMT m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

liftReduce :: ReduceM a -> TCMT m a Source #

MonadReduce m => MonadReduce (NamesT m) Source # 
Instance details

Defined in Agda.TypeChecking.Names

Methods

liftReduce :: ReduceM a -> NamesT m a Source #

MonadReduce m => MonadReduce (ListT m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

liftReduce :: ReduceM a -> ListT m a Source #

MonadReduce m => MonadReduce (ChangeT m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

liftReduce :: ReduceM a -> ChangeT m a Source #

MonadReduce m => MonadReduce (MaybeT m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

liftReduce :: ReduceM a -> MaybeT m a Source #

MonadReduce m => MonadReduce (ExceptT err m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

liftReduce :: ReduceM a -> ExceptT err m a Source #

MonadReduce m => MonadReduce (IdentityT m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

liftReduce :: ReduceM a -> IdentityT m a Source #

MonadReduce m => MonadReduce (ReaderT r m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

liftReduce :: ReduceM a -> ReaderT r m a Source #

MonadReduce m => MonadReduce (StateT w m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

liftReduce :: ReduceM a -> StateT w m a Source #

(Monoid w, MonadReduce m) => MonadReduce (WriterT w m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

liftReduce :: ReduceM a -> WriterT w m a Source #

class Monad m => MonadStConcreteNames (m :: Type -> Type) where Source #

A monad that has read and write access to the stConcreteNames part of the TCState. Basically, this is a synonym for `MonadState ConcreteNames m` (which cannot be used directly because of the limitations of Haskell's typeclass system).

Minimal complete definition

runStConcreteNames

Instances

Instances details
MonadStConcreteNames TCM Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

ReadTCState m => MonadStConcreteNames (PureConversionT m) Source # 
Instance details

Defined in Agda.TypeChecking.Conversion.Pure

MonadStConcreteNames m => MonadStConcreteNames (MaybeT m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

MonadStConcreteNames m => MonadStConcreteNames (ExceptT e m) Source #

The concrete names get lost in case of an exception.

Instance details

Defined in Agda.TypeChecking.Monad.Base

MonadStConcreteNames m => MonadStConcreteNames (IdentityT m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

MonadStConcreteNames m => MonadStConcreteNames (ReaderT r m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

MonadStConcreteNames m => MonadStConcreteNames (StateT s m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

(MonadStConcreteNames m, Monoid w) => MonadStConcreteNames (WriterT w m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

class Monad m => MonadTCEnv (m :: Type -> Type) where Source #

MonadTCEnv made into its own dedicated service class. This allows us to use MonadReader for ReaderT extensions of TCM.

Minimal complete definition

Nothing

Methods

askTC :: m TCEnv Source #

default askTC :: forall (t :: (Type -> Type) -> Type -> Type) (n :: Type -> Type). (MonadTrans t, MonadTCEnv n, t n ~ m) => m TCEnv Source #

localTC :: (TCEnv -> TCEnv) -> m a -> m a Source #

default localTC :: forall (t :: (Type -> Type) -> Type -> Type) (n :: Type -> Type) a. (MonadTransControl t, MonadTCEnv n, t n ~ m) => (TCEnv -> TCEnv) -> m a -> m a Source #

Instances

Instances details
MonadTCEnv IM Source # 
Instance details

Defined in Agda.Interaction.Monad

Methods

askTC :: IM TCEnv Source #

localTC :: (TCEnv -> TCEnv) -> IM a -> IM a Source #

MonadTCEnv TerM Source # 
Instance details

Defined in Agda.Termination.Monad

Methods

askTC :: TerM TCEnv Source #

localTC :: (TCEnv -> TCEnv) -> TerM a -> TerM a Source #

MonadTCEnv ReduceM Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

MonadTCEnv NLM Source # 
Instance details

Defined in Agda.TypeChecking.Rewriting.NonLinMatch

Methods

askTC :: NLM TCEnv Source #

localTC :: (TCEnv -> TCEnv) -> NLM a -> NLM a Source #

MonadTCEnv m => MonadTCEnv (PureConversionT m) Source # 
Instance details

Defined in Agda.TypeChecking.Conversion.Pure

MonadTCEnv m => MonadTCEnv (BlockT m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

askTC :: BlockT m TCEnv Source #

localTC :: (TCEnv -> TCEnv) -> BlockT m a -> BlockT m a Source #

MonadIO m => MonadTCEnv (TCMT m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

askTC :: TCMT m TCEnv Source #

localTC :: (TCEnv -> TCEnv) -> TCMT m a -> TCMT m a Source #

MonadTCEnv m => MonadTCEnv (NamesT m) Source # 
Instance details

Defined in Agda.TypeChecking.Names

Methods

askTC :: NamesT m TCEnv Source #

localTC :: (TCEnv -> TCEnv) -> NamesT m a -> NamesT m a Source #

MonadTCEnv m => MonadTCEnv (ListT m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

askTC :: ListT m TCEnv Source #

localTC :: (TCEnv -> TCEnv) -> ListT m a -> ListT m a Source #

MonadTCEnv m => MonadTCEnv (ChangeT m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

askTC :: ChangeT m TCEnv Source #

localTC :: (TCEnv -> TCEnv) -> ChangeT m a -> ChangeT m a Source #

MonadTCEnv m => MonadTCEnv (MaybeT m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

askTC :: MaybeT m TCEnv Source #

localTC :: (TCEnv -> TCEnv) -> MaybeT m a -> MaybeT m a Source #

MonadTCEnv m => MonadTCEnv (ExceptT err m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

askTC :: ExceptT err m TCEnv Source #

localTC :: (TCEnv -> TCEnv) -> ExceptT err m a -> ExceptT err m a Source #

MonadTCEnv m => MonadTCEnv (IdentityT m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

MonadTCEnv m => MonadTCEnv (ReaderT r m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

askTC :: ReaderT r m TCEnv Source #

localTC :: (TCEnv -> TCEnv) -> ReaderT r m a -> ReaderT r m a Source #

MonadTCEnv m => MonadTCEnv (StateT s m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

askTC :: StateT s m TCEnv Source #

localTC :: (TCEnv -> TCEnv) -> StateT s m a -> StateT s m a Source #

(Monoid w, MonadTCEnv m) => MonadTCEnv (WriterT w m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

askTC :: WriterT w m TCEnv Source #

localTC :: (TCEnv -> TCEnv) -> WriterT w m a -> WriterT w m a Source #

type MonadTCError (m :: Type -> Type) = (MonadTCEnv m, ReadTCState m, MonadError TCErr m) Source #

The constraints needed for typeError and similar.

class (Applicative tcm, MonadIO tcm, MonadTCEnv tcm, MonadTCState tcm, HasOptions tcm) => MonadTCM (tcm :: Type -> Type) where Source #

Embedding a TCM computation.

Minimal complete definition

Nothing

Methods

liftTCM :: TCM a -> tcm a Source #

default liftTCM :: forall (m :: Type -> Type) (t :: (Type -> Type) -> Type -> Type) a. (MonadTCM m, MonadTrans t, tcm ~ t m) => TCM a -> tcm a Source #

Instances

Instances details
MonadTCM IM Source # 
Instance details

Defined in Agda.Interaction.Monad

Methods

liftTCM :: TCM a -> IM a Source #

MonadTCM TerM Source # 
Instance details

Defined in Agda.Termination.Monad

Methods

liftTCM :: TCM a -> TerM a Source #

MonadTCM m => MonadTCM (BlockT m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

liftTCM :: TCM a -> BlockT m a Source #

MonadIO m => MonadTCM (TCMT m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

liftTCM :: TCM a -> TCMT m a Source #

MonadTCM m => MonadTCM (NamesT m) Source # 
Instance details

Defined in Agda.TypeChecking.Names

Methods

liftTCM :: TCM a -> NamesT m a Source #

MonadTCM tcm => MonadTCM (ListT tcm) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

liftTCM :: TCM a -> ListT tcm a Source #

MonadTCM tcm => MonadTCM (ChangeT tcm) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

liftTCM :: TCM a -> ChangeT tcm a Source #

MonadTCM tcm => MonadTCM (MaybeT tcm) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

liftTCM :: TCM a -> MaybeT tcm a Source #

MonadTCM tcm => MonadTCM (ExceptT err tcm) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

liftTCM :: TCM a -> ExceptT err tcm a Source #

MonadTCM tcm => MonadTCM (IdentityT tcm) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

liftTCM :: TCM a -> IdentityT tcm a Source #

MonadTCM tcm => MonadTCM (ReaderT r tcm) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

liftTCM :: TCM a -> ReaderT r tcm a Source #

MonadTCM tcm => MonadTCM (StateT s tcm) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

liftTCM :: TCM a -> StateT s tcm a Source #

(Monoid w, MonadTCM tcm) => MonadTCM (WriterT w tcm) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

liftTCM :: TCM a -> WriterT w tcm a Source #

class Monad m => MonadTCState (m :: Type -> Type) where Source #

MonadTCState made into its own dedicated service class. This allows us to use MonadState for StateT extensions of TCM.

Minimal complete definition

Nothing

Methods

getTC :: m TCState Source #

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 #

modifyTC :: (TCState -> TCState) -> m () Source #

default modifyTC :: forall (t :: (Type -> Type) -> Type -> Type) (n :: Type -> Type). (MonadTrans t, MonadTCState n, t n ~ m) => (TCState -> TCState) -> m () Source #

Instances

Instances details
MonadTCState IM Source # 
Instance details

Defined in Agda.Interaction.Monad

MonadTCState TerM Source # 
Instance details

Defined in Agda.Termination.Monad

MonadTCState m => MonadTCState (BlockT m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

MonadIO m => MonadTCState (TCMT m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

MonadTCState m => MonadTCState (NamesT m) Source # 
Instance details

Defined in Agda.TypeChecking.Names

MonadTCState m => MonadTCState (ListT m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

MonadTCState m => MonadTCState (ChangeT m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

MonadTCState m => MonadTCState (MaybeT m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

MonadTCState m => MonadTCState (ExceptT err m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

getTC :: ExceptT err m TCState Source #

putTC :: TCState -> ExceptT err m () Source #

modifyTC :: (TCState -> TCState) -> ExceptT err m () Source #

MonadTCState m => MonadTCState (IdentityT m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

MonadTCState m => MonadTCState (ReaderT r m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

MonadTCState m => MonadTCState (StateT s m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

getTC :: StateT s m TCState Source #

putTC :: TCState -> StateT s m () Source #

modifyTC :: (TCState -> TCState) -> StateT s m () Source #

(Monoid w, MonadTCState m) => MonadTCState (WriterT w m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

data MutualBlock Source #

A mutual block of names in the signature.

Constructors

MutualBlock 

Fields

Instances

Instances details
Null MutualBlock Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

NFData MutualBlock Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: MutualBlock -> () #

Generic MutualBlock Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep MutualBlock 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep MutualBlock = D1 ('MetaData "MutualBlock" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "MutualBlock" 'PrefixI 'True) (S1 ('MetaSel ('Just "mutualInfo") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 MutualInfo) :*: S1 ('MetaSel ('Just "mutualNames") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Set QName))))
Show MutualBlock Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Eq MutualBlock Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep MutualBlock Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep MutualBlock = D1 ('MetaData "MutualBlock" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "MutualBlock" 'PrefixI 'True) (S1 ('MetaSel ('Just "mutualInfo") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 MutualInfo) :*: S1 ('MetaSel ('Just "mutualNames") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Set QName))))

newtype MutualId Source #

Constructors

MutualId Word32 

Instances

Instances details
Pretty MutualId Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

KillRange MutualId Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

HasFresh MutualId Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

EmbPrj MutualId Source # 
Instance details

Defined in Agda.TypeChecking.Serialise.Instances.Internal

NFData MutualId Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: MutualId -> () #

Enum MutualId Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Num MutualId Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Show MutualId Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Eq MutualId Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Ord MutualId Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

data NLPSort Source #

Instances

Instances details
TermLike NLPSort Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

traverseTermM :: Monad m => (Term -> m Term) -> NLPSort -> m NLPSort Source #

foldTerm :: Monoid m => (Term -> m) -> NLPSort -> m Source #

AllMetas NLPSort Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

allMetas :: Monoid m => (MetaId -> m) -> NLPSort -> m Source #

NamesIn NLPSort Source # 
Instance details

Defined in Agda.Syntax.Internal.Names

Methods

namesAndMetasIn' :: Monoid m => (Either QName MetaId -> m) -> NLPSort -> m Source #

KillRange NLPSort Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Free NLPSort Source # 
Instance details

Defined in Agda.TypeChecking.Rewriting.NonLinPattern

Methods

freeVars' :: IsVarSet a c => NLPSort -> FreeM a c Source #

PrettyTCM NLPSort Source # 
Instance details

Defined in Agda.TypeChecking.Pretty

Methods

prettyTCM :: MonadPretty m => NLPSort -> m Doc Source #

InstantiateFull NLPSort Source # 
Instance details

Defined in Agda.TypeChecking.Reduce

GetMatchables NLPSort Source # 
Instance details

Defined in Agda.TypeChecking.Rewriting.NonLinPattern

NLPatVars NLPSort Source # 
Instance details

Defined in Agda.TypeChecking.Rewriting.NonLinPattern

EmbPrj NLPSort Source # 
Instance details

Defined in Agda.TypeChecking.Serialise.Instances.Internal

Subst NLPSort Source # 
Instance details

Defined in Agda.TypeChecking.Substitute

Associated Types

type SubstArg NLPSort 
Instance details

Defined in Agda.TypeChecking.Substitute

NFData NLPSort Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: NLPSort -> () #

Generic NLPSort Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep NLPSort 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

from :: NLPSort -> Rep NLPSort x #

to :: Rep NLPSort x -> NLPSort #

Show NLPSort Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Match NLPSort Sort Source # 
Instance details

Defined in Agda.TypeChecking.Rewriting.NonLinMatch

Methods

match :: Relevance -> Telescope -> Context -> TypeOf Sort -> NLPSort -> Sort -> NLM () Source #

NLPatToTerm NLPSort Sort Source # 
Instance details

Defined in Agda.TypeChecking.Rewriting.NonLinPattern

Methods

nlPatToTerm :: PureTCM m => NLPSort -> m Sort Source #

PatternFrom Sort NLPSort Source # 
Instance details

Defined in Agda.TypeChecking.Rewriting.NonLinPattern

type SubstArg NLPSort Source # 
Instance details

Defined in Agda.TypeChecking.Substitute

type Rep NLPSort Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

data NLPType Source #

Constructors

NLPType 

Instances

Instances details
TermLike NLPType Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

traverseTermM :: Monad m => (Term -> m Term) -> NLPType -> m NLPType Source #

foldTerm :: Monoid m => (Term -> m) -> NLPType -> m Source #

AllMetas NLPType Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

allMetas :: Monoid m => (MetaId -> m) -> NLPType -> m Source #

NamesIn NLPType Source # 
Instance details

Defined in Agda.Syntax.Internal.Names

Methods

namesAndMetasIn' :: Monoid m => (Either QName MetaId -> m) -> NLPType -> m Source #

KillRange NLPType Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Free NLPType Source # 
Instance details

Defined in Agda.TypeChecking.Rewriting.NonLinPattern

Methods

freeVars' :: IsVarSet a c => NLPType -> FreeM a c Source #

PrettyTCM NLPType Source # 
Instance details

Defined in Agda.TypeChecking.Pretty

Methods

prettyTCM :: MonadPretty m => NLPType -> m Doc Source #

InstantiateFull NLPType Source # 
Instance details

Defined in Agda.TypeChecking.Reduce

GetMatchables NLPType Source # 
Instance details

Defined in Agda.TypeChecking.Rewriting.NonLinPattern

NLPatVars NLPType Source # 
Instance details

Defined in Agda.TypeChecking.Rewriting.NonLinPattern

EmbPrj NLPType Source # 
Instance details

Defined in Agda.TypeChecking.Serialise.Instances.Internal

Subst NLPType Source # 
Instance details

Defined in Agda.TypeChecking.Substitute

Associated Types

type SubstArg NLPType 
Instance details

Defined in Agda.TypeChecking.Substitute

NFData NLPType Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: NLPType -> () #

Generic NLPType Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep NLPType 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep NLPType = D1 ('MetaData "NLPType" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "NLPType" 'PrefixI 'True) (S1 ('MetaSel ('Just "nlpTypeSort") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 NLPSort) :*: S1 ('MetaSel ('Just "nlpTypeUnEl") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 NLPat)))

Methods

from :: NLPType -> Rep NLPType x #

to :: Rep NLPType x -> NLPType #

Show NLPType Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Match NLPType Type Source # 
Instance details

Defined in Agda.TypeChecking.Rewriting.NonLinMatch

Methods

match :: Relevance -> Telescope -> Context -> TypeOf Type -> NLPType -> Type -> NLM () Source #

NLPatToTerm NLPType Type Source # 
Instance details

Defined in Agda.TypeChecking.Rewriting.NonLinPattern

Methods

nlPatToTerm :: PureTCM m => NLPType -> m Type Source #

PatternFrom Type NLPType Source # 
Instance details

Defined in Agda.TypeChecking.Rewriting.NonLinPattern

type SubstArg NLPType Source # 
Instance details

Defined in Agda.TypeChecking.Substitute

type Rep NLPType Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep NLPType = D1 ('MetaData "NLPType" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "NLPType" 'PrefixI 'True) (S1 ('MetaSel ('Just "nlpTypeSort") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 NLPSort) :*: S1 ('MetaSel ('Just "nlpTypeUnEl") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 NLPat)))

data NLPat Source #

Non-linear (non-constructor) first-order pattern.

Constructors

PVar !Int [Arg Int]

Matches anything (modulo non-linearity) that only contains bound variables that occur in the given arguments.

PDef QName PElims

Matches f es

PLam ArgInfo (Abs NLPat)

Matches λ x → t

PPi (Dom NLPType) (Abs NLPType)

Matches (x : A) → B

PSort NLPSort

Matches a sort of the given shape.

PBoundVar !Int PElims

Matches x es where x is a lambda-bound variable

PTerm Term

Matches the term modulo β (ideally βη).

Instances

Instances details
TermLike NLPat Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

traverseTermM :: Monad m => (Term -> m Term) -> NLPat -> m NLPat Source #

foldTerm :: Monoid m => (Term -> m) -> NLPat -> m Source #

AllMetas NLPat Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

allMetas :: Monoid m => (MetaId -> m) -> NLPat -> m Source #

NamesIn NLPat Source # 
Instance details

Defined in Agda.Syntax.Internal.Names

Methods

namesAndMetasIn' :: Monoid m => (Either QName MetaId -> m) -> NLPat -> m Source #

KillRange NLPat Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Free NLPat Source #

Only computes free variables that are not bound (see nlPatVars), i.e., those in a PTerm.

Instance details

Defined in Agda.TypeChecking.Rewriting.NonLinPattern

Methods

freeVars' :: IsVarSet a c => NLPat -> FreeM a c Source #

PrettyTCM NLPat Source # 
Instance details

Defined in Agda.TypeChecking.Pretty

Methods

prettyTCM :: MonadPretty m => NLPat -> m Doc Source #

InstantiateFull NLPat Source # 
Instance details

Defined in Agda.TypeChecking.Reduce

GetMatchables NLPat Source # 
Instance details

Defined in Agda.TypeChecking.Rewriting.NonLinPattern

NLPatVars NLPat Source # 
Instance details

Defined in Agda.TypeChecking.Rewriting.NonLinPattern

EmbPrj NLPat Source # 
Instance details

Defined in Agda.TypeChecking.Serialise.Instances.Internal

Subst NLPat Source # 
Instance details

Defined in Agda.TypeChecking.Substitute

Associated Types

type SubstArg NLPat 
Instance details

Defined in Agda.TypeChecking.Substitute

DeBruijn NLPat Source # 
Instance details

Defined in Agda.TypeChecking.Substitute

NFData NLPat Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: NLPat -> () #

Generic NLPat Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep NLPat 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep NLPat = D1 ('MetaData "NLPat" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) ((C1 ('MetaCons "PVar" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Int) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Arg Int])) :+: (C1 ('MetaCons "PDef" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 PElims)) :+: C1 ('MetaCons "PLam" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ArgInfo) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Abs NLPat))))) :+: ((C1 ('MetaCons "PPi" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Dom NLPType)) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Abs NLPType))) :+: C1 ('MetaCons "PSort" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 NLPSort))) :+: (C1 ('MetaCons "PBoundVar" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'SourceUnpack 'SourceStrict 'DecidedStrict) (Rec0 Int) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 PElims)) :+: C1 ('MetaCons "PTerm" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)))))

Methods

from :: NLPat -> Rep NLPat x #

to :: Rep NLPat x -> NLPat #

Show NLPat Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

showsPrec :: Int -> NLPat -> ShowS #

show :: NLPat -> String #

showList :: [NLPat] -> ShowS #

Match NLPat Level Source # 
Instance details

Defined in Agda.TypeChecking.Rewriting.NonLinMatch

Methods

match :: Relevance -> Telescope -> Context -> TypeOf Level -> NLPat -> Level -> NLM () Source #

Match NLPat Term Source # 
Instance details

Defined in Agda.TypeChecking.Rewriting.NonLinMatch

Methods

match :: Relevance -> Telescope -> Context -> TypeOf Term -> NLPat -> Term -> NLM () Source #

NLPatToTerm NLPat Level Source # 
Instance details

Defined in Agda.TypeChecking.Rewriting.NonLinPattern

Methods

nlPatToTerm :: PureTCM m => NLPat -> m Level Source #

NLPatToTerm NLPat Term Source # 
Instance details

Defined in Agda.TypeChecking.Rewriting.NonLinPattern

Methods

nlPatToTerm :: PureTCM m => NLPat -> m Term Source #

PatternFrom Level NLPat Source # 
Instance details

Defined in Agda.TypeChecking.Rewriting.NonLinPattern

PatternFrom Term NLPat Source # 
Instance details

Defined in Agda.TypeChecking.Rewriting.NonLinPattern

PatternFrom Elims [Elim' NLPat] Source # 
Instance details

Defined in Agda.TypeChecking.Rewriting.NonLinPattern

PrettyTCM (Type' NLPat) Source # 
Instance details

Defined in Agda.TypeChecking.Pretty

PrettyTCM (Elim' NLPat) Source # 
Instance details

Defined in Agda.TypeChecking.Pretty

Match [Elim' NLPat] Elims Source # 
Instance details

Defined in Agda.TypeChecking.Rewriting.NonLinMatch

Methods

match :: Relevance -> Telescope -> Context -> TypeOf Elims -> [Elim' NLPat] -> Elims -> NLM () Source #

ToNLPat (Arg DeBruijnPattern) (Elim' NLPat) Source # 
Instance details

Defined in Agda.TypeChecking.Rewriting.Clause

ToNLPat (NamedArg DeBruijnPattern) (Elim' NLPat) Source # 
Instance details

Defined in Agda.TypeChecking.Rewriting.Clause

type TypeOf NLPat Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type SubstArg NLPat Source # 
Instance details

Defined in Agda.TypeChecking.Substitute

type Rep NLPat Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep NLPat = D1 ('MetaData "NLPat" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) ((C1 ('MetaCons "PVar" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Int) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Arg Int])) :+: (C1 ('MetaCons "PDef" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 PElims)) :+: C1 ('MetaCons "PLam" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ArgInfo) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Abs NLPat))))) :+: ((C1 ('MetaCons "PPi" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Dom NLPType)) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Abs NLPType))) :+: C1 ('MetaCons "PSort" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 NLPSort))) :+: (C1 ('MetaCons "PBoundVar" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'SourceUnpack 'SourceStrict 'DecidedStrict) (Rec0 Int) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 PElims)) :+: C1 ('MetaCons "PTerm" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)))))
type TypeOf [Elim' NLPat] Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type TypeOf [Elim' NLPat] = (Type, Elims -> Term)

data NegativeUnification Source #

Instances

Instances details
PrettyTCM NegativeUnification Source # 
Instance details

Defined in Agda.TypeChecking.Errors

NFData NegativeUnification Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: NegativeUnification -> () #

Generic NegativeUnification Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Show NegativeUnification Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep NegativeUnification Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

data OpaqueBlock Source #

A block of opaque definitions.

Constructors

OpaqueBlock 

Fields

Instances

Instances details
Pretty OpaqueBlock Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

EmbPrj OpaqueBlock Source # 
Instance details

Defined in Agda.TypeChecking.Serialise.Instances.Internal

NFData OpaqueBlock Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: OpaqueBlock -> () #

Generic OpaqueBlock Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep OpaqueBlock 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep OpaqueBlock = D1 ('MetaData "OpaqueBlock" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "OpaqueBlock" 'PrefixI 'True) ((S1 ('MetaSel ('Just "opaqueId") 'SourceUnpack 'SourceStrict 'DecidedStrict) (Rec0 OpaqueId) :*: S1 ('MetaSel ('Just "opaqueUnfolding") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (HashSet QName))) :*: (S1 ('MetaSel ('Just "opaqueDecls") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (HashSet QName)) :*: (S1 ('MetaSel ('Just "opaqueParent") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe OpaqueId)) :*: S1 ('MetaSel ('Just "opaqueRange") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range)))))
Show OpaqueBlock Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Eq OpaqueBlock Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Hashable OpaqueBlock Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep OpaqueBlock Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep OpaqueBlock = D1 ('MetaData "OpaqueBlock" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "OpaqueBlock" 'PrefixI 'True) ((S1 ('MetaSel ('Just "opaqueId") 'SourceUnpack 'SourceStrict 'DecidedStrict) (Rec0 OpaqueId) :*: S1 ('MetaSel ('Just "opaqueUnfolding") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (HashSet QName))) :*: (S1 ('MetaSel ('Just "opaqueDecls") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (HashSet QName)) :*: (S1 ('MetaSel ('Just "opaqueParent") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe OpaqueId)) :*: S1 ('MetaSel ('Just "opaqueRange") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range)))))

data Open a Source #

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

Instances details
Decoration Open Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

traverseF :: Functor m => (a -> m b) -> Open a -> m (Open b) Source #

distributeF :: Functor m => Open (m a) -> m (Open a) Source #

Functor Open Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

fmap :: (a -> b) -> Open a -> Open b #

(<$) :: a -> Open b -> Open a #

Foldable Open Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

fold :: Monoid m => Open m -> m #

foldMap :: Monoid m => (a -> m) -> Open a -> m #

foldMap' :: Monoid m => (a -> m) -> Open a -> m #

foldr :: (a -> b -> b) -> b -> Open a -> b #

foldr' :: (a -> b -> b) -> b -> Open a -> b #

foldl :: (b -> a -> b) -> b -> Open a -> b #

foldl' :: (b -> a -> b) -> b -> Open a -> b #

foldr1 :: (a -> a -> a) -> Open a -> a #

foldl1 :: (a -> a -> a) -> Open a -> a #

toList :: Open a -> [a] #

null :: Open a -> Bool #

length :: Open a -> Int #

elem :: Eq a => a -> Open a -> Bool #

maximum :: Ord a => Open a -> a #

minimum :: Ord a => Open a -> a #

sum :: Num a => Open a -> a #

product :: Num a => Open a -> a #

Traversable Open Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

traverse :: Applicative f => (a -> f b) -> Open a -> f (Open b) #

sequenceA :: Applicative f => Open (f a) -> f (Open a) #

mapM :: Monad m => (a -> m b) -> Open a -> m (Open b) #

sequence :: Monad m => Open (m a) -> m (Open a) #

Pretty a => Pretty (Open a) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

pretty :: Open a -> Doc Source #

prettyPrec :: Int -> Open a -> Doc Source #

prettyList :: [Open a] -> Doc Source #

NamesIn a => NamesIn (Open a) Source # 
Instance details

Defined in Agda.Syntax.Internal.Names

Methods

namesAndMetasIn' :: Monoid m => (Either QName MetaId -> m) -> Open a -> m Source #

KillRange a => KillRange (Open a) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

ChaseDisplayForms a => ChaseDisplayForms (Open a) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Signature

InstantiateFull t => InstantiateFull (Open t) Source # 
Instance details

Defined in Agda.TypeChecking.Reduce

EmbPrj a => EmbPrj (Open a) Source # 
Instance details

Defined in Agda.TypeChecking.Serialise.Instances.Internal

Methods

icode :: Open a -> S Word32 Source #

icod_ :: Open a -> S Word32 Source #

value :: Word32 -> R (Open a) Source #

NFData a => NFData (Open a) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: Open a -> () #

Generic (Open a) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep (Open a) 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep (Open a) = D1 ('MetaData "Open" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "OpenThing" 'PrefixI 'True) ((S1 ('MetaSel ('Just "openThingCheckpoint") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 CheckpointId) :*: S1 ('MetaSel ('Just "openThingCheckpointMap") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Map CheckpointId Substitution))) :*: (S1 ('MetaSel ('Just "openThingModule") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ModuleNameHash) :*: S1 ('MetaSel ('Just "openThing") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 a))))

Methods

from :: Open a -> Rep (Open a) x #

to :: Rep (Open a) x -> Open a #

Show a => Show (Open a) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

showsPrec :: Int -> Open a -> ShowS #

show :: Open a -> String #

showList :: [Open a] -> ShowS #

type Rep (Open a) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep (Open a) = D1 ('MetaData "Open" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "OpenThing" 'PrefixI 'True) ((S1 ('MetaSel ('Just "openThingCheckpoint") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 CheckpointId) :*: S1 ('MetaSel ('Just "openThingCheckpointMap") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Map CheckpointId Substitution))) :*: (S1 ('MetaSel ('Just "openThingModule") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ModuleNameHash) :*: S1 ('MetaSel ('Just "openThing") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 a))))

data Overapplied Source #

Flag to indicate whether the meta is overapplied in the constraint. A meta is overapplied if it has more arguments than the size of the telescope in its creation environment (as stored in MetaInfo).

Instances

Instances details
NFData Overapplied Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: Overapplied -> () #

Generic Overapplied Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep Overapplied 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep Overapplied = D1 ('MetaData "Overapplied" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "Overapplied" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "NotOverapplied" 'PrefixI 'False) (U1 :: Type -> Type))
Show Overapplied Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Eq Overapplied Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep Overapplied Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep Overapplied = D1 ('MetaData "Overapplied" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "Overapplied" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "NotOverapplied" 'PrefixI 'False) (U1 :: Type -> Type))

data PersistentTCState Source #

A part of the state which is not reverted when an error is thrown or the state is reset.

Constructors

PersistentTCSt 

Fields

Instances

Instances details
LensCommandLineOptions PersistentTCState Source # 
Instance details

Defined in Agda.Interaction.Options.Lenses

LensIncludePaths PersistentTCState Source # 
Instance details

Defined in Agda.Interaction.Options.Lenses

LensPersistentVerbosity PersistentTCState Source # 
Instance details

Defined in Agda.Interaction.Options.Lenses

LensSafeMode PersistentTCState Source # 
Instance details

Defined in Agda.Interaction.Options.Lenses

NFData PersistentTCState Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: PersistentTCState -> () #

Generic PersistentTCState Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep PersistentTCState 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep PersistentTCState = D1 ('MetaData "PersistentTCState" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "PersistentTCSt" 'PrefixI 'True) ((S1 ('MetaSel ('Just "stPersistentSession") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 SessionTCState) :*: (S1 ('MetaSel ('Just "stDecodedModules") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 DecodedModules) :*: S1 ('MetaSel ('Just "stPersistentTopLevelModuleNames") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (BiMap RawTopLevelModuleName ModuleNameHash)))) :*: ((S1 ('MetaSel ('Just "stPersistentOptions") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 CommandLineOptions) :*: S1 ('MetaSel ('Just "stInteractionOutputCallback") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 InteractionOutputCallback)) :*: (S1 ('MetaSel ('Just "stAccumStatistics") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Statistics) :*: S1 ('MetaSel ('Just "stPersistLoadedFileCache") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (Maybe LoadedFileCache))))))
type Rep PersistentTCState Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep PersistentTCState = D1 ('MetaData "PersistentTCState" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "PersistentTCSt" 'PrefixI 'True) ((S1 ('MetaSel ('Just "stPersistentSession") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 SessionTCState) :*: (S1 ('MetaSel ('Just "stDecodedModules") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 DecodedModules) :*: S1 ('MetaSel ('Just "stPersistentTopLevelModuleNames") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (BiMap RawTopLevelModuleName ModuleNameHash)))) :*: ((S1 ('MetaSel ('Just "stPersistentOptions") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 CommandLineOptions) :*: S1 ('MetaSel ('Just "stInteractionOutputCallback") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 InteractionOutputCallback)) :*: (S1 ('MetaSel ('Just "stAccumStatistics") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Statistics) :*: S1 ('MetaSel ('Just "stPersistLoadedFileCache") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (Maybe LoadedFileCache))))))

data PostScopeState Source #

Constructors

PostScopeState 

Fields

Instances

Instances details
NFData PostScopeState Source #

This instance could be optimised, some things are guaranteed to be forced.

Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: PostScopeState -> () #

Generic PostScopeState Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep PostScopeState 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep PostScopeState = D1 ('MetaData "PostScopeState" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "PostScopeState" 'PrefixI 'True) (((((S1 ('MetaSel ('Just "stPostSyntaxInfo") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 HighlightingInfo) :*: S1 ('MetaSel ('Just "stPostDisambiguatedNames") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 DisambiguatedNames)) :*: (S1 ('MetaSel ('Just "stPostOpenMetaStore") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 LocalMetaStore) :*: S1 ('MetaSel ('Just "stPostSolvedMetaStore") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 LocalMetaStore))) :*: ((S1 ('MetaSel ('Just "stPostInteractionPoints") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 InteractionPoints) :*: S1 ('MetaSel ('Just "stPostAwakeConstraints") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Constraints)) :*: (S1 ('MetaSel ('Just "stPostSleepingConstraints") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Constraints) :*: (S1 ('MetaSel ('Just "stPostDirty") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Bool) :*: S1 ('MetaSel ('Just "stPostOccursCheckDefs") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (Set QName)))))) :*: (((S1 ('MetaSel ('Just "stPostSignature") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Signature) :*: S1 ('MetaSel ('Just "stPostModuleCheckpoints") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (Map ModuleName CheckpointId))) :*: (S1 ('MetaSel ('Just "stPostImportsDisplayForms") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 DisplayForms) :*: (S1 ('MetaSel ('Just "stPostForeignCode") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 BackendForeignCode) :*: S1 ('MetaSel ('Just "stPostCurrentModule") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (Maybe (ModuleName, TopLevelModuleName)))))) :*: ((S1 ('MetaSel ('Just "stPostPendingInstances") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (Set QName)) :*: S1 ('MetaSel ('Just "stPostTemporaryInstances") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (Set QName))) :*: (S1 ('MetaSel ('Just "stPostConcreteNames") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 ConcreteNames) :*: (S1 ('MetaSel ('Just "stPostUsedNames") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 UsedNames) :*: S1 ('MetaSel ('Just "stPostShadowingNames") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 ShadowingNames)))))) :*: ((((S1 ('MetaSel ('Just "stPostStatistics") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Statistics) :*: S1 ('MetaSel ('Just "stPostTCWarnings") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (Set TCWarning))) :*: (S1 ('MetaSel ('Just "stPostMutualBlocks") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 MutualBlocks) :*: S1 ('MetaSel ('Just "stPostLocalBuiltins") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 BuiltinThings))) :*: ((S1 ('MetaSel ('Just "stPostFreshMetaId") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 MetaId) :*: S1 ('MetaSel ('Just "stPostFreshMutualId") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 MutualId)) :*: (S1 ('MetaSel ('Just "stPostFreshProblemId") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 ProblemId) :*: (S1 ('MetaSel ('Just "stPostFreshCheckpointId") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 CheckpointId) :*: S1 ('MetaSel ('Just "stPostFreshInt") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Int))))) :*: (((S1 ('MetaSel ('Just "stPostFreshNameId") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 NameId) :*: S1 ('MetaSel ('Just "stPostFreshOpaqueId") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 OpaqueId)) :*: (S1 ('MetaSel ('Just "stPostAreWeCaching") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Bool) :*: (S1 ('MetaSel ('Just "stPostPostponeInstanceSearch") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Bool) :*: S1 ('MetaSel ('Just "stPostConsideringInstance") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Bool)))) :*: ((S1 ('MetaSel ('Just "stPostInstantiateBlocking") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Bool) :*: S1 ('MetaSel ('Just "stPostLocalPartialDefs") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (Set QName))) :*: (S1 ('MetaSel ('Just "stPostOpaqueBlocks") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Map OpaqueId OpaqueBlock)) :*: (S1 ('MetaSel ('Just "stPostOpaqueIds") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Map QName OpaqueId)) :*: S1 ('MetaSel ('Just "stPostInstanceHack") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Bool))))))))
type Rep PostScopeState Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep PostScopeState = D1 ('MetaData "PostScopeState" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "PostScopeState" 'PrefixI 'True) (((((S1 ('MetaSel ('Just "stPostSyntaxInfo") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 HighlightingInfo) :*: S1 ('MetaSel ('Just "stPostDisambiguatedNames") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 DisambiguatedNames)) :*: (S1 ('MetaSel ('Just "stPostOpenMetaStore") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 LocalMetaStore) :*: S1 ('MetaSel ('Just "stPostSolvedMetaStore") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 LocalMetaStore))) :*: ((S1 ('MetaSel ('Just "stPostInteractionPoints") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 InteractionPoints) :*: S1 ('MetaSel ('Just "stPostAwakeConstraints") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Constraints)) :*: (S1 ('MetaSel ('Just "stPostSleepingConstraints") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Constraints) :*: (S1 ('MetaSel ('Just "stPostDirty") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Bool) :*: S1 ('MetaSel ('Just "stPostOccursCheckDefs") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (Set QName)))))) :*: (((S1 ('MetaSel ('Just "stPostSignature") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Signature) :*: S1 ('MetaSel ('Just "stPostModuleCheckpoints") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (Map ModuleName CheckpointId))) :*: (S1 ('MetaSel ('Just "stPostImportsDisplayForms") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 DisplayForms) :*: (S1 ('MetaSel ('Just "stPostForeignCode") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 BackendForeignCode) :*: S1 ('MetaSel ('Just "stPostCurrentModule") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (Maybe (ModuleName, TopLevelModuleName)))))) :*: ((S1 ('MetaSel ('Just "stPostPendingInstances") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (Set QName)) :*: S1 ('MetaSel ('Just "stPostTemporaryInstances") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (Set QName))) :*: (S1 ('MetaSel ('Just "stPostConcreteNames") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 ConcreteNames) :*: (S1 ('MetaSel ('Just "stPostUsedNames") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 UsedNames) :*: S1 ('MetaSel ('Just "stPostShadowingNames") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 ShadowingNames)))))) :*: ((((S1 ('MetaSel ('Just "stPostStatistics") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Statistics) :*: S1 ('MetaSel ('Just "stPostTCWarnings") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (Set TCWarning))) :*: (S1 ('MetaSel ('Just "stPostMutualBlocks") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 MutualBlocks) :*: S1 ('MetaSel ('Just "stPostLocalBuiltins") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 BuiltinThings))) :*: ((S1 ('MetaSel ('Just "stPostFreshMetaId") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 MetaId) :*: S1 ('MetaSel ('Just "stPostFreshMutualId") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 MutualId)) :*: (S1 ('MetaSel ('Just "stPostFreshProblemId") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 ProblemId) :*: (S1 ('MetaSel ('Just "stPostFreshCheckpointId") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 CheckpointId) :*: S1 ('MetaSel ('Just "stPostFreshInt") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Int))))) :*: (((S1 ('MetaSel ('Just "stPostFreshNameId") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 NameId) :*: S1 ('MetaSel ('Just "stPostFreshOpaqueId") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 OpaqueId)) :*: (S1 ('MetaSel ('Just "stPostAreWeCaching") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Bool) :*: (S1 ('MetaSel ('Just "stPostPostponeInstanceSearch") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Bool) :*: S1 ('MetaSel ('Just "stPostConsideringInstance") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Bool)))) :*: ((S1 ('MetaSel ('Just "stPostInstantiateBlocking") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Bool) :*: S1 ('MetaSel ('Just "stPostLocalPartialDefs") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (Set QName))) :*: (S1 ('MetaSel ('Just "stPostOpaqueBlocks") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Map OpaqueId OpaqueBlock)) :*: (S1 ('MetaSel ('Just "stPostOpaqueIds") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Map QName OpaqueId)) :*: S1 ('MetaSel ('Just "stPostInstanceHack") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Bool))))))))

data PreScopeState Source #

Constructors

PreScopeState 

Fields

Instances

Instances details
LensPragmaOptions PreScopeState Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

NFData PreScopeState Source #

This instance could be optimised, some things are guaranteed to be forced.

Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: PreScopeState -> () #

Generic PreScopeState Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep PreScopeState 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep PreScopeState = D1 ('MetaData "PreScopeState" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "PreScopeState" 'PrefixI 'True) ((((S1 ('MetaSel ('Just "stPreTokens") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 HighlightingInfo) :*: S1 ('MetaSel ('Just "stPreImports") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Signature)) :*: (S1 ('MetaSel ('Just "stPreImportedModules") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 ImportedModules) :*: (S1 ('MetaSel ('Just "stPreModuleToSourceId") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 ModuleToSourceId) :*: S1 ('MetaSel ('Just "stPreVisitedModules") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 VisitedModules)))) :*: ((S1 ('MetaSel ('Just "stPreScope") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 ScopeInfo) :*: S1 ('MetaSel ('Just "stPrePatternSyns") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 PatternSynDefns)) :*: (S1 ('MetaSel ('Just "stPrePatternSynImports") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 PatternSynDefns) :*: (S1 ('MetaSel ('Just "stPreGeneralizedVars") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (Maybe (Set QName))) :*: S1 ('MetaSel ('Just "stPrePragmaOptions") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 PragmaOptions))))) :*: (((S1 ('MetaSel ('Just "stPreImportedBuiltins") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 BuiltinThings) :*: S1 ('MetaSel ('Just "stPreImportedDisplayForms") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 DisplayForms)) :*: (S1 ('MetaSel ('Just "stPreFreshInteractionId") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 InteractionId) :*: (S1 ('MetaSel ('Just "stPreImportedUserWarnings") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 UserWarnings) :*: S1 ('MetaSel ('Just "stPreLocalUserWarnings") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 UserWarnings)))) :*: ((S1 ('MetaSel ('Just "stPreWarningOnImport") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (Maybe Text)) :*: (S1 ('MetaSel ('Just "stPreImportedPartialDefs") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (Set QName)) :*: S1 ('MetaSel ('Just "stPreLibCache") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 LibCache))) :*: (S1 ('MetaSel ('Just "stPreImportedMetaStore") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 RemoteMetaStore) :*: (S1 ('MetaSel ('Just "stPreCopiedNames") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (HashMap QName QName)) :*: S1 ('MetaSel ('Just "stPreNameCopies") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (HashMap QName (HashSet QName)))))))))
type Rep PreScopeState Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep PreScopeState = D1 ('MetaData "PreScopeState" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "PreScopeState" 'PrefixI 'True) ((((S1 ('MetaSel ('Just "stPreTokens") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 HighlightingInfo) :*: S1 ('MetaSel ('Just "stPreImports") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Signature)) :*: (S1 ('MetaSel ('Just "stPreImportedModules") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 ImportedModules) :*: (S1 ('MetaSel ('Just "stPreModuleToSourceId") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 ModuleToSourceId) :*: S1 ('MetaSel ('Just "stPreVisitedModules") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 VisitedModules)))) :*: ((S1 ('MetaSel ('Just "stPreScope") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 ScopeInfo) :*: S1 ('MetaSel ('Just "stPrePatternSyns") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 PatternSynDefns)) :*: (S1 ('MetaSel ('Just "stPrePatternSynImports") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 PatternSynDefns) :*: (S1 ('MetaSel ('Just "stPreGeneralizedVars") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (Maybe (Set QName))) :*: S1 ('MetaSel ('Just "stPrePragmaOptions") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 PragmaOptions))))) :*: (((S1 ('MetaSel ('Just "stPreImportedBuiltins") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 BuiltinThings) :*: S1 ('MetaSel ('Just "stPreImportedDisplayForms") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 DisplayForms)) :*: (S1 ('MetaSel ('Just "stPreFreshInteractionId") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 InteractionId) :*: (S1 ('MetaSel ('Just "stPreImportedUserWarnings") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 UserWarnings) :*: S1 ('MetaSel ('Just "stPreLocalUserWarnings") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 UserWarnings)))) :*: ((S1 ('MetaSel ('Just "stPreWarningOnImport") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (Maybe Text)) :*: (S1 ('MetaSel ('Just "stPreImportedPartialDefs") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (Set QName)) :*: S1 ('MetaSel ('Just "stPreLibCache") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 LibCache))) :*: (S1 ('MetaSel ('Just "stPreImportedMetaStore") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 RemoteMetaStore) :*: (S1 ('MetaSel ('Just "stPreCopiedNames") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (HashMap QName QName)) :*: S1 ('MetaSel ('Just "stPreNameCopies") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (HashMap QName (HashSet QName)))))))))

data PrimFun Source #

Instances

Instances details
NamesIn PrimFun Source #

Note that the primFunImplementation is skipped.

Instance details

Defined in Agda.Syntax.Internal.Names

Methods

namesAndMetasIn' :: Monoid m => (Either QName MetaId -> m) -> PrimFun -> m Source #

Abstract PrimFun Source # 
Instance details

Defined in Agda.TypeChecking.Substitute

Apply PrimFun Source # 
Instance details

Defined in Agda.TypeChecking.Substitute

NFData PrimFun Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: PrimFun -> () #

Generic PrimFun Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep PrimFun 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep PrimFun = D1 ('MetaData "PrimFun" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "PrimFun" 'PrefixI 'True) ((S1 ('MetaSel ('Just "primFunName") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Just "primFunArity") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Arity)) :*: (S1 ('MetaSel ('Just "primFunArgOccurrences") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Occurrence]) :*: S1 ('MetaSel ('Just "primFunImplementation") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ([Arg Term] -> Int -> ReduceM (Reduced MaybeReducedArgs Term))))))

Methods

from :: PrimFun -> Rep PrimFun x #

to :: Rep PrimFun x -> PrimFun #

type Rep PrimFun Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep PrimFun = D1 ('MetaData "PrimFun" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "PrimFun" 'PrefixI 'True) ((S1 ('MetaSel ('Just "primFunName") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Just "primFunArity") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Arity)) :*: (S1 ('MetaSel ('Just "primFunArgOccurrences") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Occurrence]) :*: S1 ('MetaSel ('Just "primFunImplementation") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ([Arg Term] -> Int -> ReduceM (Reduced MaybeReducedArgs Term))))))

data PrimitiveData Source #

Constructors

PrimitiveData 

Fields

Instances

Instances details
Pretty PrimitiveData Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

NFData PrimitiveData Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: PrimitiveData -> () #

Generic PrimitiveData Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Show PrimitiveData Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep PrimitiveData Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

data PrimitiveImpl Source #

Primitives

Constructors

PrimImpl Type PrimFun 

data PrimitiveSortData Source #

Instances

Instances details
Pretty PrimitiveSortData Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

NFData PrimitiveSortData Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: PrimitiveSortData -> () #

Generic PrimitiveSortData Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep PrimitiveSortData 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep PrimitiveSortData = D1 ('MetaData "PrimitiveSortData" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "PrimitiveSortData" 'PrefixI 'True) (S1 ('MetaSel ('Just "_primSortName") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 BuiltinSort) :*: S1 ('MetaSel ('Just "_primSortSort") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Sort)))
Show PrimitiveSortData Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep PrimitiveSortData Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep PrimitiveSortData = D1 ('MetaData "PrimitiveSortData" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "PrimitiveSortData" 'PrefixI 'True) (S1 ('MetaSel ('Just "_primSortName") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 BuiltinSort) :*: S1 ('MetaSel ('Just "_primSortSort") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Sort)))

data PrincipalArgTypeMetas Source #

Constructors

PrincipalArgTypeMetas 

Fields

  • patmMetas :: Args

    metas created for hidden and instance arguments in the principal argument's type

  • patmRemainder :: Type

    principal argument's type, stripped of hidden and instance arguments

Instances

Instances details
NFData PrincipalArgTypeMetas Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: PrincipalArgTypeMetas -> () #

Generic PrincipalArgTypeMetas Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep PrincipalArgTypeMetas 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep PrincipalArgTypeMetas = D1 ('MetaData "PrincipalArgTypeMetas" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "PrincipalArgTypeMetas" 'PrefixI 'True) (S1 ('MetaSel ('Just "patmMetas") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Args) :*: S1 ('MetaSel ('Just "patmRemainder") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)))
type Rep PrincipalArgTypeMetas Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep PrincipalArgTypeMetas = D1 ('MetaData "PrincipalArgTypeMetas" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "PrincipalArgTypeMetas" 'PrefixI 'True) (S1 ('MetaSel ('Just "patmMetas") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Args) :*: S1 ('MetaSel ('Just "patmRemainder") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)))

data ProblemConstraint Source #

Instances

Instances details
HasRange ProblemConstraint Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Reify ProblemConstraint Source # 
Instance details

Defined in Agda.Interaction.BasicOps

MentionsMeta ProblemConstraint Source # 
Instance details

Defined in Agda.TypeChecking.MetaVars.Mention

PrettyTCM ProblemConstraint Source # 
Instance details

Defined in Agda.TypeChecking.Pretty.Constraint

InstantiateFull ProblemConstraint Source # 
Instance details

Defined in Agda.TypeChecking.Reduce

Normalise ProblemConstraint Source # 
Instance details

Defined in Agda.TypeChecking.Reduce

Simplify ProblemConstraint Source # 
Instance details

Defined in Agda.TypeChecking.Reduce

NFData ProblemConstraint Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: ProblemConstraint -> () #

Generic ProblemConstraint Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep ProblemConstraint 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep ProblemConstraint = D1 ('MetaData "ProblemConstraint" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "PConstr" 'PrefixI 'True) (S1 ('MetaSel ('Just "constraintProblems") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Set ProblemId)) :*: (S1 ('MetaSel ('Just "constraintUnblocker") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Blocker) :*: S1 ('MetaSel ('Just "theConstraint") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Closure Constraint)))))
Show ProblemConstraint Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type ReifiesTo ProblemConstraint Source # 
Instance details

Defined in Agda.Interaction.BasicOps

type Rep ProblemConstraint Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep ProblemConstraint = D1 ('MetaData "ProblemConstraint" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "PConstr" 'PrefixI 'True) (S1 ('MetaSel ('Just "constraintProblems") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Set ProblemId)) :*: (S1 ('MetaSel ('Just "constraintUnblocker") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Blocker) :*: S1 ('MetaSel ('Just "theConstraint") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Closure Constraint)))))

newtype ProjLams Source #

Abstractions to build projection function (dropping parameters).

Constructors

ProjLams 

Fields

Instances

Instances details
Pretty ProjLams Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

KillRange ProjLams Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

EmbPrj ProjLams Source # 
Instance details

Defined in Agda.TypeChecking.Serialise.Instances.Internal

Abstract ProjLams Source # 
Instance details

Defined in Agda.TypeChecking.Substitute

Apply ProjLams Source # 
Instance details

Defined in Agda.TypeChecking.Substitute

Null ProjLams Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

NFData ProjLams Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: ProjLams -> () #

Generic ProjLams Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep ProjLams 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep ProjLams = D1 ('MetaData "ProjLams" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'True) (C1 ('MetaCons "ProjLams" 'PrefixI 'True) (S1 ('MetaSel ('Just "getProjLams") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Arg ArgName])))

Methods

from :: ProjLams -> Rep ProjLams x #

to :: Rep ProjLams x -> ProjLams #

Show ProjLams Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep ProjLams Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep ProjLams = D1 ('MetaData "ProjLams" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'True) (C1 ('MetaCons "ProjLams" 'PrefixI 'True) (S1 ('MetaSel ('Just "getProjLams") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Arg ArgName])))

data Projection Source #

Additional information for projection Functions.

Constructors

Projection 

Fields

  • projProper :: Maybe QName

    Nothing if only projection-like, Just r if record projection. The r is the name of the record type projected from. This field is updated by module application.

  • projOrig :: QName

    The original projection name (current name could be from module application).

  • projFromType :: Arg QName

    Type projected from. Original record type if projProper = Just{}. Also stores ArgInfo of the principal argument. This field is unchanged by module application.

  • projIndex :: Int

    Index of the record argument. Start counting with 1, because 0 means that it is already applied to the record value. This can happen in module instantiation, but then either the record value is var 0, or funProjection == Left _.

  • projLams :: ProjLams

    Term t to be be applied to record parameters and record value. The parameters will be dropped. In case of a proper projection, a postfix projection application will be created: t = pars r -> r .p (Invariant: the number of abstractions equals projIndex.) In case of a projection-like function, just the function symbol is returned as Def: t = pars -> f.

Instances

Instances details
Pretty Projection Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

KillRange Projection Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

EmbPrj Projection Source # 
Instance details

Defined in Agda.TypeChecking.Serialise.Instances.Internal

Abstract Projection Source # 
Instance details

Defined in Agda.TypeChecking.Substitute

Apply Projection Source # 
Instance details

Defined in Agda.TypeChecking.Substitute

NFData Projection Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: Projection -> () #

Generic Projection Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep Projection 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Show Projection Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep Projection Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

data ProjectionLikenessMissing Source #

Indicates the reason behind a function having not been marked projection-like.

Constructors

MaybeProjection

Projection-likeness analysis has not run on this function yet. It may do so in the future.

NeverProjection

The user has requested that this function be not be marked projection-like. The analysis may already have run on this function, but the results have been discarded, and it will not be run again.

Instances

Instances details
Pretty ProjectionLikenessMissing Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

KillRange ProjectionLikenessMissing Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

EmbPrj ProjectionLikenessMissing Source # 
Instance details

Defined in Agda.TypeChecking.Serialise.Instances.Internal

NFData ProjectionLikenessMissing Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Bounded ProjectionLikenessMissing Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Enum ProjectionLikenessMissing Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Generic ProjectionLikenessMissing Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep ProjectionLikenessMissing 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep ProjectionLikenessMissing = D1 ('MetaData "ProjectionLikenessMissing" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "MaybeProjection" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "NeverProjection" 'PrefixI 'False) (U1 :: Type -> Type))
Show ProjectionLikenessMissing Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep ProjectionLikenessMissing Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep ProjectionLikenessMissing = D1 ('MetaData "ProjectionLikenessMissing" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "MaybeProjection" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "NeverProjection" 'PrefixI 'False) (U1 :: Type -> Type))

class Monad m => ReadTCState (m :: Type -> Type) where Source #

Minimal complete definition

Nothing

Methods

getTCState :: m TCState Source #

default getTCState :: forall (t :: (Type -> Type) -> Type -> Type) (n :: Type -> Type). (MonadTrans t, ReadTCState n, t n ~ m) => m TCState Source #

locallyTCState :: Lens' TCState a -> (a -> a) -> m b -> m b Source #

default locallyTCState :: forall (t :: (Type -> Type) -> Type -> Type) (n :: Type -> Type) a b. (MonadTransControl t, ReadTCState n, t n ~ m) => Lens' TCState a -> (a -> a) -> m b -> m b Source #

withTCState :: (TCState -> TCState) -> m a -> m a Source #

Instances

Instances details
ReadTCState IM Source # 
Instance details

Defined in Agda.Interaction.Monad

Methods

getTCState :: IM TCState Source #

locallyTCState :: Lens' TCState a -> (a -> a) -> IM b -> IM b Source #

withTCState :: (TCState -> TCState) -> IM a -> IM a Source #

ReadTCState TerM Source # 
Instance details

Defined in Agda.Termination.Monad

ReadTCState ReduceM Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

ReadTCState NLM Source # 
Instance details

Defined in Agda.TypeChecking.Rewriting.NonLinMatch

Methods

getTCState :: NLM TCState Source #

locallyTCState :: Lens' TCState a -> (a -> a) -> NLM b -> NLM b Source #

withTCState :: (TCState -> TCState) -> NLM a -> NLM a Source #

ReadTCState m => ReadTCState (PureConversionT m) Source # 
Instance details

Defined in Agda.TypeChecking.Conversion.Pure

ReadTCState m => ReadTCState (BlockT m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

getTCState :: BlockT m TCState Source #

locallyTCState :: Lens' TCState a -> (a -> a) -> BlockT m b -> BlockT m b Source #

withTCState :: (TCState -> TCState) -> BlockT m a -> BlockT m a Source #

MonadIO m => ReadTCState (TCMT m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

getTCState :: TCMT m TCState Source #

locallyTCState :: Lens' TCState a -> (a -> a) -> TCMT m b -> TCMT m b Source #

withTCState :: (TCState -> TCState) -> TCMT m a -> TCMT m a Source #

ReadTCState m => ReadTCState (NamesT m) Source # 
Instance details

Defined in Agda.TypeChecking.Names

Methods

getTCState :: NamesT m TCState Source #

locallyTCState :: Lens' TCState a -> (a -> a) -> NamesT m b -> NamesT m b Source #

withTCState :: (TCState -> TCState) -> NamesT m a -> NamesT m a Source #

ReadTCState m => ReadTCState (ListT m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

getTCState :: ListT m TCState Source #

locallyTCState :: Lens' TCState a -> (a -> a) -> ListT m b -> ListT m b Source #

withTCState :: (TCState -> TCState) -> ListT m a -> ListT m a Source #

ReadTCState m => ReadTCState (ChangeT m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

ReadTCState m => ReadTCState (MaybeT m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

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 # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

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 # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

ReadTCState m => ReadTCState (ReaderT r m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

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 # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

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 # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

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 RecordData Source #

Constructors

RecordData 

Fields

  • _recPars :: Nat

    Number of parameters.

  • _recClause :: Maybe Clause

    Was this record type created by a module application? If yes, the clause is its definition (linking back to the original record type).

  • _recConHead :: ConHead

    Constructor name and fields.

  • _recNamedCon :: Bool

    Does this record have a constructor?

  • _recFields :: [Dom QName]

    The record field names.

  • _recTel :: Telescope

    The record field telescope. (Includes record parameters.) Note: TelV recTel _ == telView' recConType. Thus, recTel is redundant.

  • _recMutual :: Maybe [QName]

    Mutually recursive functions, datas and records. Does include this record. Empty if not recursive. Nothing if not yet computed (by positivity checker).

  • _recEtaEquality' :: EtaEquality

    Eta-expand at this record type? False for unguarded recursive records and coinductive records unless the user specifies otherwise.

  • _recPatternMatching :: PatternOrCopattern

    In case eta-equality is off, do we allow pattern matching on the constructor or construction by copattern matching? Having both loses subject reduction, see issue #4560. After positivity checking, this field is obsolete, part of EtaEquality.

  • _recInduction :: Maybe Induction

    Inductive or CoInductive? Matters only for recursive records. Nothing means that the user did not specify it, which is an error for recursive records.

  • _recTerminates :: Maybe Bool

    'Just True' means that unfolding of the recursive record terminates, 'Just False' means that we have no evidence for termination, and Nothing means we have not run the termination checker yet.

  • _recAbstr :: IsAbstract
     
  • _recComp :: CompKit
     

Instances

Instances details
Pretty RecordData Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

NFData RecordData Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: RecordData -> () #

Generic RecordData Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep RecordData 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep RecordData = D1 ('MetaData "RecordData" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "RecordData" 'PrefixI 'True) (((S1 ('MetaSel ('Just "_recPars") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Nat) :*: (S1 ('MetaSel ('Just "_recClause") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe Clause)) :*: S1 ('MetaSel ('Just "_recConHead") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ConHead))) :*: (S1 ('MetaSel ('Just "_recNamedCon") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool) :*: (S1 ('MetaSel ('Just "_recFields") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Dom QName]) :*: S1 ('MetaSel ('Just "_recTel") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Telescope)))) :*: ((S1 ('MetaSel ('Just "_recMutual") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe [QName])) :*: (S1 ('MetaSel ('Just "_recEtaEquality'") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 EtaEquality) :*: S1 ('MetaSel ('Just "_recPatternMatching") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 PatternOrCopattern))) :*: ((S1 ('MetaSel ('Just "_recInduction") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe Induction)) :*: S1 ('MetaSel ('Just "_recTerminates") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe Bool))) :*: (S1 ('MetaSel ('Just "_recAbstr") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 IsAbstract) :*: S1 ('MetaSel ('Just "_recComp") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 CompKit))))))
Show RecordData Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep RecordData Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep RecordData = D1 ('MetaData "RecordData" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "RecordData" 'PrefixI 'True) (((S1 ('MetaSel ('Just "_recPars") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Nat) :*: (S1 ('MetaSel ('Just "_recClause") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe Clause)) :*: S1 ('MetaSel ('Just "_recConHead") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ConHead))) :*: (S1 ('MetaSel ('Just "_recNamedCon") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool) :*: (S1 ('MetaSel ('Just "_recFields") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Dom QName]) :*: S1 ('MetaSel ('Just "_recTel") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Telescope)))) :*: ((S1 ('MetaSel ('Just "_recMutual") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe [QName])) :*: (S1 ('MetaSel ('Just "_recEtaEquality'") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 EtaEquality) :*: S1 ('MetaSel ('Just "_recPatternMatching") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 PatternOrCopattern))) :*: ((S1 ('MetaSel ('Just "_recInduction") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe Induction)) :*: S1 ('MetaSel ('Just "_recTerminates") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe Bool))) :*: (S1 ('MetaSel ('Just "_recAbstr") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 IsAbstract) :*: S1 ('MetaSel ('Just "_recComp") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 CompKit))))))

data ReduceDefs Source #

Instances

Instances details
NFData ReduceDefs Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: ReduceDefs -> () #

Generic ReduceDefs Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep ReduceDefs 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep ReduceDefs = D1 ('MetaData "ReduceDefs" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "OnlyReduceDefs" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Set QName))) :+: C1 ('MetaCons "DontReduceDefs" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Set QName))))
type Rep ReduceDefs Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep ReduceDefs = D1 ('MetaData "ReduceDefs" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "OnlyReduceDefs" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Set QName))) :+: C1 ('MetaCons "DontReduceDefs" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Set QName))))

data ReduceEnv Source #

Environment of the reduce monad.

Constructors

ReduceEnv 

Fields

  • redEnv :: TCEnv

    Read only access to environment.

  • redSt :: TCState

    Read only access to state (signature, metas...).

  • redPred :: Maybe (MetaId -> ReduceM Bool)

    An optional predicate that is used by instantiate' and instantiateFull': meta-variables are only instantiated if they satisfy this predicate.

newtype ReduceM a Source #

Constructors

ReduceM 

Fields

Instances

Instances details
HasOptions ReduceM Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

MonadReduce ReduceM Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

liftReduce :: ReduceM a -> ReduceM a Source #

MonadTCEnv ReduceM Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

ReadTCState ReduceM Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

HasBuiltins ReduceM Source # 
Instance details

Defined in Agda.TypeChecking.Reduce.Monad

MonadAddContext ReduceM Source # 
Instance details

Defined in Agda.TypeChecking.Reduce.Monad

MonadDebug ReduceM Source # 
Instance details

Defined in Agda.TypeChecking.Reduce.Monad

PureTCM ReduceM Source # 
Instance details

Defined in Agda.TypeChecking.Reduce.Monad

HasConstInfo ReduceM Source # 
Instance details

Defined in Agda.TypeChecking.Reduce.Monad

Applicative ReduceM Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

pure :: a -> ReduceM a #

(<*>) :: ReduceM (a -> b) -> ReduceM a -> ReduceM b #

liftA2 :: (a -> b -> c) -> ReduceM a -> ReduceM b -> ReduceM c #

(*>) :: ReduceM a -> ReduceM b -> ReduceM b #

(<*) :: ReduceM a -> ReduceM b -> ReduceM a #

Functor ReduceM Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

fmap :: (a -> b) -> ReduceM a -> ReduceM b #

(<$) :: a -> ReduceM b -> ReduceM a #

Monad ReduceM Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

(>>=) :: ReduceM a -> (a -> ReduceM b) -> ReduceM b #

(>>) :: ReduceM a -> ReduceM b -> ReduceM b #

return :: a -> ReduceM a #

MonadFail ReduceM Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

fail :: String -> ReduceM a #

(Semigroup a, Monoid a) => Monoid (OccM a) Source # 
Instance details

Defined in Agda.TypeChecking.Positivity

Methods

mempty :: OccM a #

mappend :: OccM a -> OccM a -> OccM a #

mconcat :: [OccM a] -> OccM a #

data Reduced no yes Source #

Instances

Instances details
Functor (Reduced no) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

fmap :: (a -> b) -> Reduced no a -> Reduced no b #

(<$) :: a -> Reduced no b -> Reduced no a #

type RemoteMetaStore = HashMap MetaId RemoteMetaVariable Source #

Used for meta-variables from other modules (and in Interfaces).

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 the miModality field. The miMetaOccursCheck and miGeneralizable fields are omitted, because the meta-variable has already been instantiated. The Range that is part of the miClosRange field and the miNameSuggestion field are omitted because instantiated meta-variables are typically not presented to users. Finally the Closure part of the miClosRange field is omitted because it can be large (at least if we ignore potential sharing).

Instances

Instances details
LensModalPolarity RemoteMetaVariable Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

LensModality RemoteMetaVariable Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

LensQuantity RemoteMetaVariable Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

LensRelevance RemoteMetaVariable Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

InstantiateFull RemoteMetaVariable Source # 
Instance details

Defined in Agda.TypeChecking.Reduce

EmbPrj RemoteMetaVariable Source # 
Instance details

Defined in Agda.TypeChecking.Serialise.Instances.Internal

NFData RemoteMetaVariable Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: RemoteMetaVariable -> () #

Generic RemoteMetaVariable Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep RemoteMetaVariable 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep RemoteMetaVariable = D1 ('MetaData "RemoteMetaVariable" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "RemoteMetaVariable" 'PrefixI 'True) (S1 ('MetaSel ('Just "rmvInstantiation") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Instantiation) :*: (S1 ('MetaSel ('Just "rmvModality") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Modality) :*: S1 ('MetaSel ('Just "rmvJudgement") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Judgement MetaId)))))
Show RemoteMetaVariable Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep RemoteMetaVariable Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep RemoteMetaVariable = D1 ('MetaData "RemoteMetaVariable" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "RemoteMetaVariable" 'PrefixI 'True) (S1 ('MetaSel ('Just "rmvInstantiation") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Instantiation) :*: (S1 ('MetaSel ('Just "rmvModality") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Modality) :*: S1 ('MetaSel ('Just "rmvJudgement") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Judgement MetaId)))))

data RewriteRule Source #

Rewrite rules can be added independently from function clauses.

Constructors

RewriteRule 

Fields

Instances

Instances details
NamesIn RewriteRule Source # 
Instance details

Defined in Agda.Syntax.Internal.Names

KillRange RewriteRule Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

KillRange RewriteRuleMap Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

PrettyTCM RewriteRule Source # 
Instance details

Defined in Agda.TypeChecking.Pretty

InstantiateFull RewriteRule Source # 
Instance details

Defined in Agda.TypeChecking.Reduce

GetMatchables RewriteRule Source # 
Instance details

Defined in Agda.TypeChecking.Rewriting.NonLinPattern

EmbPrj RewriteRule Source # 
Instance details

Defined in Agda.TypeChecking.Serialise.Instances.Internal

Abstract RewriteRule Source #

tel ⊢ (Γ ⊢ lhs ↦ rhs : t) becomes tel, Γ ⊢ lhs ↦ rhs : t) we do not need to change lhs, rhs, and t since they live in Γ. See 'Abstract Clause'.

Instance details

Defined in Agda.TypeChecking.Substitute

Apply RewriteRule Source # 
Instance details

Defined in Agda.TypeChecking.Substitute

Subst RewriteRule Source # 
Instance details

Defined in Agda.TypeChecking.Substitute

Associated Types

type SubstArg RewriteRule 
Instance details

Defined in Agda.TypeChecking.Substitute

NFData RewriteRule Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: RewriteRule -> () #

Generic RewriteRule Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Show RewriteRule Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type SubstArg RewriteRule Source # 
Instance details

Defined in Agda.TypeChecking.Substitute

type Rep RewriteRule Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

data RunMetaOccursCheck Source #

Instances

Instances details
NFData RunMetaOccursCheck Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: RunMetaOccursCheck -> () #

Generic RunMetaOccursCheck Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep RunMetaOccursCheck 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep RunMetaOccursCheck = D1 ('MetaData "RunMetaOccursCheck" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "RunMetaOccursCheck" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "DontRunMetaOccursCheck" 'PrefixI 'False) (U1 :: Type -> Type))
Show RunMetaOccursCheck Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Eq RunMetaOccursCheck Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Ord RunMetaOccursCheck Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep RunMetaOccursCheck Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep RunMetaOccursCheck = D1 ('MetaData "RunMetaOccursCheck" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "RunMetaOccursCheck" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "DontRunMetaOccursCheck" 'PrefixI 'False) (U1 :: Type -> Type))

newtype Section Source #

Constructors

Section 

data SessionTCState Source #

A part of the state that grows monotonically over the whole Agda session. Never reset.

Constructors

SessionTCState 

Fields

Instances

Instances details
NFData SessionTCState Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: SessionTCState -> () #

Generic SessionTCState Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep SessionTCState 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep SessionTCState = D1 ('MetaData "SessionTCState" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "SessionTCState" 'PrefixI 'True) (S1 ('MetaSel ('Just "stSessionBenchmark") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Benchmark) :*: (S1 ('MetaSel ('Just "stSessionBackends") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Backend]) :*: S1 ('MetaSel ('Just "stSessionFileDict") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 FileDictWithBuiltins))))
type Rep SessionTCState Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep SessionTCState = D1 ('MetaData "SessionTCState" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "SessionTCState" 'PrefixI 'True) (S1 ('MetaSel ('Just "stSessionBenchmark") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Benchmark) :*: (S1 ('MetaSel ('Just "stSessionBackends") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Backend]) :*: S1 ('MetaSel ('Just "stSessionFileDict") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 FileDictWithBuiltins))))

data Signature Source #

Constructors

Sig 

Instances

Instances details
KillRange Signature Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

InstantiateFull Signature Source # 
Instance details

Defined in Agda.TypeChecking.Reduce

EmbPrj Signature Source # 
Instance details

Defined in Agda.TypeChecking.Serialise.Instances.Internal

NFData Signature Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: Signature -> () #

Generic Signature Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep Signature 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep Signature = D1 ('MetaData "Signature" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "Sig" 'PrefixI 'True) ((S1 ('MetaSel ('Just "_sigSections") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Sections) :*: S1 ('MetaSel ('Just "_sigDefinitions") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Definitions)) :*: (S1 ('MetaSel ('Just "_sigRewriteRules") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 RewriteRuleMap) :*: S1 ('MetaSel ('Just "_sigInstances") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 InstanceTable))))
Show Signature Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep Signature Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep Signature = D1 ('MetaData "Signature" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "Sig" 'PrefixI 'True) ((S1 ('MetaSel ('Just "_sigSections") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Sections) :*: S1 ('MetaSel ('Just "_sigDefinitions") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Definitions)) :*: (S1 ('MetaSel ('Just "_sigRewriteRules") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 RewriteRuleMap) :*: S1 ('MetaSel ('Just "_sigInstances") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 InstanceTable))))

data Simplification Source #

Did we encounter a simplifying reduction? In terms of CIC, that would be a iota-reduction. In terms of Agda, this is a constructor or literal pattern that matched. Just beta-reduction (substitution) or delta-reduction (unfolding of definitions) does not count as simplifying?

Instances

Instances details
Null Simplification Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

NFData Simplification Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: Simplification -> () #

Monoid Simplification Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Semigroup Simplification Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Generic Simplification Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep Simplification 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep Simplification = D1 ('MetaData "Simplification" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "YesSimplification" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "NoSimplification" 'PrefixI 'False) (U1 :: Type -> Type))
Show Simplification Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Eq Simplification Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep Simplification Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep Simplification = D1 ('MetaData "Simplification" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "YesSimplification" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "NoSimplification" 'PrefixI 'False) (U1 :: Type -> Type))

data SplitError Source #

Error when splitting a pattern variable into possible constructor patterns.

Constructors

NotADatatype (Closure Type)

Neither data type nor record.

BlockedType Blocker (Closure Type)

Type could not be sufficiently reduced.

ErasedDatatype ErasedDatatypeReason (Closure Type)

Data type, but in erased position.

CoinductiveDatatype (Closure Type)

Split on codata not allowed. UNUSED, but keep! -- | NoRecordConstructor Type -- ^ record type, but no constructor

UnificationStuck 

Fields

CosplitCatchall

Copattern split with a catchall

CosplitNoTarget

We do not know the target type of the clause.

CosplitNoRecordType (Closure Type)

Target type is not a record type.

CannotCreateMissingClause QName (Telescope, [NamedArg DeBruijnPattern]) Doc (Closure (Abs Type)) 
GenericSplitError String 

Instances

Instances details
PrettyTCM SplitError Source # 
Instance details

Defined in Agda.TypeChecking.Errors

NFData SplitError Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: SplitError -> () #

Generic SplitError Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep SplitError 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep SplitError = D1 ('MetaData "SplitError" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (((C1 ('MetaCons "NotADatatype" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Closure Type))) :+: C1 ('MetaCons "BlockedType" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Blocker) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Closure Type)))) :+: (C1 ('MetaCons "ErasedDatatype" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ErasedDatatypeReason) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Closure Type))) :+: (C1 ('MetaCons "CoinductiveDatatype" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Closure Type))) :+: C1 ('MetaCons "UnificationStuck" 'PrefixI 'True) ((S1 ('MetaSel ('Just "cantSplitBlocker") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe Blocker)) :*: (S1 ('MetaSel ('Just "cantSplitConName") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Just "cantSplitTel") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Telescope))) :*: (S1 ('MetaSel ('Just "cantSplitConIdx") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Args) :*: (S1 ('MetaSel ('Just "cantSplitGivenIdx") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Args) :*: S1 ('MetaSel ('Just "cantSplitFailures") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [UnificationFailure]))))))) :+: ((C1 ('MetaCons "CosplitCatchall" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "CosplitNoTarget" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "CosplitNoRecordType" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Closure Type))) :+: (C1 ('MetaCons "CannotCreateMissingClause" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Telescope, [NamedArg DeBruijnPattern]))) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Doc) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Closure (Abs Type))))) :+: C1 ('MetaCons "GenericSplitError" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String))))))
Show SplitError Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep SplitError Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep SplitError = D1 ('MetaData "SplitError" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (((C1 ('MetaCons "NotADatatype" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Closure Type))) :+: C1 ('MetaCons "BlockedType" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Blocker) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Closure Type)))) :+: (C1 ('MetaCons "ErasedDatatype" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ErasedDatatypeReason) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Closure Type))) :+: (C1 ('MetaCons "CoinductiveDatatype" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Closure Type))) :+: C1 ('MetaCons "UnificationStuck" 'PrefixI 'True) ((S1 ('MetaSel ('Just "cantSplitBlocker") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe Blocker)) :*: (S1 ('MetaSel ('Just "cantSplitConName") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Just "cantSplitTel") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Telescope))) :*: (S1 ('MetaSel ('Just "cantSplitConIdx") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Args) :*: (S1 ('MetaSel ('Just "cantSplitGivenIdx") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Args) :*: S1 ('MetaSel ('Just "cantSplitFailures") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [UnificationFailure]))))))) :+: ((C1 ('MetaCons "CosplitCatchall" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "CosplitNoTarget" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "CosplitNoRecordType" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Closure Type))) :+: (C1 ('MetaCons "CannotCreateMissingClause" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Telescope, [NamedArg DeBruijnPattern]))) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Doc) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Closure (Abs Type))))) :+: C1 ('MetaCons "GenericSplitError" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String))))))

data System Source #

An alternative representation of partial elements in a telescope: Γ ⊢ λ Δ. [φ₁ u₁, ... , φₙ uₙ] : Δ → PartialP (∨_ᵢ φᵢ) T see cubicaltt paper (however we do not store the type T).

Constructors

System 

Fields

  • systemTel :: Telescope

    the telescope Δ, binding vars for the clauses, Γ ⊢ Δ

  • systemClauses :: [(Face, Term)]

    a system [φ₁ u₁, ... , φₙ uₙ] where Γ, Δ ⊢ φᵢ and Γ, Δ, φᵢ ⊢ uᵢ

Instances

Instances details
NamesIn System Source # 
Instance details

Defined in Agda.Syntax.Internal.Names

Methods

namesAndMetasIn' :: Monoid m => (Either QName MetaId -> m) -> System -> m Source #

KillRange System Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

InstantiateFull System Source # 
Instance details

Defined in Agda.TypeChecking.Reduce

EmbPrj System Source # 
Instance details

Defined in Agda.TypeChecking.Serialise.Instances.Internal

Abstract System Source # 
Instance details

Defined in Agda.TypeChecking.Substitute

Apply System Source # 
Instance details

Defined in Agda.TypeChecking.Substitute

NFData System Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: System -> () #

Generic System Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep System 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep System = D1 ('MetaData "System" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "System" 'PrefixI 'True) (S1 ('MetaSel ('Just "systemTel") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Telescope) :*: S1 ('MetaSel ('Just "systemClauses") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [(Face, Term)])))

Methods

from :: System -> Rep System x #

to :: Rep System x -> System #

Show System Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Reify (QNamed System) Source # 
Instance details

Defined in Agda.Syntax.Translation.InternalToAbstract

Associated Types

type ReifiesTo (QNamed System) 
Instance details

Defined in Agda.Syntax.Translation.InternalToAbstract

type Rep System Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep System = D1 ('MetaData "System" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "System" 'PrefixI 'True) (S1 ('MetaSel ('Just "systemTel") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Telescope) :*: S1 ('MetaSel ('Just "systemClauses") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [(Face, Term)])))
type ReifiesTo (QNamed System) Source # 
Instance details

Defined in Agda.Syntax.Translation.InternalToAbstract

data TCEnv Source #

Constructors

TCEnv 

Fields

Instances

Instances details
LensIsAbstract TCEnv Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

LensIsOpaque TCEnv Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

LensTCEnv TCEnv Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

NFData TCEnv Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: TCEnv -> () #

Generic TCEnv Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep TCEnv 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep TCEnv = D1 ('MetaData "TCEnv" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "TCEnv" 'PrefixI 'True) (((((S1 ('MetaSel ('Just "envContext") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Context) :*: (S1 ('MetaSel ('Just "envLetBindings") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 LetBindings) :*: S1 ('MetaSel ('Just "envCurrentModule") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ModuleName))) :*: (S1 ('MetaSel ('Just "envCurrentPath") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe FileId)) :*: (S1 ('MetaSel ('Just "envAnonymousModules") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [(ModuleName, Nat)]) :*: S1 ('MetaSel ('Just "envImportPath") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [TopLevelModuleName])))) :*: ((S1 ('MetaSel ('Just "envMutualBlock") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe MutualId)) :*: (S1 ('MetaSel ('Just "envTerminationCheck") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (TerminationCheck ())) :*: S1 ('MetaSel ('Just "envCoverageCheck") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 CoverageCheck))) :*: ((S1 ('MetaSel ('Just "envMakeCase") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool) :*: S1 ('MetaSel ('Just "envSolvingConstraints") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool)) :*: (S1 ('MetaSel ('Just "envCheckingWhere") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 WhereClause_) :*: S1 ('MetaSel ('Just "envWorkingOnTypes") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool))))) :*: (((S1 ('MetaSel ('Just "envAssignMetas") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool) :*: (S1 ('MetaSel ('Just "envActiveProblems") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Set ProblemId)) :*: S1 ('MetaSel ('Just "envUnquoteProblem") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe ProblemId)))) :*: ((S1 ('MetaSel ('Just "envAbstractMode") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 AbstractMode) :*: S1 ('MetaSel ('Just "envRelevance") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Relevance)) :*: (S1 ('MetaSel ('Just "envQuantity") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Quantity) :*: S1 ('MetaSel ('Just "envHardCompileTimeMode") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool)))) :*: ((S1 ('MetaSel ('Just "envSplitOnStrict") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool) :*: (S1 ('MetaSel ('Just "envDisplayFormsEnabled") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool) :*: S1 ('MetaSel ('Just "envFoldLetBindings") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool))) :*: ((S1 ('MetaSel ('Just "envRange") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range) :*: S1 ('MetaSel ('Just "envHighlightingRange") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range)) :*: (S1 ('MetaSel ('Just "envClause") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 IPClause) :*: S1 ('MetaSel ('Just "envCall") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe (Closure Call)))))))) :*: ((((S1 ('MetaSel ('Just "envHighlightingLevel") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 HighlightingLevel) :*: (S1 ('MetaSel ('Just "envHighlightingMethod") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 HighlightingMethod) :*: S1 ('MetaSel ('Just "envExpandLast") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ExpandHidden))) :*: ((S1 ('MetaSel ('Just "envAppDef") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe QName)) :*: S1 ('MetaSel ('Just "envSimplification") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Simplification)) :*: (S1 ('MetaSel ('Just "envAllowedReductions") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 AllowedReductions) :*: S1 ('MetaSel ('Just "envReduceDefs") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ReduceDefs)))) :*: ((S1 ('MetaSel ('Just "envReconstructed") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool) :*: (S1 ('MetaSel ('Just "envInjectivityDepth") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Int) :*: S1 ('MetaSel ('Just "envCompareBlocked") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool))) :*: ((S1 ('MetaSel ('Just "envPrintDomainFreePi") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool) :*: S1 ('MetaSel ('Just "envPrintMetasBare") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool)) :*: (S1 ('MetaSel ('Just "envInsideDotPattern") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool) :*: S1 ('MetaSel ('Just "envUnquoteFlags") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 UnquoteFlags))))) :*: (((S1 ('MetaSel ('Just "envInstanceDepth") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Int) :*: (S1 ('MetaSel ('Just "envIsDebugPrinting") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool) :*: S1 ('MetaSel ('Just "envPrintingPatternLambdas") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [QName]))) :*: ((S1 ('MetaSel ('Just "envCallByNeed") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool) :*: S1 ('MetaSel ('Just "envCurrentCheckpoint") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 CheckpointId)) :*: (S1 ('MetaSel ('Just "envCheckpoints") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Map CheckpointId Substitution)) :*: S1 ('MetaSel ('Just "envGeneralizeMetas") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 DoGeneralize)))) :*: ((S1 ('MetaSel ('Just "envGeneralizedVars") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Map QName GeneralizedValue)) :*: (S1 ('MetaSel ('Just "envActiveBackendName") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe BackendName)) :*: S1 ('MetaSel ('Just "envConflComputingOverlap") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool))) :*: ((S1 ('MetaSel ('Just "envCurrentlyElaborating") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool) :*: S1 ('MetaSel ('Just "envSyntacticEqualityFuel") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (Maybe Int))) :*: (S1 ('MetaSel ('Just "envCurrentOpaqueId") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (Maybe OpaqueId)) :*: S1 ('MetaSel ('Just "envTermCheckReducing") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool))))))))

Methods

from :: TCEnv -> Rep TCEnv x #

to :: Rep TCEnv x -> TCEnv #

type Rep TCEnv Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep TCEnv = D1 ('MetaData "TCEnv" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "TCEnv" 'PrefixI 'True) (((((S1 ('MetaSel ('Just "envContext") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Context) :*: (S1 ('MetaSel ('Just "envLetBindings") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 LetBindings) :*: S1 ('MetaSel ('Just "envCurrentModule") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ModuleName))) :*: (S1 ('MetaSel ('Just "envCurrentPath") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe FileId)) :*: (S1 ('MetaSel ('Just "envAnonymousModules") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [(ModuleName, Nat)]) :*: S1 ('MetaSel ('Just "envImportPath") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [TopLevelModuleName])))) :*: ((S1 ('MetaSel ('Just "envMutualBlock") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe MutualId)) :*: (S1 ('MetaSel ('Just "envTerminationCheck") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (TerminationCheck ())) :*: S1 ('MetaSel ('Just "envCoverageCheck") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 CoverageCheck))) :*: ((S1 ('MetaSel ('Just "envMakeCase") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool) :*: S1 ('MetaSel ('Just "envSolvingConstraints") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool)) :*: (S1 ('MetaSel ('Just "envCheckingWhere") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 WhereClause_) :*: S1 ('MetaSel ('Just "envWorkingOnTypes") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool))))) :*: (((S1 ('MetaSel ('Just "envAssignMetas") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool) :*: (S1 ('MetaSel ('Just "envActiveProblems") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Set ProblemId)) :*: S1 ('MetaSel ('Just "envUnquoteProblem") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe ProblemId)))) :*: ((S1 ('MetaSel ('Just "envAbstractMode") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 AbstractMode) :*: S1 ('MetaSel ('Just "envRelevance") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Relevance)) :*: (S1 ('MetaSel ('Just "envQuantity") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Quantity) :*: S1 ('MetaSel ('Just "envHardCompileTimeMode") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool)))) :*: ((S1 ('MetaSel ('Just "envSplitOnStrict") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool) :*: (S1 ('MetaSel ('Just "envDisplayFormsEnabled") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool) :*: S1 ('MetaSel ('Just "envFoldLetBindings") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool))) :*: ((S1 ('MetaSel ('Just "envRange") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range) :*: S1 ('MetaSel ('Just "envHighlightingRange") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range)) :*: (S1 ('MetaSel ('Just "envClause") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 IPClause) :*: S1 ('MetaSel ('Just "envCall") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe (Closure Call)))))))) :*: ((((S1 ('MetaSel ('Just "envHighlightingLevel") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 HighlightingLevel) :*: (S1 ('MetaSel ('Just "envHighlightingMethod") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 HighlightingMethod) :*: S1 ('MetaSel ('Just "envExpandLast") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ExpandHidden))) :*: ((S1 ('MetaSel ('Just "envAppDef") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe QName)) :*: S1 ('MetaSel ('Just "envSimplification") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Simplification)) :*: (S1 ('MetaSel ('Just "envAllowedReductions") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 AllowedReductions) :*: S1 ('MetaSel ('Just "envReduceDefs") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ReduceDefs)))) :*: ((S1 ('MetaSel ('Just "envReconstructed") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool) :*: (S1 ('MetaSel ('Just "envInjectivityDepth") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Int) :*: S1 ('MetaSel ('Just "envCompareBlocked") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool))) :*: ((S1 ('MetaSel ('Just "envPrintDomainFreePi") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool) :*: S1 ('MetaSel ('Just "envPrintMetasBare") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool)) :*: (S1 ('MetaSel ('Just "envInsideDotPattern") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool) :*: S1 ('MetaSel ('Just "envUnquoteFlags") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 UnquoteFlags))))) :*: (((S1 ('MetaSel ('Just "envInstanceDepth") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 Int) :*: (S1 ('MetaSel ('Just "envIsDebugPrinting") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool) :*: S1 ('MetaSel ('Just "envPrintingPatternLambdas") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [QName]))) :*: ((S1 ('MetaSel ('Just "envCallByNeed") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool) :*: S1 ('MetaSel ('Just "envCurrentCheckpoint") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 CheckpointId)) :*: (S1 ('MetaSel ('Just "envCheckpoints") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Map CheckpointId Substitution)) :*: S1 ('MetaSel ('Just "envGeneralizeMetas") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 DoGeneralize)))) :*: ((S1 ('MetaSel ('Just "envGeneralizedVars") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Map QName GeneralizedValue)) :*: (S1 ('MetaSel ('Just "envActiveBackendName") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe BackendName)) :*: S1 ('MetaSel ('Just "envConflComputingOverlap") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool))) :*: ((S1 ('MetaSel ('Just "envCurrentlyElaborating") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool) :*: S1 ('MetaSel ('Just "envSyntacticEqualityFuel") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (Maybe Int))) :*: (S1 ('MetaSel ('Just "envCurrentOpaqueId") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 (Maybe OpaqueId)) :*: S1 ('MetaSel ('Just "envTermCheckReducing") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool))))))))

data TCErr Source #

Type-checking errors.

Constructors

TypeError 

Fields

ParserError ParseError

Error raised by the Happy parser.

GenericException String

Unspecific error without Range.

IOException (Maybe TCState) Range IOException

The first argument is the state in which the error was raised.

PatternErr Blocker

The exception which is usually caught. Raised for pattern violations during unification (assignV) but also in other situations where we want to backtrack. Contains an unblocker to control when the computation should be retried.

Instances

Instances details
EncodeTCM DisplayInfo Source # 
Instance details

Defined in Agda.Interaction.JSONTop

EncodeTCM Info_Error Source # 
Instance details

Defined in Agda.Interaction.JSONTop

EncodeTCM Response Source # 
Instance details

Defined in Agda.Interaction.JSONTop

EncodeTCM TCErr Source # 
Instance details

Defined in Agda.Interaction.JSONTop

HasRange TCErr Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

getRange :: TCErr -> Range Source #

PrettyTCM TCErr Source # 
Instance details

Defined in Agda.TypeChecking.Errors

Methods

prettyTCM :: MonadPretty m => TCErr -> m Doc Source #

NFData TCErr Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: TCErr -> () #

Exception TCErr Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Show TCErr Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

showsPrec :: Int -> TCErr -> ShowS #

show :: TCErr -> String #

showList :: [TCErr] -> ShowS #

MonadError TCErr IM Source # 
Instance details

Defined in Agda.Interaction.Monad

Methods

throwError :: TCErr -> IM a #

catchError :: IM a -> (TCErr -> IM a) -> IM a #

MonadError TCErr TerM Source # 
Instance details

Defined in Agda.Termination.Monad

Methods

throwError :: TCErr -> TerM a #

catchError :: TerM a -> (TCErr -> TerM a) -> TerM a #

Monad m => MonadError TCErr (PureConversionT m) Source # 
Instance details

Defined in Agda.TypeChecking.Conversion.Pure

(CatchIO m, MonadIO m) => MonadError TCErr (TCMT m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

throwError :: TCErr -> TCMT m a #

catchError :: TCMT m a -> (TCErr -> TCMT m a) -> TCMT m a #

EncodeTCM (OutputForm Expr Expr) Source # 
Instance details

Defined in Agda.Interaction.JSONTop

(Pretty a, Pretty b) => Pretty (OutputConstraint a b) Source # 
Instance details

Defined in Agda.Interaction.BasicOps

(Pretty a, Pretty b) => Pretty (OutputForm a b) Source # 
Instance details

Defined in Agda.Interaction.BasicOps

(ToConcrete a, ToConcrete b) => ToConcrete (OutputConstraint a b) Source # 
Instance details

Defined in Agda.Interaction.BasicOps

Associated Types

type ConOfAbs (OutputConstraint a b) 
Instance details

Defined in Agda.Interaction.BasicOps

Methods

toConcrete :: MonadToConcrete m => OutputConstraint a b -> m (ConOfAbs (OutputConstraint a b)) Source #

bindToConcrete :: MonadToConcrete m => OutputConstraint a b -> (ConOfAbs (OutputConstraint a b) -> m b0) -> m b0 Source #

(ToConcrete a, ToConcrete b) => ToConcrete (OutputForm a b) Source # 
Instance details

Defined in Agda.Interaction.BasicOps

Associated Types

type ConOfAbs (OutputForm a b) 
Instance details

Defined in Agda.Interaction.BasicOps

Methods

toConcrete :: MonadToConcrete m => OutputForm a b -> m (ConOfAbs (OutputForm a b)) Source #

bindToConcrete :: MonadToConcrete m => OutputForm a b -> (ConOfAbs (OutputForm a b) -> m b0) -> m b0 Source #

Monad m => MonadBlock (ExceptT TCErr m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

(Pretty a, Pretty b) => PrettyTCM (OutputForm a b) Source # 
Instance details

Defined in Agda.Interaction.BasicOps

Methods

prettyTCM :: MonadPretty m => OutputForm a b -> m Doc Source #

type ConOfAbs (OutputConstraint a b) Source # 
Instance details

Defined in Agda.Interaction.BasicOps

type ConOfAbs (OutputForm a b) Source # 
Instance details

Defined in Agda.Interaction.BasicOps

type TCM = TCMT IO Source #

Type checking monad.

newtype TCMT (m :: Type -> Type) a Source #

The type checking monad transformer. Adds readonly TCEnv and mutable TCState.

Constructors

TCM 

Fields

Instances

Instances details
MonadFixityError ScopeM Source # 
Instance details

Defined in Agda.Syntax.Scope.Monad

MonadBlock TCM Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

MonadStConcreteNames TCM Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

MonadConstraint TCM Source # 
Instance details

Defined in Agda.TypeChecking.Constraints

MonadAddContext TCM Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Context

Methods

addCtx :: Name -> Dom Type -> TCM a -> TCM a Source #

addLetBinding' :: Origin -> Name -> Term -> Dom Type -> TCM a -> TCM a Source #

updateContext :: Substitution -> (Context -> Context) -> TCM a -> TCM a Source #

withFreshName :: Range -> ArgName -> (Name -> TCM a) -> TCM a Source #

MonadDebug TCM Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Debug

MonadInteractionPoints TCM Source # 
Instance details

Defined in Agda.TypeChecking.Monad.MetaVars

MonadMetaSolver TCM Source # 
Instance details

Defined in Agda.TypeChecking.MetaVars

PureTCM TCM Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Pure

MonadStatistics TCM Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Statistics

Methods

modifyCounter :: String -> (Integer -> Integer) -> TCM () Source #

MonadTrace TCM Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Trace

MonadWarning TCM Source # 
Instance details

Defined in Agda.TypeChecking.Warnings

Methods

addWarning :: Bool -> TCWarning -> TCM () Source #

MonadBench TCM Source #

We store benchmark statistics in an IORef. This enables benchmarking pure computation, see Agda.Benchmarking.

Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type BenchPhase TCM 
Instance details

Defined in Agda.TypeChecking.Monad.Base

CatchImpossible TCM Source #

Like catchError, but resets the state completely before running the handler. This means it also loses changes to the stPersistentState.

The intended use is to catch internal errors during debug printing. In debug printing, we are not expecting state changes.

Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

catchImpossible :: TCM a -> (Impossible -> TCM a) -> TCM a Source #

catchImpossibleJust :: (Impossible -> Maybe b) -> TCM a -> (b -> TCM a) -> TCM a Source #

handleImpossible :: (Impossible -> TCM a) -> TCM a -> TCM a Source #

handleImpossibleJust :: (Impossible -> Maybe b) -> (b -> TCM a) -> TCM a -> TCM a Source #

MonadTrans TCMT Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

lift :: Monad m => m a -> TCMT m a #

HasFresh i => MonadFresh i TCM Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

fresh :: TCM i Source #

(CatchIO m, MonadIO m) => MonadError TCErr (TCMT m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

throwError :: TCErr -> TCMT m a #

catchError :: TCMT m a -> (TCErr -> TCMT m a) -> TCMT m a #

MonadIO m => HasOptions (TCMT m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

MonadIO m => MonadReduce (TCMT m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

liftReduce :: ReduceM a -> TCMT m a Source #

MonadIO m => MonadTCEnv (TCMT m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

askTC :: TCMT m TCEnv Source #

localTC :: (TCEnv -> TCEnv) -> TCMT m a -> TCMT m a Source #

MonadIO m => MonadTCM (TCMT m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

liftTCM :: TCM a -> TCMT m a Source #

MonadIO m => MonadTCState (TCMT m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

MonadIO m => ReadTCState (TCMT m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

getTCState :: TCMT m TCState Source #

locallyTCState :: Lens' TCState a -> (a -> a) -> TCMT m b -> TCMT m b Source #

withTCState :: (TCState -> TCState) -> TCMT m a -> TCMT m a Source #

MonadIO m => HasBuiltins (TCMT m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Builtin

ReportS (TCM Doc) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Debug

Methods

reportS :: MonadDebug m => VerboseKey -> VerboseLevel -> TCM Doc -> m () Source #

ReportS [TCM Doc] Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Debug

Methods

reportS :: MonadDebug m => VerboseKey -> VerboseLevel -> [TCM Doc] -> m () Source #

TraceS (TCM Doc) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Debug

Methods

traceS :: MonadDebug m => VerboseKey -> VerboseLevel -> TCM Doc -> m c -> m c Source #

TraceS [TCM Doc] Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Debug

Methods

traceS :: MonadDebug m => VerboseKey -> VerboseLevel -> [TCM Doc] -> m c -> m c Source #

HasConstInfo (TCMT IO) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Signature

MonadIO m => MonadFileId (TCMT m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.State

Null (TCM Doc) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Applicative m => Applicative (TCMT m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

pure :: a -> TCMT m a #

(<*>) :: TCMT m (a -> b) -> TCMT m a -> TCMT m b #

liftA2 :: (a -> b -> c) -> TCMT m a -> TCMT m b -> TCMT m c #

(*>) :: TCMT m a -> TCMT m b -> TCMT m b #

(<*) :: TCMT m a -> TCMT m b -> TCMT m a #

Functor m => Functor (TCMT m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

fmap :: (a -> b) -> TCMT m a -> TCMT m b #

(<$) :: a -> TCMT m b -> TCMT m a #

Monad m => Monad (TCMT m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

(>>=) :: TCMT m a -> (a -> TCMT m b) -> TCMT m b #

(>>) :: TCMT m a -> TCMT m b -> TCMT m b #

return :: a -> TCMT m a #

Semigroup (TCM Doc) Source #

This instance is more specific than a generic instance Semigroup a => Semigroup (TCM a).

Instance details

Defined in Agda.TypeChecking.Pretty

Methods

(<>) :: TCM Doc -> TCM Doc -> TCM Doc #

sconcat :: NonEmpty (TCM Doc) -> TCM Doc #

stimes :: Integral b => b -> TCM Doc -> TCM Doc #

(CatchIO m, MonadIO m) => MonadFail (TCMT m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

fail :: String -> TCMT m a #

MonadIO m => MonadIO (TCMT m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

liftIO :: IO a -> TCMT m a #

(MonadIO m, Null a) => Null (TCMT m a) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

empty :: TCMT m a Source #

null :: TCMT m a -> Bool Source #

(MonadIO m, Semigroup a, Monoid a) => Monoid (TCMT m a) Source #

Strict (non-shortcut) monoid.

Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

mempty :: TCMT m a #

mappend :: TCMT m a -> TCMT m a -> TCMT m a #

mconcat :: [TCMT m a] -> TCMT m a #

(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 and2M as concatenation, to shortcut conjunction in case we already have False.

Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

(<>) :: TCMT m a -> TCMT m a -> TCMT m a #

sconcat :: NonEmpty (TCMT m a) -> TCMT m a #

stimes :: Integral b => b -> TCMT m a -> TCMT m a #

(IsString a, MonadIO m) => IsString (TCMT m a) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

fromString :: String -> TCMT m a #

type BenchPhase TCM Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

data TCState Source #

Constructors

TCSt 

Fields

Instances

Instances details
LensCommandLineOptions TCState Source # 
Instance details

Defined in Agda.Interaction.Options.Lenses

LensIncludePaths TCState Source # 
Instance details

Defined in Agda.Interaction.Options.Lenses

LensPersistentVerbosity TCState Source # 
Instance details

Defined in Agda.Interaction.Options.Lenses

LensSafeMode TCState Source # 
Instance details

Defined in Agda.Interaction.Options.Lenses

LensVerbosity TCState Source # 
Instance details

Defined in Agda.Interaction.Options.Lenses

LensPragmaOptions TCState Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

NFData TCState Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: TCState -> () #

Generic TCState Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep TCState 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep TCState = D1 ('MetaData "TCState" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "TCSt" 'PrefixI 'True) (S1 ('MetaSel ('Just "stPersistentState") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 PersistentTCState) :*: (S1 ('MetaSel ('Just "stPreScopeState") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 PreScopeState) :*: S1 ('MetaSel ('Just "stPostScopeState") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 PostScopeState))))

Methods

from :: TCState -> Rep TCState x #

to :: Rep TCState x -> TCState #

Show TCState Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep TCState Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep TCState = D1 ('MetaData "TCState" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "TCSt" 'PrefixI 'True) (S1 ('MetaSel ('Just "stPersistentState") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 PersistentTCState) :*: (S1 ('MetaSel ('Just "stPreScopeState") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 PreScopeState) :*: S1 ('MetaSel ('Just "stPostScopeState") 'NoSourceUnpackedness 'SourceStrict 'DecidedStrict) (Rec0 PostScopeState))))

data TCWarning Source #

Constructors

TCWarning 

Fields

Instances

Instances details
EncodeTCM DisplayInfo Source # 
Instance details

Defined in Agda.Interaction.JSONTop

EncodeTCM Info_Error Source # 
Instance details

Defined in Agda.Interaction.JSONTop

EncodeTCM Response Source # 
Instance details

Defined in Agda.Interaction.JSONTop

EncodeTCM TCWarning Source # 
Instance details

Defined in Agda.Interaction.JSONTop

HasRange TCWarning Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

PrettyTCM TCWarning Source # 
Instance details

Defined in Agda.TypeChecking.Pretty.Warning

EmbPrj TCWarning Source # 
Instance details

Defined in Agda.TypeChecking.Serialise.Instances.Errors

NFData TCWarning Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: TCWarning -> () #

Generic TCWarning Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep TCWarning 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep TCWarning = D1 ('MetaData "TCWarning" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "TCWarning" 'PrefixI 'True) ((S1 ('MetaSel ('Just "tcWarningLocation") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 CallStack) :*: (S1 ('MetaSel ('Just "tcWarningRange") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range) :*: S1 ('MetaSel ('Just "tcWarning") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Warning))) :*: (S1 ('MetaSel ('Just "tcWarningDoc") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Doc) :*: (S1 ('MetaSel ('Just "tcWarningString") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String) :*: S1 ('MetaSel ('Just "tcWarningCached") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool)))))
Show TCWarning Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Eq TCWarning Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Ord TCWarning Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep TCWarning Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep TCWarning = D1 ('MetaData "TCWarning" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "TCWarning" 'PrefixI 'True) ((S1 ('MetaSel ('Just "tcWarningLocation") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 CallStack) :*: (S1 ('MetaSel ('Just "tcWarningRange") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range) :*: S1 ('MetaSel ('Just "tcWarning") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Warning))) :*: (S1 ('MetaSel ('Just "tcWarningDoc") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Doc) :*: (S1 ('MetaSel ('Just "tcWarningString") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String) :*: S1 ('MetaSel ('Just "tcWarningCached") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool)))))

type TempInstanceTable = (InstanceTable, Set QName) Source #

When typechecking something of the following form:

instance x : _ x = y

it's not yet known where to add x, so we add it to a list of unresolved instances and we'll deal with it later.

data TermHead Source #

Instances

Instances details
Pretty TermHead Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

KillRange TermHead Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

EmbPrj TermHead Source # 
Instance details

Defined in Agda.TypeChecking.Serialise.Instances.Internal

NFData TermHead Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: TermHead -> () #

Generic TermHead Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep TermHead 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep TermHead = D1 ('MetaData "TermHead" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) ((C1 ('MetaCons "SortHead" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "PiHead" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "ConsHead" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: (C1 ('MetaCons "VarHead" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Nat)) :+: C1 ('MetaCons "UnknownHead" 'PrefixI 'False) (U1 :: Type -> Type))))

Methods

from :: TermHead -> Rep TermHead x #

to :: Rep TermHead x -> TermHead #

Show TermHead Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Eq TermHead Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Ord TermHead Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep TermHead Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep TermHead = D1 ('MetaData "TermHead" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) ((C1 ('MetaCons "SortHead" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "PiHead" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "ConsHead" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: (C1 ('MetaCons "VarHead" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Nat)) :+: C1 ('MetaCons "UnknownHead" 'PrefixI 'False) (U1 :: Type -> Type))))

data TerminationError Source #

Information about a mutual block which did not pass the termination checker.

Constructors

TerminationError 

Fields

Instances

Instances details
NFData TerminationError Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: TerminationError -> () #

Generic TerminationError Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep TerminationError 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep TerminationError = D1 ('MetaData "TerminationError" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "TerminationError" 'PrefixI 'True) (S1 ('MetaSel ('Just "termErrFunctions") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [QName]) :*: (S1 ('MetaSel ('Just "termErrCalls") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [CallInfo]) :*: S1 ('MetaSel ('Just "termErrGuardednessHelps") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 GuardednessHelps))))
Show TerminationError Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep TerminationError Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep TerminationError = D1 ('MetaData "TerminationError" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "TerminationError" 'PrefixI 'True) (S1 ('MetaSel ('Just "termErrFunctions") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [QName]) :*: (S1 ('MetaSel ('Just "termErrCalls") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [CallInfo]) :*: S1 ('MetaSel ('Just "termErrGuardednessHelps") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 GuardednessHelps))))

data TypeCheckAction Source #

A complete log for a module will look like this:

Instances

Instances details
NFData TypeCheckAction Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: TypeCheckAction -> () #

Generic TypeCheckAction Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep TypeCheckAction Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

data TypeCheckingProblem Source #

Constructors

CheckExpr Comparison Expr Type 
CheckArgs Comparison ExpandHidden Expr [NamedArg Expr] Type Type (ArgsCheckState CheckedTarget -> TCM Term) 
CheckProjAppToKnownPrincipalArg Comparison Expr ProjOrigin (List1 QName) Expr Args Type Int Term Type PrincipalArgTypeMetas 
CheckLambda Comparison (Arg (List1 (WithHiding Name), Maybe Type)) Expr Type

(λ (xs : t₀) → e) : t This is not an instance of CheckExpr as the domain type has already been checked. For example, when checking (λ (x y : Fin _) → e) : (x : Fin n) → ? we want to postpone (λ (y : Fin n) → e) : ? where Fin n is a Type rather than an Expr.

DisambiguateConstructor ConstructorDisambiguationData (ConHead -> TCM Term)

A stuck constructor disambiguation with the bits to retry it on and the success continuation.

DoQuoteTerm Comparison Term Type

Quote the given term and check type against Term

Instances

Instances details
PrettyTCM TypeCheckingProblem Source # 
Instance details

Defined in Agda.TypeChecking.Pretty

NFData TypeCheckingProblem Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: TypeCheckingProblem -> () #

Generic TypeCheckingProblem Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep TypeCheckingProblem 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep TypeCheckingProblem = D1 ('MetaData "TypeCheckingProblem" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) ((C1 ('MetaCons "CheckExpr" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type))) :+: (C1 ('MetaCons "CheckArgs" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ExpandHidden) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr))) :*: ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [NamedArg Expr]) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (ArgsCheckState CheckedTarget -> TCM Term))))) :+: C1 ('MetaCons "CheckProjAppToKnownPrincipalArg" 'PrefixI 'False) (((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ProjOrigin) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 QName)) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr)))) :*: ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Args) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Int))) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 PrincipalArgTypeMetas))))))) :+: (C1 ('MetaCons "CheckLambda" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Arg (List1 (WithHiding Name), Maybe Type)))) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type))) :+: (C1 ('MetaCons "DisambiguateConstructor" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ConstructorDisambiguationData) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (ConHead -> TCM Term))) :+: C1 ('MetaCons "DoQuoteTerm" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type))))))
type Rep TypeCheckingProblem Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep TypeCheckingProblem = D1 ('MetaData "TypeCheckingProblem" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) ((C1 ('MetaCons "CheckExpr" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type))) :+: (C1 ('MetaCons "CheckArgs" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ExpandHidden) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr))) :*: ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [NamedArg Expr]) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (ArgsCheckState CheckedTarget -> TCM Term))))) :+: C1 ('MetaCons "CheckProjAppToKnownPrincipalArg" 'PrefixI 'False) (((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ProjOrigin) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 QName)) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr)))) :*: ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Args) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Int))) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 PrincipalArgTypeMetas))))))) :+: (C1 ('MetaCons "CheckLambda" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Arg (List1 (WithHiding Name), Maybe Type)))) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type))) :+: (C1 ('MetaCons "DisambiguateConstructor" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ConstructorDisambiguationData) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (ConHead -> TCM Term))) :+: C1 ('MetaCons "DoQuoteTerm" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type))))))

data TypeError Source #

Constructors

InternalError String 
NotImplemented String 
NotSupported String 
CompilationError String 
SyntaxError String

Essential syntax error thrown after successful parsing. Description in String.

OptionError OptionError

Error thrown by the option parser.

NicifierError DeclarationException'

Error thrown in the nicifier phase Definitions.

DoNotationError String

Error during unsugaring some do notation. Error message in String.

IdiomBracketError String

Error during (operator) parsing and interpreting the contents of idiom brackets. Error message in String.

NoKnownRecordWithSuchFields [Name]

The user has given a record expression with the given fields, but no record type known to type inference has all these fields. The list can be empty.

ShouldEndInApplicationOfTheDatatype Type

The target of a constructor isn't an application of its datatype. The Type records what it does target.

ConstructorPatternInWrongDatatype QName QName

constructor, datatype

CantResolveOverloadedConstructorsTargetingSameDatatype QName (List1 QName)

Datatype, constructors.

ConstructorDoesNotTargetGivenType QName Type

constructor, type

InvalidDottedExpression

.e in non-argument position.

LiteralTooBig

An integer literal that would be too costly to expand to unary.

NegativeLiteralInPattern

Negative literals are not supported in patterns.

WrongHidingInLHS

The left hand side of a function definition has a hidden argument where a non-hidden was expected.

WrongHidingInLambda Type

Expected a non-hidden function and found a hidden lambda.

WrongHidingInApplication Type

A function is applied to a hidden argument where a non-hidden was expected.

WrongHidingInProjection QName 
IllegalHidingInPostfixProjection (NamedArg Expr) 
WrongNamedArgument (NamedArg Expr) (List1 NamedName)

A function is applied to a hidden named argument it does not have. The list contains names of possible hidden arguments at this point.

WrongAnnotationInLambda

Wrong user-given (lock/tick) annotation in lambda.

WrongIrrelevanceInLambda

Wrong user-given relevance annotation in lambda.

WrongQuantityInLambda

Wrong user-given quantity annotation in lambda.

WrongCohesionInLambda

Wrong user-given cohesion annotation in lambda.

WrongPolarityInLambda

Wrong user-given polarity annotation in lambda.

QuantityMismatch Quantity Quantity

The given quantity does not correspond to the expected quantity.

HidingMismatch Hiding Hiding

The given hiding does not correspond to the expected hiding.

RelevanceMismatch Relevance Relevance

The given relevance does not correspond to the expected relevance.

ForcedConstructorNotInstantiated Pattern 
IllformedProjectionPatternAbstract Pattern 
IllformedProjectionPatternConcrete Pattern 
CannotEliminateWithPattern (Maybe Blocker) (NamedArg Pattern) Type 
CannotEliminateWithProjection (Arg Type) Bool QName 
WrongNumberOfConstructorArguments QName Nat Nat 
ShouldBeEmpty Type [DeBruijnPattern]

Type should be empty. The list gives possible patterns that match, but can be empty.

ShouldBeASort Type

The given type should have been a sort.

ShouldBePi Type

The given type should have been a pi.

ShouldBePath Type 
ShouldBeRecordType Type 
ShouldBeRecordPattern DeBruijnPattern 
CannotApply Expr Type

The given expression is used as a function but its type is not a function type.

InvalidTypeSort Sort

This sort is not a type expression.

SplitOnCoinductive 
SplitOnIrrelevant (Dom Type) 
SplitOnUnusableCohesion (Dom Type) 
SplitOnUnusablePolarity (Dom Type) 
SplitOnNonVariable Term Type 
SplitOnNonEtaRecord QName 
SplitOnAbstract QName 
SplitOnUnchecked QName 
SplitOnPartial (Dom Type) 
SplitInProp DataOrRecordE 
DefinitionIsIrrelevant QName 
DefinitionIsErased QName 
ProjectionIsIrrelevant QName 
VariableIsIrrelevant Name 
VariableIsErased Name 
VariableIsOfUnusableCohesion Name Cohesion 
LambdaIsErased 
RecordIsErased 
InvalidModalTelescopeUse Term Modality Modality Definition 
VariableIsOfUnusablePolarity Name PolarityModality 
UnequalLevel Comparison Level Level 
UnequalTerms Comparison Term Term CompareAs 
UnequalRelevance Comparison Term Term

The two function types have different relevance.

UnequalQuantity Comparison Term Term

The two function types have different relevance.

UnequalCohesion Comparison Term Term

The two function types have different cohesion.

UnequalPolarity Comparison Term Term

The two function types have different polarity.

UnequalFiniteness Comparison Term Term

One of the function types has a finite domain (i.e. is a Partial@) and the other isonot.

UnequalHiding Term Term

The two function types have different hiding.

UnequalSorts Sort Sort 
NotLeqSort Sort Sort 
MetaCannotDependOn MetaId Term Nat

The arguments are the meta variable, the proposed solution, and the parameter that it wants to depend on.

MetaIrrelevantSolution MetaId Term

When solving MetaId ... := Term, part of the Term is invalid as it was created in an irrelevant context.

MetaErasedSolution MetaId Term

When solving MetaId ... := Term, part of the Term is invalid as it was created in an erased context.

GenericError String 
GenericDocError Doc 
SortOfSplitVarError (Maybe Blocker) Doc

the meta is what we might be blocked on.

WrongSharpArity QName 
BuiltinMustBeConstructor BuiltinId Expr 
BuiltinMustBeData BuiltinId Int 
BuiltinMustBeDef BuiltinId 
BuiltinMustBeFunction BuiltinId 
BuiltinMustBePostulate BuiltinId 
NoSuchBuiltinName String 
InvalidBuiltin String 
DuplicateBuiltinBinding BuiltinId Term Term 
NoBindingForBuiltin BuiltinId 
NoBindingForPrimitive PrimitiveId 
NoSuchPrimitiveFunction String 
DuplicatePrimitiveBinding PrimitiveId QName QName 
WrongArgInfoForPrimitive PrimitiveId ArgInfo ArgInfo 
ShadowedModule Name (List1 ModuleName) 
BuiltinInParameterisedModule BuiltinId 
IllegalDeclarationInDataDefinition (List1 Declaration)

The declaration list comes from a single NiceDeclaration.

IllegalLetInTelescope TypedBinding 
IllegalPatternInTelescope Binder 
AbsentRHSRequiresAbsurdPattern 
TooManyFields QName [Name] (List1 Name)

Record type, fields not supplied by user, possibly non-fields but supplied.

DuplicateFields (List1 Name) 
DuplicateConstructors (List1 Name) 
DuplicateOverlapPragma QName OverlapMode OverlapMode 
WithOnFreeVariable Expr Term 
UnexpectedWithPatterns (List1 Pattern) 
WithClausePatternMismatch Pattern (NamedArg DeBruijnPattern) 
IllTypedPatternAfterWithAbstraction Pattern 
TooFewPatternsInWithClause 
TooManyPatternsInWithClause 
PathAbstractionFailed (Abs Type) 
FieldOutsideRecord 
ModuleArityMismatch ModuleName Telescope (Either (List1 (NamedArg Expr)) Args) 
GeneralizeCyclicDependency 
ReferencesFutureVariables Term (List1 Int) (Arg Term) Int

The first term references the given list of variables, which are in "the future" with respect to the given lock (and its leftmost variable)

DoesNotMentionTicks Term Type (Arg Term)

Arguments: later term, its type, lock term. The lock term does not mention any @lock variables.

MismatchedProjectionsError QName QName 
AttributeKindNotEnabled String String String 
InvalidProjectionParameter (NamedArg Expr) 
TacticAttributeNotAllowed 
CannotRewriteByNonEquation Type 
MacroResultTypeMismatch Type 
NamedWhereModuleInRefinedContext [Term] [String]

The lists should have the same length. TODO: enforce this by construction.

ComatchingDisabledForRecord QName 
IncorrectTypeForRewriteRelation Term IncorrectTypeForRewriteRelationReason 
CannotGenerateHCompClause Type

Cannot generate hcomp clause because type is not fibrant.

CannotGenerateTransportClause QName (Closure (Abs Type))

Cannot generate transport clause because type is not fibrant.

CubicalNotErasure QName

Name was defined for --cubical and for use in --erased-cubical the option --erasure is needed.

CubicalPrimitiveNotFullyApplied QName 
ExpectedIntervalLiteral Expr

Expected an interval literal (0 or 1) but found Expr.

FaceConstraintDisjunction 
FaceConstraintUnsatisfiable 
PatternInPathLambda

Attempt to pattern match in an abstraction of interval type.

PatternInSystem

Attempt to pattern or copattern match in a system. Data errors

UnexpectedParameter LamBinding 
NoParameterOfName ArgName 
UnexpectedModalityAnnotationInParameter LamBinding 
ExpectedBindingForParameter (Dom Type) (Abs Type) 
UnexpectedTypeSignatureForParameter (List1 (NamedArg Binder)) 
SortDoesNotAdmitDataDefinitions QName Sort 
SortCannotDependOnItsIndex QName Type 
UnusableAtModality WhyCheckModality Modality Term 
SplitError SplitError 
ImpossibleConstructor QName NegativeUnification 
DatatypeIndexPolarity

An index of a data type has a polarity different from Mixed.

RecursiveRecordNeedsInductivity QName

A record type inferred as recursive needs a manual declaration whether it should be inductively or coinductively.

CannotSolveSizeConstraints (List1 (ProblemConstraint, HypSizeConstraint)) Doc

The list of constraints is given redundantly as pairs of ProblemConstraint (original constraint) and HypSizeConstraint (form with size assumptions in context spelled out). The Doc is some extra reason for why solving failed.

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 SizeUniv, which is not allowed.

PostulatedSizeInModule 
LibraryError LibErrors

Collected errors when processing the .agda-lib file.

LibTooFarDown TopLevelModuleName AgdaLibFile

The .agda-lib file for the given module is not on the right level.

SolvedButOpenHoles

Some interaction points (holes) have not been filled by user. There are not UnsolvedMetas since unification solved them. This is an error, since interaction points are never filled without user interaction.

CyclicModuleDependency (List2 TopLevelModuleName)

The cycle starts and ends with the same module.

FileNotFound TopLevelModuleName [AbsolutePath]

The list can be empty.

OverlappingProjects AbsolutePath TopLevelModuleName TopLevelModuleName 
AmbiguousTopLevelModuleName TopLevelModuleName (List2 AbsolutePath)

The given module has at least 2 file locations.

ModuleNameUnexpected TopLevelModuleName TopLevelModuleName

Found module name, expected module name.

ModuleNameDoesntMatchFileName TopLevelModuleName [AbsolutePath]

The list can be empty.

ModuleDefinedInOtherFile TopLevelModuleName AbsolutePath AbsolutePath

Module name, file from which it was loaded, file which the include path says contains the module.

InvalidFileName AbsolutePath InvalidFileNameReason

The file name does not correspond to a module name.

ModuleNameHashCollision RawTopLevelModuleName (Maybe RawTopLevelModuleName) 
BothWithAndRHS 
AbstractConstructorNotInScope QName 
CopatternHeadNotProjection QName 
NotAllowedInDotPatterns NotAllowedInDotPatterns 
NotInScope QName 
NoSuchModule QName 
AmbiguousName QName AmbiguousNameReason 
AmbiguousModule QName (List1 ModuleName) 
AmbiguousField Name (List2 ModuleName) 
AmbiguousConstructor QName (List2 QName)

The list contains all interpretations of the name.

ClashingDefinition QName QName (Maybe NiceDeclaration) 
ClashingModule ModuleName ModuleName 
DefinitionInDifferentModule QName

The given data/record definition rests in a different module than its signature.

DuplicateImports QName (List1 ImportedName) 
InvalidPattern Pattern 
InvalidPun ConstructorOrPatternSynonym QName

Expected the identifier to be a variable, not a constructor or pattern synonym.

RepeatedNamesInImportDirective (List1 (List2 ImportedName))

Some names are bound several times by an import/open directive.

RepeatedVariablesInPattern (List1 Name) 
GeneralizeNotSupportedHere QName 
GeneralizedVarInLetOpenedModule QName 
MultipleFixityDecls (List1 (Name, Pair Fixity')) 
MultiplePolarityPragmas (List1 Name) 
ExplicitPolarityVsPragma QName 
ConstructorNameOfNonRecord ResolvedName 
CannotQuote CannotQuote 
CannotQuoteTerm CannotQuoteTerm 
DeclarationsAfterTopLevelModule 
IllegalDeclarationBeforeTopLevelModule 
MissingTypeSignature MissingTypeSignatureInfo 
NotAnExpression Expr 
NotAValidLetBinding (Maybe NotAValidLetBinding) 
NotAValidLetExpression NotAValidLetExpression 
NotValidBeforeField NiceDeclaration 
OpenEverythingInRecordWhere 
PrivateRecordField 
QualifiedLocalModule 
AsPatternInPatternSynonym 
DotPatternInPatternSynonym 
BadArgumentsToPatternSynonym AmbiguousQName 
TooFewArgumentsToPatternSynonym AmbiguousQName 
CannotResolveAmbiguousPatternSynonym (List1 (QName, PatternSynDefn)) 
IllegalInstanceVariableInPatternSynonym Name

This variable is bound in the lhs of the pattern synonym in instance position, but not on the rhs. This is forbidden because expansion of pattern synonyms would not be faithful to availability of instances in instance search.

PatternSynonymArgumentShadows ConstructorOrPatternSynonym Name (List1 AbstractName)

A variable to be bound in the pattern synonym resolved on the rhs as name of a constructor or a pattern synonym. The resolvents are given in the list.

UnusedVariableInPatternSynonym Name

This variable is only bound on the lhs of the pattern synonym, not on the rhs.

UnboundVariablesInPatternSynonym (List1 Name)

These variables are only bound on the rhs of the pattern synonym, not on the lhs. Operator errors

NoParseForApplication (List2 Expr) 
AmbiguousParseForApplication (List2 Expr) (List1 Expr) 
NoParseForLHS LHSOrPatSyn [Pattern] Pattern

The list contains patterns that failed to be interpreted. If it is non-empty, the first entry could be printed as error hint.

AmbiguousParseForLHS LHSOrPatSyn Pattern (List2 Pattern)

Pattern and its possible interpretations.

AmbiguousProjection QName (List1 QName)

The list contains alternative interpretations of the name.

AmbiguousOverloadedProjection (List1 QName) Doc 
OperatorInformation [NotationSection] TypeError

The list of notations can be empty.

InstanceNoCandidate Type [(Term, TCErr)]

The list can be empty. Reflection errors

ExecError ExecError 
UnquoteFailed UnquoteError 
DeBruijnIndexOutOfScope Nat Telescope [Name]

The list can be empty. Language option errors

NeedOptionAllowExec 
NeedOptionCopatterns 
NeedOptionCubical Cubical String

Flavor of cubical needed for the given reason.

NeedOptionPatternMatching 
NeedOptionProp 
NeedOptionRewriting 
NeedOptionSizedTypes String

Need --sized-types for the given reason.

NeedOptionTwoLevel 
NeedOptionUniversePolymorphism 
NonFatalErrors (Set1 TCWarning) 
InstanceSearchDepthExhausted Term Type Int 
TriedToCopyConstrainedPrim QName 
InvalidInstanceHeadType Type WhyInvalidInstanceType 
InteractionError InteractionError 
BackendDoesNotSupportOnlyScopeChecking BackendName

The given backend does not support --only-scope-checking.

CubicalCompilationNotSupported Cubical

NYI: Compilation of files using the given flavor of Cubical.

CustomBackendError BackendName Doc

Used for backend-specific errors. The string is the backend name.

GHCBackendError GHCBackendError

Errors raised by the GHC backend.

JSBackendError JSBackendError

Errors raised by the JS backend.

UnknownBackend BackendName (Set BackendName)

Unknown backend requested, known ones are in the Set.

Instances

Instances details
PrettyTCM TypeError Source # 
Instance details

Defined in Agda.TypeChecking.Errors

NFData TypeError Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: TypeError -> () #

Generic TypeError Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep TypeError 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep TypeError = D1 ('MetaData "TypeError" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (((((((C1 ('MetaCons "InternalError" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String)) :+: (C1 ('MetaCons "NotImplemented" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String)) :+: C1 ('MetaCons "NotSupported" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String)))) :+: ((C1 ('MetaCons "CompilationError" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String)) :+: C1 ('MetaCons "SyntaxError" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String))) :+: (C1 ('MetaCons "OptionError" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 OptionError)) :+: C1 ('MetaCons "NicifierError" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 DeclarationException'))))) :+: (((C1 ('MetaCons "DoNotationError" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String)) :+: C1 ('MetaCons "IdiomBracketError" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String))) :+: (C1 ('MetaCons "NoKnownRecordWithSuchFields" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Name])) :+: C1 ('MetaCons "ShouldEndInApplicationOfTheDatatype" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)))) :+: ((C1 ('MetaCons "ConstructorPatternInWrongDatatype" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "CantResolveOverloadedConstructorsTargetingSameDatatype" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 QName)))) :+: (C1 ('MetaCons "ConstructorDoesNotTargetGivenType" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)) :+: C1 ('MetaCons "InvalidDottedExpression" 'PrefixI 'False) (U1 :: Type -> Type))))) :+: (((C1 ('MetaCons "LiteralTooBig" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "NegativeLiteralInPattern" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "WrongHidingInLHS" 'PrefixI 'False) (U1 :: Type -> Type))) :+: ((C1 ('MetaCons "WrongHidingInLambda" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)) :+: C1 ('MetaCons "WrongHidingInApplication" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type))) :+: (C1 ('MetaCons "WrongHidingInProjection" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "IllegalHidingInPostfixProjection" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (NamedArg Expr)))))) :+: (((C1 ('MetaCons "WrongNamedArgument" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (NamedArg Expr)) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 NamedName))) :+: C1 ('MetaCons "WrongAnnotationInLambda" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "WrongIrrelevanceInLambda" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "WrongQuantityInLambda" 'PrefixI 'False) (U1 :: Type -> Type))) :+: ((C1 ('MetaCons "WrongCohesionInLambda" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "WrongPolarityInLambda" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "QuantityMismatch" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Quantity) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Quantity)) :+: C1 ('MetaCons "HidingMismatch" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Hiding) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Hiding))))))) :+: ((((C1 ('MetaCons "RelevanceMismatch" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Relevance) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Relevance)) :+: (C1 ('MetaCons "ForcedConstructorNotInstantiated" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Pattern)) :+: C1 ('MetaCons "IllformedProjectionPatternAbstract" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Pattern)))) :+: ((C1 ('MetaCons "IllformedProjectionPatternConcrete" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Pattern)) :+: C1 ('MetaCons "CannotEliminateWithPattern" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe Blocker)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (NamedArg Pattern)) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)))) :+: (C1 ('MetaCons "CannotEliminateWithProjection" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Arg Type)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName))) :+: C1 ('MetaCons "WrongNumberOfConstructorArguments" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Nat) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Nat)))))) :+: (((C1 ('MetaCons "ShouldBeEmpty" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [DeBruijnPattern])) :+: C1 ('MetaCons "ShouldBeASort" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type))) :+: (C1 ('MetaCons "ShouldBePi" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)) :+: C1 ('MetaCons "ShouldBePath" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)))) :+: ((C1 ('MetaCons "ShouldBeRecordType" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)) :+: C1 ('MetaCons "ShouldBeRecordPattern" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 DeBruijnPattern))) :+: (C1 ('MetaCons "CannotApply" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)) :+: C1 ('MetaCons "InvalidTypeSort" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Sort)))))) :+: (((C1 ('MetaCons "SplitOnCoinductive" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "SplitOnIrrelevant" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Dom Type))) :+: C1 ('MetaCons "SplitOnUnusableCohesion" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Dom Type))))) :+: ((C1 ('MetaCons "SplitOnUnusablePolarity" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Dom Type))) :+: C1 ('MetaCons "SplitOnNonVariable" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type))) :+: (C1 ('MetaCons "SplitOnNonEtaRecord" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "SplitOnAbstract" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName))))) :+: (((C1 ('MetaCons "SplitOnUnchecked" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "SplitOnPartial" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Dom Type)))) :+: (C1 ('MetaCons "SplitInProp" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 DataOrRecordE)) :+: C1 ('MetaCons "DefinitionIsIrrelevant" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)))) :+: ((C1 ('MetaCons "DefinitionIsErased" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "ProjectionIsIrrelevant" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName))) :+: (C1 ('MetaCons "VariableIsIrrelevant" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Name)) :+: C1 ('MetaCons "VariableIsErased" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Name)))))))) :+: (((((C1 ('MetaCons "VariableIsOfUnusableCohesion" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Name) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Cohesion)) :+: (C1 ('MetaCons "LambdaIsErased" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "RecordIsErased" 'PrefixI 'False) (U1 :: Type -> Type))) :+: ((C1 ('MetaCons "InvalidModalTelescopeUse" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Modality)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Modality) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Definition))) :+: C1 ('MetaCons "VariableIsOfUnusablePolarity" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Name) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 PolarityModality))) :+: (C1 ('MetaCons "UnequalLevel" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Level) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Level))) :+: C1 ('MetaCons "UnequalTerms" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 CompareAs)))))) :+: (((C1 ('MetaCons "UnequalRelevance" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term))) :+: C1 ('MetaCons "UnequalQuantity" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)))) :+: (C1 ('MetaCons "UnequalCohesion" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term))) :+: C1 ('MetaCons "UnequalPolarity" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term))))) :+: ((C1 ('MetaCons "UnequalFiniteness" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term))) :+: C1 ('MetaCons "UnequalHiding" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term))) :+: (C1 ('MetaCons "UnequalSorts" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Sort) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Sort)) :+: C1 ('MetaCons "NotLeqSort" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Sort) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Sort)))))) :+: (((C1 ('MetaCons "MetaCannotDependOn" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 MetaId) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Nat))) :+: (C1 ('MetaCons "MetaIrrelevantSolution" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 MetaId) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)) :+: C1 ('MetaCons "MetaErasedSolution" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 MetaId) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)))) :+: ((C1 ('MetaCons "GenericError" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String)) :+: C1 ('MetaCons "GenericDocError" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Doc))) :+: (C1 ('MetaCons "SortOfSplitVarError" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe Blocker)) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Doc)) :+: C1 ('MetaCons "WrongSharpArity" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName))))) :+: (((C1 ('MetaCons "BuiltinMustBeConstructor" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 BuiltinId) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr)) :+: C1 ('MetaCons "BuiltinMustBeData" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 BuiltinId) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Int))) :+: (C1 ('MetaCons "BuiltinMustBeDef" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 BuiltinId)) :+: C1 ('MetaCons "BuiltinMustBeFunction" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 BuiltinId)))) :+: ((C1 ('MetaCons "BuiltinMustBePostulate" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 BuiltinId)) :+: C1 ('MetaCons "NoSuchBuiltinName" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String))) :+: (C1 ('MetaCons "InvalidBuiltin" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String)) :+: C1 ('MetaCons "DuplicateBuiltinBinding" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 BuiltinId) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)))))))) :+: ((((C1 ('MetaCons "NoBindingForBuiltin" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 BuiltinId)) :+: (C1 ('MetaCons "NoBindingForPrimitive" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 PrimitiveId)) :+: C1 ('MetaCons "NoSuchPrimitiveFunction" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String)))) :+: ((C1 ('MetaCons "DuplicatePrimitiveBinding" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 PrimitiveId) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName))) :+: C1 ('MetaCons "WrongArgInfoForPrimitive" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 PrimitiveId) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ArgInfo) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ArgInfo)))) :+: (C1 ('MetaCons "ShadowedModule" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Name) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 ModuleName))) :+: C1 ('MetaCons "BuiltinInParameterisedModule" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 BuiltinId))))) :+: (((C1 ('MetaCons "IllegalDeclarationInDataDefinition" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 Declaration))) :+: C1 ('MetaCons "IllegalLetInTelescope" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 TypedBinding))) :+: (C1 ('MetaCons "IllegalPatternInTelescope" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Binder)) :+: C1 ('MetaCons "AbsentRHSRequiresAbsurdPattern" 'PrefixI 'False) (U1 :: Type -> Type))) :+: ((C1 ('MetaCons "TooManyFields" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Name]) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 Name)))) :+: C1 ('MetaCons "DuplicateFields" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 Name)))) :+: (C1 ('MetaCons "DuplicateConstructors" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 Name))) :+: C1 ('MetaCons "DuplicateOverlapPragma" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 OverlapMode) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 OverlapMode))))))) :+: (((C1 ('MetaCons "WithOnFreeVariable" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)) :+: (C1 ('MetaCons "UnexpectedWithPatterns" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 Pattern))) :+: C1 ('MetaCons "WithClausePatternMismatch" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Pattern) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (NamedArg DeBruijnPattern))))) :+: ((C1 ('MetaCons "IllTypedPatternAfterWithAbstraction" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Pattern)) :+: C1 ('MetaCons "TooFewPatternsInWithClause" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "TooManyPatternsInWithClause" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "PathAbstractionFailed" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Abs Type)))))) :+: (((C1 ('MetaCons "FieldOutsideRecord" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "ModuleArityMismatch" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ModuleName) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Telescope) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Either (List1 (NamedArg Expr)) Args))))) :+: (C1 ('MetaCons "GeneralizeCyclicDependency" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "ReferencesFutureVariables" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 Int))) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Arg Term)) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Int))))) :+: ((C1 ('MetaCons "DoesNotMentionTicks" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Arg Term)))) :+: C1 ('MetaCons "MismatchedProjectionsError" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName))) :+: (C1 ('MetaCons "AttributeKindNotEnabled" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String))) :+: C1 ('MetaCons "InvalidProjectionParameter" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (NamedArg Expr)))))))))) :+: ((((((C1 ('MetaCons "TacticAttributeNotAllowed" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "CannotRewriteByNonEquation" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)) :+: C1 ('MetaCons "MacroResultTypeMismatch" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)))) :+: ((C1 ('MetaCons "NamedWhereModuleInRefinedContext" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Term]) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [String])) :+: C1 ('MetaCons "ComatchingDisabledForRecord" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName))) :+: (C1 ('MetaCons "IncorrectTypeForRewriteRelation" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 IncorrectTypeForRewriteRelationReason)) :+: C1 ('MetaCons "CannotGenerateHCompClause" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type))))) :+: (((C1 ('MetaCons "CannotGenerateTransportClause" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Closure (Abs Type)))) :+: C1 ('MetaCons "CubicalNotErasure" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName))) :+: (C1 ('MetaCons "CubicalPrimitiveNotFullyApplied" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "ExpectedIntervalLiteral" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr)))) :+: ((C1 ('MetaCons "FaceConstraintDisjunction" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "FaceConstraintUnsatisfiable" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "PatternInPathLambda" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "PatternInSystem" 'PrefixI 'False) (U1 :: Type -> Type))))) :+: (((C1 ('MetaCons "UnexpectedParameter" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 LamBinding)) :+: (C1 ('MetaCons "NoParameterOfName" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ArgName)) :+: C1 ('MetaCons "UnexpectedModalityAnnotationInParameter" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 LamBinding)))) :+: ((C1 ('MetaCons "ExpectedBindingForParameter" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Dom Type)) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Abs Type))) :+: C1 ('MetaCons "UnexpectedTypeSignatureForParameter" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 (NamedArg Binder))))) :+: (C1 ('MetaCons "SortDoesNotAdmitDataDefinitions" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Sort)) :+: C1 ('MetaCons "SortCannotDependOnItsIndex" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type))))) :+: (((C1 ('MetaCons "UnusableAtModality" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 WhyCheckModality) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Modality) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term))) :+: C1 ('MetaCons "SplitError" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 SplitError))) :+: (C1 ('MetaCons "ImpossibleConstructor" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 NegativeUnification)) :+: C1 ('MetaCons "DatatypeIndexPolarity" 'PrefixI 'False) (U1 :: Type -> Type))) :+: ((C1 ('MetaCons "RecursiveRecordNeedsInductivity" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "CannotSolveSizeConstraints" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 (ProblemConstraint, HypSizeConstraint))) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Doc))) :+: (C1 ('MetaCons "ContradictorySizeConstraint" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (ProblemConstraint, HypSizeConstraint))) :+: C1 ('MetaCons "EmptyTypeOfSizes" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term))))))) :+: ((((C1 ('MetaCons "FunctionTypeInSizeUniv" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)) :+: (C1 ('MetaCons "PostulatedSizeInModule" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "LibraryError" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 LibErrors)))) :+: ((C1 ('MetaCons "LibTooFarDown" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 TopLevelModuleName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 AgdaLibFile)) :+: C1 ('MetaCons "SolvedButOpenHoles" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "CyclicModuleDependency" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List2 TopLevelModuleName))) :+: C1 ('MetaCons "FileNotFound" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 TopLevelModuleName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [AbsolutePath]))))) :+: (((C1 ('MetaCons "OverlappingProjects" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 AbsolutePath) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 TopLevelModuleName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 TopLevelModuleName))) :+: C1 ('MetaCons "AmbiguousTopLevelModuleName" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 TopLevelModuleName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List2 AbsolutePath)))) :+: (C1 ('MetaCons "ModuleNameUnexpected" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 TopLevelModuleName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 TopLevelModuleName)) :+: C1 ('MetaCons "ModuleNameDoesntMatchFileName" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 TopLevelModuleName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [AbsolutePath])))) :+: ((C1 ('MetaCons "ModuleDefinedInOtherFile" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 TopLevelModuleName) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 AbsolutePath) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 AbsolutePath))) :+: C1 ('MetaCons "InvalidFileName" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 AbsolutePath) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 InvalidFileNameReason))) :+: (C1 ('MetaCons "ModuleNameHashCollision" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 RawTopLevelModuleName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe RawTopLevelModuleName))) :+: C1 ('MetaCons "BothWithAndRHS" 'PrefixI 'False) (U1 :: Type -> Type))))) :+: (((C1 ('MetaCons "AbstractConstructorNotInScope" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: (C1 ('MetaCons "CopatternHeadNotProjection" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "NotAllowedInDotPatterns" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 NotAllowedInDotPatterns)))) :+: ((C1 ('MetaCons "NotInScope" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "NoSuchModule" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName))) :+: (C1 ('MetaCons "AmbiguousName" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 AmbiguousNameReason)) :+: C1 ('MetaCons "AmbiguousModule" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 ModuleName)))))) :+: (((C1 ('MetaCons "AmbiguousField" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Name) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List2 ModuleName))) :+: C1 ('MetaCons "AmbiguousConstructor" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List2 QName)))) :+: (C1 ('MetaCons "ClashingDefinition" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe NiceDeclaration)))) :+: C1 ('MetaCons "ClashingModule" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ModuleName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ModuleName)))) :+: ((C1 ('MetaCons "DefinitionInDifferentModule" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "DuplicateImports" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 ImportedName)))) :+: (C1 ('MetaCons "InvalidPattern" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Pattern)) :+: C1 ('MetaCons "InvalidPun" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ConstructorOrPatternSynonym) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)))))))) :+: (((((C1 ('MetaCons "RepeatedNamesInImportDirective" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 (List2 ImportedName)))) :+: (C1 ('MetaCons "RepeatedVariablesInPattern" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 Name))) :+: C1 ('MetaCons "GeneralizeNotSupportedHere" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)))) :+: ((C1 ('MetaCons "GeneralizedVarInLetOpenedModule" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "MultipleFixityDecls" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 (Name, Pair Fixity'))))) :+: (C1 ('MetaCons "MultiplePolarityPragmas" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 Name))) :+: C1 ('MetaCons "ExplicitPolarityVsPragma" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName))))) :+: (((C1 ('MetaCons "ConstructorNameOfNonRecord" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ResolvedName)) :+: C1 ('MetaCons "CannotQuote" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 CannotQuote))) :+: (C1 ('MetaCons "CannotQuoteTerm" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 CannotQuoteTerm)) :+: C1 ('MetaCons "DeclarationsAfterTopLevelModule" 'PrefixI 'False) (U1 :: Type -> Type))) :+: ((C1 ('MetaCons "IllegalDeclarationBeforeTopLevelModule" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "MissingTypeSignature" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 MissingTypeSignatureInfo))) :+: (C1 ('MetaCons "NotAnExpression" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr)) :+: C1 ('MetaCons "NotAValidLetBinding" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe NotAValidLetBinding))))))) :+: (((C1 ('MetaCons "NotAValidLetExpression" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 NotAValidLetExpression)) :+: (C1 ('MetaCons "NotValidBeforeField" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 NiceDeclaration)) :+: C1 ('MetaCons "OpenEverythingInRecordWhere" 'PrefixI 'False) (U1 :: Type -> Type))) :+: ((C1 ('MetaCons "PrivateRecordField" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "QualifiedLocalModule" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "AsPatternInPatternSynonym" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "DotPatternInPatternSynonym" 'PrefixI 'False) (U1 :: Type -> Type)))) :+: (((C1 ('MetaCons "BadArgumentsToPatternSynonym" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 AmbiguousQName)) :+: C1 ('MetaCons "TooFewArgumentsToPatternSynonym" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 AmbiguousQName))) :+: (C1 ('MetaCons "CannotResolveAmbiguousPatternSynonym" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 (QName, PatternSynDefn)))) :+: C1 ('MetaCons "IllegalInstanceVariableInPatternSynonym" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Name)))) :+: ((C1 ('MetaCons "PatternSynonymArgumentShadows" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ConstructorOrPatternSynonym) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Name) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 AbstractName)))) :+: C1 ('MetaCons "UnusedVariableInPatternSynonym" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Name))) :+: (C1 ('MetaCons "UnboundVariablesInPatternSynonym" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 Name))) :+: C1 ('MetaCons "NoParseForApplication" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List2 Expr)))))))) :+: ((((C1 ('MetaCons "AmbiguousParseForApplication" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List2 Expr)) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 Expr))) :+: (C1 ('MetaCons "NoParseForLHS" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 LHSOrPatSyn) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Pattern]) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Pattern))) :+: C1 ('MetaCons "AmbiguousParseForLHS" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 LHSOrPatSyn) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Pattern) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List2 Pattern)))))) :+: ((C1 ('MetaCons "AmbiguousProjection" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 QName))) :+: C1 ('MetaCons "AmbiguousOverloadedProjection" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 QName)) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Doc))) :+: (C1 ('MetaCons "OperatorInformation" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [NotationSection]) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 TypeError)) :+: C1 ('MetaCons "InstanceNoCandidate" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [(Term, TCErr)]))))) :+: (((C1 ('MetaCons "ExecError" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ExecError)) :+: C1 ('MetaCons "UnquoteFailed" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 UnquoteError))) :+: (C1 ('MetaCons "DeBruijnIndexOutOfScope" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Nat) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Telescope) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Name]))) :+: C1 ('MetaCons "NeedOptionAllowExec" 'PrefixI 'False) (U1 :: Type -> Type))) :+: ((C1 ('MetaCons "NeedOptionCopatterns" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "NeedOptionCubical" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Cubical) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String))) :+: (C1 ('MetaCons "NeedOptionPatternMatching" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "NeedOptionProp" 'PrefixI 'False) (U1 :: Type -> Type))))) :+: (((C1 ('MetaCons "NeedOptionRewriting" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "NeedOptionSizedTypes" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String)) :+: C1 ('MetaCons "NeedOptionTwoLevel" 'PrefixI 'False) (U1 :: Type -> Type))) :+: ((C1 ('MetaCons "NeedOptionUniversePolymorphism" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "NonFatalErrors" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Set1 TCWarning)))) :+: (C1 ('MetaCons "InstanceSearchDepthExhausted" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Int))) :+: C1 ('MetaCons "TriedToCopyConstrainedPrim" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName))))) :+: (((C1 ('MetaCons "InvalidInstanceHeadType" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 WhyInvalidInstanceType)) :+: C1 ('MetaCons "InteractionError" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 InteractionError))) :+: (C1 ('MetaCons "BackendDoesNotSupportOnlyScopeChecking" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 BackendName)) :+: C1 ('MetaCons "CubicalCompilationNotSupported" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Cubical)))) :+: ((C1 ('MetaCons "CustomBackendError" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 BackendName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Doc)) :+: C1 ('MetaCons "GHCBackendError" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 GHCBackendError))) :+: (C1 ('MetaCons "JSBackendError" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 JSBackendError)) :+: C1 ('MetaCons "UnknownBackend" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 BackendName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Set BackendName)))))))))))
Show TypeError Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep TypeError Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep TypeError = D1 ('MetaData "TypeError" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (((((((C1 ('MetaCons "InternalError" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String)) :+: (C1 ('MetaCons "NotImplemented" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String)) :+: C1 ('MetaCons "NotSupported" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String)))) :+: ((C1 ('MetaCons "CompilationError" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String)) :+: C1 ('MetaCons "SyntaxError" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String))) :+: (C1 ('MetaCons "OptionError" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 OptionError)) :+: C1 ('MetaCons "NicifierError" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 DeclarationException'))))) :+: (((C1 ('MetaCons "DoNotationError" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String)) :+: C1 ('MetaCons "IdiomBracketError" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String))) :+: (C1 ('MetaCons "NoKnownRecordWithSuchFields" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Name])) :+: C1 ('MetaCons "ShouldEndInApplicationOfTheDatatype" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)))) :+: ((C1 ('MetaCons "ConstructorPatternInWrongDatatype" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "CantResolveOverloadedConstructorsTargetingSameDatatype" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 QName)))) :+: (C1 ('MetaCons "ConstructorDoesNotTargetGivenType" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)) :+: C1 ('MetaCons "InvalidDottedExpression" 'PrefixI 'False) (U1 :: Type -> Type))))) :+: (((C1 ('MetaCons "LiteralTooBig" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "NegativeLiteralInPattern" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "WrongHidingInLHS" 'PrefixI 'False) (U1 :: Type -> Type))) :+: ((C1 ('MetaCons "WrongHidingInLambda" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)) :+: C1 ('MetaCons "WrongHidingInApplication" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type))) :+: (C1 ('MetaCons "WrongHidingInProjection" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "IllegalHidingInPostfixProjection" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (NamedArg Expr)))))) :+: (((C1 ('MetaCons "WrongNamedArgument" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (NamedArg Expr)) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 NamedName))) :+: C1 ('MetaCons "WrongAnnotationInLambda" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "WrongIrrelevanceInLambda" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "WrongQuantityInLambda" 'PrefixI 'False) (U1 :: Type -> Type))) :+: ((C1 ('MetaCons "WrongCohesionInLambda" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "WrongPolarityInLambda" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "QuantityMismatch" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Quantity) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Quantity)) :+: C1 ('MetaCons "HidingMismatch" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Hiding) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Hiding))))))) :+: ((((C1 ('MetaCons "RelevanceMismatch" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Relevance) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Relevance)) :+: (C1 ('MetaCons "ForcedConstructorNotInstantiated" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Pattern)) :+: C1 ('MetaCons "IllformedProjectionPatternAbstract" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Pattern)))) :+: ((C1 ('MetaCons "IllformedProjectionPatternConcrete" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Pattern)) :+: C1 ('MetaCons "CannotEliminateWithPattern" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe Blocker)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (NamedArg Pattern)) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)))) :+: (C1 ('MetaCons "CannotEliminateWithProjection" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Arg Type)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName))) :+: C1 ('MetaCons "WrongNumberOfConstructorArguments" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Nat) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Nat)))))) :+: (((C1 ('MetaCons "ShouldBeEmpty" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [DeBruijnPattern])) :+: C1 ('MetaCons "ShouldBeASort" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type))) :+: (C1 ('MetaCons "ShouldBePi" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)) :+: C1 ('MetaCons "ShouldBePath" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)))) :+: ((C1 ('MetaCons "ShouldBeRecordType" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)) :+: C1 ('MetaCons "ShouldBeRecordPattern" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 DeBruijnPattern))) :+: (C1 ('MetaCons "CannotApply" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)) :+: C1 ('MetaCons "InvalidTypeSort" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Sort)))))) :+: (((C1 ('MetaCons "SplitOnCoinductive" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "SplitOnIrrelevant" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Dom Type))) :+: C1 ('MetaCons "SplitOnUnusableCohesion" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Dom Type))))) :+: ((C1 ('MetaCons "SplitOnUnusablePolarity" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Dom Type))) :+: C1 ('MetaCons "SplitOnNonVariable" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type))) :+: (C1 ('MetaCons "SplitOnNonEtaRecord" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "SplitOnAbstract" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName))))) :+: (((C1 ('MetaCons "SplitOnUnchecked" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "SplitOnPartial" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Dom Type)))) :+: (C1 ('MetaCons "SplitInProp" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 DataOrRecordE)) :+: C1 ('MetaCons "DefinitionIsIrrelevant" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)))) :+: ((C1 ('MetaCons "DefinitionIsErased" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "ProjectionIsIrrelevant" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName))) :+: (C1 ('MetaCons "VariableIsIrrelevant" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Name)) :+: C1 ('MetaCons "VariableIsErased" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Name)))))))) :+: (((((C1 ('MetaCons "VariableIsOfUnusableCohesion" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Name) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Cohesion)) :+: (C1 ('MetaCons "LambdaIsErased" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "RecordIsErased" 'PrefixI 'False) (U1 :: Type -> Type))) :+: ((C1 ('MetaCons "InvalidModalTelescopeUse" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Modality)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Modality) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Definition))) :+: C1 ('MetaCons "VariableIsOfUnusablePolarity" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Name) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 PolarityModality))) :+: (C1 ('MetaCons "UnequalLevel" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Level) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Level))) :+: C1 ('MetaCons "UnequalTerms" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 CompareAs)))))) :+: (((C1 ('MetaCons "UnequalRelevance" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term))) :+: C1 ('MetaCons "UnequalQuantity" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)))) :+: (C1 ('MetaCons "UnequalCohesion" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term))) :+: C1 ('MetaCons "UnequalPolarity" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term))))) :+: ((C1 ('MetaCons "UnequalFiniteness" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Comparison) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term))) :+: C1 ('MetaCons "UnequalHiding" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term))) :+: (C1 ('MetaCons "UnequalSorts" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Sort) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Sort)) :+: C1 ('MetaCons "NotLeqSort" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Sort) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Sort)))))) :+: (((C1 ('MetaCons "MetaCannotDependOn" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 MetaId) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Nat))) :+: (C1 ('MetaCons "MetaIrrelevantSolution" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 MetaId) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)) :+: C1 ('MetaCons "MetaErasedSolution" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 MetaId) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)))) :+: ((C1 ('MetaCons "GenericError" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String)) :+: C1 ('MetaCons "GenericDocError" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Doc))) :+: (C1 ('MetaCons "SortOfSplitVarError" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe Blocker)) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Doc)) :+: C1 ('MetaCons "WrongSharpArity" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName))))) :+: (((C1 ('MetaCons "BuiltinMustBeConstructor" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 BuiltinId) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr)) :+: C1 ('MetaCons "BuiltinMustBeData" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 BuiltinId) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Int))) :+: (C1 ('MetaCons "BuiltinMustBeDef" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 BuiltinId)) :+: C1 ('MetaCons "BuiltinMustBeFunction" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 BuiltinId)))) :+: ((C1 ('MetaCons "BuiltinMustBePostulate" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 BuiltinId)) :+: C1 ('MetaCons "NoSuchBuiltinName" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String))) :+: (C1 ('MetaCons "InvalidBuiltin" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String)) :+: C1 ('MetaCons "DuplicateBuiltinBinding" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 BuiltinId) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)))))))) :+: ((((C1 ('MetaCons "NoBindingForBuiltin" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 BuiltinId)) :+: (C1 ('MetaCons "NoBindingForPrimitive" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 PrimitiveId)) :+: C1 ('MetaCons "NoSuchPrimitiveFunction" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String)))) :+: ((C1 ('MetaCons "DuplicatePrimitiveBinding" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 PrimitiveId) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName))) :+: C1 ('MetaCons "WrongArgInfoForPrimitive" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 PrimitiveId) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ArgInfo) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ArgInfo)))) :+: (C1 ('MetaCons "ShadowedModule" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Name) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 ModuleName))) :+: C1 ('MetaCons "BuiltinInParameterisedModule" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 BuiltinId))))) :+: (((C1 ('MetaCons "IllegalDeclarationInDataDefinition" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 Declaration))) :+: C1 ('MetaCons "IllegalLetInTelescope" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 TypedBinding))) :+: (C1 ('MetaCons "IllegalPatternInTelescope" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Binder)) :+: C1 ('MetaCons "AbsentRHSRequiresAbsurdPattern" 'PrefixI 'False) (U1 :: Type -> Type))) :+: ((C1 ('MetaCons "TooManyFields" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Name]) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 Name)))) :+: C1 ('MetaCons "DuplicateFields" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 Name)))) :+: (C1 ('MetaCons "DuplicateConstructors" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 Name))) :+: C1 ('MetaCons "DuplicateOverlapPragma" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 OverlapMode) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 OverlapMode))))))) :+: (((C1 ('MetaCons "WithOnFreeVariable" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)) :+: (C1 ('MetaCons "UnexpectedWithPatterns" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 Pattern))) :+: C1 ('MetaCons "WithClausePatternMismatch" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Pattern) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (NamedArg DeBruijnPattern))))) :+: ((C1 ('MetaCons "IllTypedPatternAfterWithAbstraction" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Pattern)) :+: C1 ('MetaCons "TooFewPatternsInWithClause" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "TooManyPatternsInWithClause" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "PathAbstractionFailed" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Abs Type)))))) :+: (((C1 ('MetaCons "FieldOutsideRecord" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "ModuleArityMismatch" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ModuleName) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Telescope) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Either (List1 (NamedArg Expr)) Args))))) :+: (C1 ('MetaCons "GeneralizeCyclicDependency" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "ReferencesFutureVariables" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 Int))) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Arg Term)) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Int))))) :+: ((C1 ('MetaCons "DoesNotMentionTicks" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Arg Term)))) :+: C1 ('MetaCons "MismatchedProjectionsError" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName))) :+: (C1 ('MetaCons "AttributeKindNotEnabled" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String))) :+: C1 ('MetaCons "InvalidProjectionParameter" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (NamedArg Expr)))))))))) :+: ((((((C1 ('MetaCons "TacticAttributeNotAllowed" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "CannotRewriteByNonEquation" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)) :+: C1 ('MetaCons "MacroResultTypeMismatch" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)))) :+: ((C1 ('MetaCons "NamedWhereModuleInRefinedContext" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Term]) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [String])) :+: C1 ('MetaCons "ComatchingDisabledForRecord" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName))) :+: (C1 ('MetaCons "IncorrectTypeForRewriteRelation" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 IncorrectTypeForRewriteRelationReason)) :+: C1 ('MetaCons "CannotGenerateHCompClause" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type))))) :+: (((C1 ('MetaCons "CannotGenerateTransportClause" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Closure (Abs Type)))) :+: C1 ('MetaCons "CubicalNotErasure" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName))) :+: (C1 ('MetaCons "CubicalPrimitiveNotFullyApplied" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "ExpectedIntervalLiteral" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr)))) :+: ((C1 ('MetaCons "FaceConstraintDisjunction" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "FaceConstraintUnsatisfiable" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "PatternInPathLambda" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "PatternInSystem" 'PrefixI 'False) (U1 :: Type -> Type))))) :+: (((C1 ('MetaCons "UnexpectedParameter" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 LamBinding)) :+: (C1 ('MetaCons "NoParameterOfName" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ArgName)) :+: C1 ('MetaCons "UnexpectedModalityAnnotationInParameter" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 LamBinding)))) :+: ((C1 ('MetaCons "ExpectedBindingForParameter" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Dom Type)) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Abs Type))) :+: C1 ('MetaCons "UnexpectedTypeSignatureForParameter" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 (NamedArg Binder))))) :+: (C1 ('MetaCons "SortDoesNotAdmitDataDefinitions" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Sort)) :+: C1 ('MetaCons "SortCannotDependOnItsIndex" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type))))) :+: (((C1 ('MetaCons "UnusableAtModality" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 WhyCheckModality) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Modality) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term))) :+: C1 ('MetaCons "SplitError" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 SplitError))) :+: (C1 ('MetaCons "ImpossibleConstructor" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 NegativeUnification)) :+: C1 ('MetaCons "DatatypeIndexPolarity" 'PrefixI 'False) (U1 :: Type -> Type))) :+: ((C1 ('MetaCons "RecursiveRecordNeedsInductivity" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "CannotSolveSizeConstraints" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 (ProblemConstraint, HypSizeConstraint))) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Doc))) :+: (C1 ('MetaCons "ContradictorySizeConstraint" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (ProblemConstraint, HypSizeConstraint))) :+: C1 ('MetaCons "EmptyTypeOfSizes" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term))))))) :+: ((((C1 ('MetaCons "FunctionTypeInSizeUniv" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)) :+: (C1 ('MetaCons "PostulatedSizeInModule" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "LibraryError" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 LibErrors)))) :+: ((C1 ('MetaCons "LibTooFarDown" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 TopLevelModuleName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 AgdaLibFile)) :+: C1 ('MetaCons "SolvedButOpenHoles" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "CyclicModuleDependency" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List2 TopLevelModuleName))) :+: C1 ('MetaCons "FileNotFound" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 TopLevelModuleName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [AbsolutePath]))))) :+: (((C1 ('MetaCons "OverlappingProjects" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 AbsolutePath) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 TopLevelModuleName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 TopLevelModuleName))) :+: C1 ('MetaCons "AmbiguousTopLevelModuleName" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 TopLevelModuleName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List2 AbsolutePath)))) :+: (C1 ('MetaCons "ModuleNameUnexpected" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 TopLevelModuleName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 TopLevelModuleName)) :+: C1 ('MetaCons "ModuleNameDoesntMatchFileName" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 TopLevelModuleName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [AbsolutePath])))) :+: ((C1 ('MetaCons "ModuleDefinedInOtherFile" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 TopLevelModuleName) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 AbsolutePath) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 AbsolutePath))) :+: C1 ('MetaCons "InvalidFileName" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 AbsolutePath) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 InvalidFileNameReason))) :+: (C1 ('MetaCons "ModuleNameHashCollision" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 RawTopLevelModuleName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe RawTopLevelModuleName))) :+: C1 ('MetaCons "BothWithAndRHS" 'PrefixI 'False) (U1 :: Type -> Type))))) :+: (((C1 ('MetaCons "AbstractConstructorNotInScope" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: (C1 ('MetaCons "CopatternHeadNotProjection" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "NotAllowedInDotPatterns" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 NotAllowedInDotPatterns)))) :+: ((C1 ('MetaCons "NotInScope" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "NoSuchModule" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName))) :+: (C1 ('MetaCons "AmbiguousName" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 AmbiguousNameReason)) :+: C1 ('MetaCons "AmbiguousModule" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 ModuleName)))))) :+: (((C1 ('MetaCons "AmbiguousField" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Name) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List2 ModuleName))) :+: C1 ('MetaCons "AmbiguousConstructor" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List2 QName)))) :+: (C1 ('MetaCons "ClashingDefinition" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe NiceDeclaration)))) :+: C1 ('MetaCons "ClashingModule" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ModuleName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ModuleName)))) :+: ((C1 ('MetaCons "DefinitionInDifferentModule" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "DuplicateImports" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 ImportedName)))) :+: (C1 ('MetaCons "InvalidPattern" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Pattern)) :+: C1 ('MetaCons "InvalidPun" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ConstructorOrPatternSynonym) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)))))))) :+: (((((C1 ('MetaCons "RepeatedNamesInImportDirective" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 (List2 ImportedName)))) :+: (C1 ('MetaCons "RepeatedVariablesInPattern" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 Name))) :+: C1 ('MetaCons "GeneralizeNotSupportedHere" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)))) :+: ((C1 ('MetaCons "GeneralizedVarInLetOpenedModule" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "MultipleFixityDecls" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 (Name, Pair Fixity'))))) :+: (C1 ('MetaCons "MultiplePolarityPragmas" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 Name))) :+: C1 ('MetaCons "ExplicitPolarityVsPragma" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName))))) :+: (((C1 ('MetaCons "ConstructorNameOfNonRecord" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ResolvedName)) :+: C1 ('MetaCons "CannotQuote" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 CannotQuote))) :+: (C1 ('MetaCons "CannotQuoteTerm" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 CannotQuoteTerm)) :+: C1 ('MetaCons "DeclarationsAfterTopLevelModule" 'PrefixI 'False) (U1 :: Type -> Type))) :+: ((C1 ('MetaCons "IllegalDeclarationBeforeTopLevelModule" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "MissingTypeSignature" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 MissingTypeSignatureInfo))) :+: (C1 ('MetaCons "NotAnExpression" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr)) :+: C1 ('MetaCons "NotAValidLetBinding" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Maybe NotAValidLetBinding))))))) :+: (((C1 ('MetaCons "NotAValidLetExpression" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 NotAValidLetExpression)) :+: (C1 ('MetaCons "NotValidBeforeField" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 NiceDeclaration)) :+: C1 ('MetaCons "OpenEverythingInRecordWhere" 'PrefixI 'False) (U1 :: Type -> Type))) :+: ((C1 ('MetaCons "PrivateRecordField" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "QualifiedLocalModule" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "AsPatternInPatternSynonym" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "DotPatternInPatternSynonym" 'PrefixI 'False) (U1 :: Type -> Type)))) :+: (((C1 ('MetaCons "BadArgumentsToPatternSynonym" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 AmbiguousQName)) :+: C1 ('MetaCons "TooFewArgumentsToPatternSynonym" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 AmbiguousQName))) :+: (C1 ('MetaCons "CannotResolveAmbiguousPatternSynonym" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 (QName, PatternSynDefn)))) :+: C1 ('MetaCons "IllegalInstanceVariableInPatternSynonym" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Name)))) :+: ((C1 ('MetaCons "PatternSynonymArgumentShadows" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ConstructorOrPatternSynonym) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Name) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 AbstractName)))) :+: C1 ('MetaCons "UnusedVariableInPatternSynonym" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Name))) :+: (C1 ('MetaCons "UnboundVariablesInPatternSynonym" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 Name))) :+: C1 ('MetaCons "NoParseForApplication" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List2 Expr)))))))) :+: ((((C1 ('MetaCons "AmbiguousParseForApplication" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List2 Expr)) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 Expr))) :+: (C1 ('MetaCons "NoParseForLHS" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 LHSOrPatSyn) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Pattern]) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Pattern))) :+: C1 ('MetaCons "AmbiguousParseForLHS" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 LHSOrPatSyn) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Pattern) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List2 Pattern)))))) :+: ((C1 ('MetaCons "AmbiguousProjection" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 QName))) :+: C1 ('MetaCons "AmbiguousOverloadedProjection" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 QName)) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Doc))) :+: (C1 ('MetaCons "OperatorInformation" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [NotationSection]) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 TypeError)) :+: C1 ('MetaCons "InstanceNoCandidate" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [(Term, TCErr)]))))) :+: (((C1 ('MetaCons "ExecError" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ExecError)) :+: C1 ('MetaCons "UnquoteFailed" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 UnquoteError))) :+: (C1 ('MetaCons "DeBruijnIndexOutOfScope" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Nat) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Telescope) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Name]))) :+: C1 ('MetaCons "NeedOptionAllowExec" 'PrefixI 'False) (U1 :: Type -> Type))) :+: ((C1 ('MetaCons "NeedOptionCopatterns" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "NeedOptionCubical" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Cubical) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String))) :+: (C1 ('MetaCons "NeedOptionPatternMatching" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "NeedOptionProp" 'PrefixI 'False) (U1 :: Type -> Type))))) :+: (((C1 ('MetaCons "NeedOptionRewriting" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "NeedOptionSizedTypes" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String)) :+: C1 ('MetaCons "NeedOptionTwoLevel" 'PrefixI 'False) (U1 :: Type -> Type))) :+: ((C1 ('MetaCons "NeedOptionUniversePolymorphism" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "NonFatalErrors" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Set1 TCWarning)))) :+: (C1 ('MetaCons "InstanceSearchDepthExhausted" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Int))) :+: C1 ('MetaCons "TriedToCopyConstrainedPrim" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName))))) :+: (((C1 ('MetaCons "InvalidInstanceHeadType" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 WhyInvalidInstanceType)) :+: C1 ('MetaCons "InteractionError" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 InteractionError))) :+: (C1 ('MetaCons "BackendDoesNotSupportOnlyScopeChecking" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 BackendName)) :+: C1 ('MetaCons "CubicalCompilationNotSupported" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Cubical)))) :+: ((C1 ('MetaCons "CustomBackendError" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 BackendName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Doc)) :+: C1 ('MetaCons "GHCBackendError" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 GHCBackendError))) :+: (C1 ('MetaCons "JSBackendError" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 JSBackendError)) :+: C1 ('MetaCons "UnknownBackend" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 BackendName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Set BackendName)))))))))))

data UnificationFailure Source #

Constructors

UnifyIndicesNotVars Telescope Type Term Term Args

Failed to apply injectivity to constructor of indexed datatype

UnifyRecursiveEq Telescope Type Int Term

Can't solve equation because variable occurs in (type of) lhs

UnifyReflexiveEq Telescope Type Term

Can't solve reflexive equation because --without-K is enabled

UnifyUnusableModality Telescope Type Int Term Modality

Can't solve equation because solution modality is less "usable"

Instances

Instances details
PrettyTCM UnificationFailure Source # 
Instance details

Defined in Agda.TypeChecking.Errors

NFData UnificationFailure Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: UnificationFailure -> () #

Generic UnificationFailure Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep UnificationFailure 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep UnificationFailure = D1 ('MetaData "UnificationFailure" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) ((C1 ('MetaCons "UnifyIndicesNotVars" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Telescope) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Args)))) :+: C1 ('MetaCons "UnifyRecursiveEq" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Telescope) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Int) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)))) :+: (C1 ('MetaCons "UnifyReflexiveEq" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Telescope) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term))) :+: C1 ('MetaCons "UnifyUnusableModality" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Telescope) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Int) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Modality))))))
Show UnificationFailure Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep UnificationFailure Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep UnificationFailure = D1 ('MetaData "UnificationFailure" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) ((C1 ('MetaCons "UnifyIndicesNotVars" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Telescope) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Args)))) :+: C1 ('MetaCons "UnifyRecursiveEq" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Telescope) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Int) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)))) :+: (C1 ('MetaCons "UnifyReflexiveEq" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Telescope) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term))) :+: C1 ('MetaCons "UnifyUnusableModality" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Telescope) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Type)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Int) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Modality))))))

data UnquoteError Source #

Instances

Instances details
PrettyTCM UnquoteError Source # 
Instance details

Defined in Agda.TypeChecking.Errors

NFData UnquoteError Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: UnquoteError -> () #

Generic UnquoteError Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep UnquoteError 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep UnquoteError = D1 ('MetaData "UnquoteError" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (((C1 ('MetaCons "BlockedOnMeta" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 TCState) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Blocker)) :+: (C1 ('MetaCons "CannotDeclareHiddenFunction" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "CommitAfterDef" 'PrefixI 'False) (U1 :: Type -> Type))) :+: ((C1 ('MetaCons "ConInsteadOfDef" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String))) :+: C1 ('MetaCons "DefineDataNotData" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName))) :+: (C1 ('MetaCons "DefInsteadOfCon" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String))) :+: C1 ('MetaCons "MissingDeclaration" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName))))) :+: ((C1 ('MetaCons "MissingDefinition" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: (C1 ('MetaCons "NakedUnquote" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "NonCanonical" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)))) :+: ((C1 ('MetaCons "PatLamWithoutClauses" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)) :+: C1 ('MetaCons "StaleMeta" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 TopLevelModuleName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 MetaId))) :+: (C1 ('MetaCons "TooManyParameters" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Nat) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr)) :+: C1 ('MetaCons "UnboundName" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName))))))
Show UnquoteError Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep UnquoteError Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep UnquoteError = D1 ('MetaData "UnquoteError" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (((C1 ('MetaCons "BlockedOnMeta" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 TCState) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Blocker)) :+: (C1 ('MetaCons "CannotDeclareHiddenFunction" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "CommitAfterDef" 'PrefixI 'False) (U1 :: Type -> Type))) :+: ((C1 ('MetaCons "ConInsteadOfDef" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String))) :+: C1 ('MetaCons "DefineDataNotData" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName))) :+: (C1 ('MetaCons "DefInsteadOfCon" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String))) :+: C1 ('MetaCons "MissingDeclaration" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName))))) :+: ((C1 ('MetaCons "MissingDefinition" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: (C1 ('MetaCons "NakedUnquote" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "NonCanonical" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)))) :+: ((C1 ('MetaCons "PatLamWithoutClauses" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)) :+: C1 ('MetaCons "StaleMeta" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 TopLevelModuleName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 MetaId))) :+: (C1 ('MetaCons "TooManyParameters" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Nat) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr)) :+: C1 ('MetaCons "UnboundName" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName))))))

data UnquoteFlags Source #

Constructors

UnquoteFlags 

Instances

Instances details
NFData UnquoteFlags Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: UnquoteFlags -> () #

Generic UnquoteFlags Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep UnquoteFlags 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep UnquoteFlags = D1 ('MetaData "UnquoteFlags" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "UnquoteFlags" 'PrefixI 'True) (S1 ('MetaSel ('Just "_unquoteNormalise") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool)))
type Rep UnquoteFlags Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep UnquoteFlags = D1 ('MetaData "UnquoteFlags" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "UnquoteFlags" 'PrefixI 'True) (S1 ('MetaSel ('Just "_unquoteNormalise") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Bool)))

data Warning Source #

A non-fatal error is an error which does not prevent us from checking the document further and interacting with the user.

Constructors

NicifierIssue DeclarationWarning 
TerminationIssue (List1 TerminationError) 
UnreachableClauses QName (List1 Range)

UnreachableClauses f rs means that the clauses in f whose ranges are rs are unreachable.

CoverageIssue QName (List1 (Telescope, [NamedArg DeBruijnPattern]))

`CoverageIssue f pss` means that pss are not covered in f

CoverageNoExactSplit QName (List1 Clause) 
InlineNoExactSplit QName Clause

Clause was turned into copattern matching clause(s) by an {-# INLINE constructor #-} and thus is not a definitional equality any more.

NotStrictlyPositive QName (Seq OccursWhere) 
ConstructorDoesNotFitInData DataOrRecord_ QName Sort Sort TCErr

Checking whether constructor QName Sort fits into data Sort produced TCErr.

CoinductiveEtaRecord QName

A record type declared as both coinductive and having eta-equality.

UnsolvedMetaVariables (Set1 Range)

Do not use directly with warning

UnsolvedInteractionMetas (Set1 Range)

Do not use directly with warning

UnsolvedConstraints (List1 ProblemConstraint)

Do not use directly with warning

InteractionMetaBoundaries (Set1 Range)

Do not use directly with warning

CantGeneralizeOverSorts (Set1 MetaId) 
AbsurdPatternRequiresAbsentRHS 
OldBuiltin BuiltinId BuiltinId

In `OldBuiltin old new`, the BUILTIN old has been replaced by new.

BuiltinDeclaresIdentifier BuiltinId

The builtin declares a new identifier, so it should not be in scope.

DuplicateRecordDirective RecordDirective

The given record directive is conflicting with a prior one in the same record declaration.

EmptyRewritePragma

If the user wrote just {-# REWRITE #-}.

EmptyWhere

An empty where block is dead code. TODO: linearity -- | FixingQuantity String Quantity Quantity -- -- ^ Auto-correcting quantity pertaining to String from to.

FixingRelevance String Relevance Relevance

Auto-correcting relevance pertaining to String from to.

FixingCohesion String Cohesion Cohesion

Auto-correcting cohesion pertaining to String from to.

FixingPolarity String PolarityModality PolarityModality

Auto-correcting polarity pertaining to String from to.

IllformedAsClause String

If the user wrote something other than an unqualified name in the as clause of an import statement. The String gives optionally extra explanation.

InvalidCharacterLiteral Char

A character literal Agda does not support, e.g. surrogate code points.

ClashesViaRenaming NameOrModule (Set1 Name)

If a renaming import directive introduces a name or module name clash in the exported names of a module. (See issue #4154.)

UselessPatternDeclarationForRecord String

The 'pattern' declaration is useless in the presence of either coinductive or eta-equality. Content of String is "coinductive" or "eta", resp.

UselessPragma Range Doc

Warning when pragma is useless and thus ignored. Range is for dead code highlighting.

UselessPublic UselessPublicReason

If the user opens a module public before the module header. (See issue #2377.)

UselessHiding (List1 ImportedName)

Names in hiding directive that don't hide anything imported by a using directive.

UselessInline QName 
UselessTactic

A tactic attribute applied to a non-hidden (visible or instance) argument.

WrongInstanceDeclaration 
InstanceWithExplicitArg QName

An instance was declared with an implicit argument, which means it will never actually be considered by instance search.

InstanceNoOutputTypeName Doc

The type of an instance argument doesn't end in a named or variable type, so it will never be considered by instance search.

InstanceArgWithExplicitArg Doc

As InstanceWithExplicitArg, but for local bindings rather than top-level instances.

InversionDepthReached QName

The --inversion-max-depth was reached.

SafeFlagPostulate QName 
SafeFlagPragma (Set String)

Unsafe OPTIONS.

SafeFlagWithoutKFlagPrimEraseEquality 
WithoutKFlagPrimEraseEquality 
ConflictingPragmaOptions String String

`ConflictingPragmaOptions a b`: Inconsistent options `--a` and `--no-b`, since `--a` implies `--b`. Ignoring `--no-b`.

OptionWarning OptionWarning 
ParseWarning ParseWarning 
LibraryWarning LibWarning 
DeprecationWarning String String String

`DeprecationWarning old new version`: old is deprecated, use new instead. This will be an error in Agda 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 using directive(s).

FixityInRenamingModule (List1 Range)

Fixity of modules cannot be changed via renaming (since modules have no fixity).

ModuleDoesntExport QName [Name] [Name] (List1 ImportedName)

Some imported names are not actually exported by the source module. The second argument is the names that could be exported. The third argument is the module names that could be exported.

InfectiveImport Doc

Importing a file using an infective option into one which doesn't

CoInfectiveImport Doc

Importing a file not using a coinfective option from one which does

ConfluenceCheckingIncompleteBecauseOfMeta QName

Confluence checking incomplete because the definition of the QName contains unsolved metavariables.

ConfluenceForCubicalNotSupported

Confluence checking with --cubical might be incomplete.

NotARewriteRule QName IsAmbiguous

IllegalRewriteRule detected during scope checking.

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 u that reduces to both v1 and v2 and there is no rule to resolve the ambiguity.

RewriteMissingRule Term Term Term

The global confluence checker found a term u that reduces to v, but v does not reduce to rho(u).

PragmaCompileErased BackendName QName

COMPILE directive for an erased symbol.

PragmaCompileList

COMPILE GHC pragma for lists; ignored.

PragmaCompileMaybe

COMPILE GHC pragma for MAYBE; ignored.

PragmaCompileUnparsable String

COMPILE GHC pragma String not parsable; ignored.

PragmaCompileWrong QName String

Wrong COMPILE GHC given for QName; explanation is in String.

PragmaCompileWrongName QName IsAmbiguous

COMPILE pragma with name QName that is not an unambiguous constructor or definition.

PragmaExpectsDefinedSymbol String QName

Pragma String with name QName that is not an Def.

PragmaExpectsUnambiguousConstructorOrFunction String QName IsAmbiguous

Pragma String with name QName that is not an unambiguous constructor or definition. General form of PragmaCompileWrongName and NotARewriteRule.

PragmaExpectsUnambiguousProjectionOrFunction String QName IsAmbiguous

Pragma String with name QName that is not an unambiguous projection or function.

NoMain TopLevelModuleName

Compiler run on module that does not have a main function.

NotInScopeW QName

Out of scope error we can recover from.

UnsupportedIndexedMatch Doc

Was not able to compute a full equivalence when splitting.

AsPatternShadowsConstructorOrPatternSynonym ConstructorOrPatternSynonym

The as-name in an as-pattern may not shadow a constructor or pattern synonym name, because this can be confusing to read.

PatternShadowsConstructor Name QName

A pattern variable has the name of a constructor (data constructor or matchable record constructor).

PlentyInHardCompileTimeMode QωOrigin

Explicit use of @ω or @plenty in hard compile-time mode.

RecordFieldWarning RecordFieldWarning 
MissingTypeSignatureForOpaque QName IsOpaque

An abstract or opaque definition lacks a type signature.

NotAffectedByOpaque 
UnfoldingWrongName QName

Name in unfolding clause does not resolve to unambiguous defined name.

UnfoldTransparentName QName 
UselessOpaque 
HiddenNotInArgumentPosition Expr 
InstanceNotInArgumentPosition Expr 
MacroInLetBindings 
AbstractInLetBindings 
InvalidDisplayForm QName String

DISPLAY form for QName is invalid because String.

UnusedVariablesInDisplayForm (List1 Name)

The given names are bound in the lhs of the display form but not used on the rhs. This can indicate a user misunderstanding of display forms.

TooManyArgumentsToSort QName (List1 (NamedArg Expr))

Extra arguments to sort (will be ignored).

WithClauseProjectionFixityMismatch

The with-clause uses projection in a different fixity style than the parent clause.

TooManyPolarities QName PragmaPolarities

Too many polarities given in POLARITY pragma for QName. PragmaPolarities contains the (likely) excessive polarities.

TopLevelPolarity QName PolarityModality

Definition with non-default polarity annotation.

FaceConstraintCannotBeHidden ArgInfo

Face constraint patterns (i = 0) must be visible arguments.

FaceConstraintCannotBeNamed NamedName

Face constraint patterns (i = 0) must be unnamed arguments.

CustomBackendWarning String Doc

Used for backend-specific warnings. The string is the backend name.

Instances

Instances details
PrettyTCM Warning Source # 
Instance details

Defined in Agda.TypeChecking.Pretty.Warning

Methods

prettyTCM :: MonadPretty m => Warning -> m Doc Source #

EmbPrj Warning Source # 
Instance details

Defined in Agda.TypeChecking.Serialise.Instances.Errors

NFData Warning Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: Warning -> () #

Generic Warning Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep Warning 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep Warning = D1 ('MetaData "Warning" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) ((((((C1 ('MetaCons "NicifierIssue" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 DeclarationWarning)) :+: C1 ('MetaCons "TerminationIssue" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 TerminationError)))) :+: (C1 ('MetaCons "UnreachableClauses" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 Range))) :+: (C1 ('MetaCons "CoverageIssue" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 (Telescope, [NamedArg DeBruijnPattern])))) :+: C1 ('MetaCons "CoverageNoExactSplit" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 Clause)))))) :+: ((C1 ('MetaCons "InlineNoExactSplit" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Clause)) :+: (C1 ('MetaCons "NotStrictlyPositive" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Seq OccursWhere))) :+: C1 ('MetaCons "ConstructorDoesNotFitInData" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 DataOrRecord_) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Sort) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Sort) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 TCErr)))))) :+: (C1 ('MetaCons "CoinductiveEtaRecord" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: (C1 ('MetaCons "UnsolvedMetaVariables" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Set1 Range))) :+: C1 ('MetaCons "UnsolvedInteractionMetas" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Set1 Range))))))) :+: (((C1 ('MetaCons "UnsolvedConstraints" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 ProblemConstraint))) :+: (C1 ('MetaCons "InteractionMetaBoundaries" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Set1 Range))) :+: C1 ('MetaCons "CantGeneralizeOverSorts" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Set1 MetaId))))) :+: (C1 ('MetaCons "AbsurdPatternRequiresAbsentRHS" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "OldBuiltin" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 BuiltinId) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 BuiltinId)) :+: C1 ('MetaCons "BuiltinDeclaresIdentifier" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 BuiltinId))))) :+: ((C1 ('MetaCons "DuplicateRecordDirective" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 RecordDirective)) :+: (C1 ('MetaCons "EmptyRewritePragma" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "EmptyWhere" 'PrefixI 'False) (U1 :: Type -> Type))) :+: (C1 ('MetaCons "FixingRelevance" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Relevance) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Relevance))) :+: (C1 ('MetaCons "FixingCohesion" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Cohesion) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Cohesion))) :+: C1 ('MetaCons "FixingPolarity" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 PolarityModality) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 PolarityModality)))))))) :+: ((((C1 ('MetaCons "IllformedAsClause" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String)) :+: (C1 ('MetaCons "InvalidCharacterLiteral" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Char)) :+: C1 ('MetaCons "ClashesViaRenaming" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 NameOrModule) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Set1 Name))))) :+: (C1 ('MetaCons "UselessPatternDeclarationForRecord" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String)) :+: (C1 ('MetaCons "UselessPragma" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Doc)) :+: C1 ('MetaCons "UselessPublic" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 UselessPublicReason))))) :+: ((C1 ('MetaCons "UselessHiding" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 ImportedName))) :+: (C1 ('MetaCons "UselessInline" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "UselessTactic" 'PrefixI 'False) (U1 :: Type -> Type))) :+: (C1 ('MetaCons "WrongInstanceDeclaration" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "InstanceWithExplicitArg" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "InstanceNoOutputTypeName" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Doc)))))) :+: (((C1 ('MetaCons "InstanceArgWithExplicitArg" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Doc)) :+: (C1 ('MetaCons "InversionDepthReached" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "SafeFlagPostulate" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)))) :+: (C1 ('MetaCons "SafeFlagPragma" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Set String))) :+: (C1 ('MetaCons "SafeFlagWithoutKFlagPrimEraseEquality" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "WithoutKFlagPrimEraseEquality" 'PrefixI 'False) (U1 :: Type -> Type)))) :+: ((C1 ('MetaCons "ConflictingPragmaOptions" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String)) :+: (C1 ('MetaCons "OptionWarning" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 OptionWarning)) :+: C1 ('MetaCons "ParseWarning" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ParseWarning)))) :+: (C1 ('MetaCons "LibraryWarning" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 LibWarning)) :+: (C1 ('MetaCons "DeprecationWarning" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String))) :+: C1 ('MetaCons "UserWarning" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Text)))))))) :+: (((((C1 ('MetaCons "DuplicateUsing" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 ImportedName))) :+: C1 ('MetaCons "FixityInRenamingModule" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 Range)))) :+: (C1 ('MetaCons "ModuleDoesntExport" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Name])) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Name]) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 ImportedName)))) :+: (C1 ('MetaCons "InfectiveImport" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Doc)) :+: C1 ('MetaCons "CoInfectiveImport" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Doc))))) :+: ((C1 ('MetaCons "ConfluenceCheckingIncompleteBecauseOfMeta" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: (C1 ('MetaCons "ConfluenceForCubicalNotSupported" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "NotARewriteRule" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 IsAmbiguous)))) :+: (C1 ('MetaCons "IllegalRewriteRule" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 IllegalRewriteRuleReason)) :+: (C1 ('MetaCons "RewriteNonConfluent" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Doc))) :+: C1 ('MetaCons "RewriteMaybeNonConfluent" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Doc]))))))) :+: (((C1 ('MetaCons "RewriteAmbiguousRules" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term))) :+: (C1 ('MetaCons "RewriteMissingRule" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term))) :+: C1 ('MetaCons "PragmaCompileErased" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 BackendName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)))) :+: (C1 ('MetaCons "PragmaCompileList" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "PragmaCompileMaybe" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "PragmaCompileUnparsable" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String))))) :+: ((C1 ('MetaCons "PragmaCompileWrong" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String)) :+: (C1 ('MetaCons "PragmaCompileWrongName" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 IsAmbiguous)) :+: C1 ('MetaCons "PragmaExpectsDefinedSymbol" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)))) :+: (C1 ('MetaCons "PragmaExpectsUnambiguousConstructorOrFunction" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 IsAmbiguous))) :+: (C1 ('MetaCons "PragmaExpectsUnambiguousProjectionOrFunction" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 IsAmbiguous))) :+: C1 ('MetaCons "NoMain" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 TopLevelModuleName))))))) :+: ((((C1 ('MetaCons "NotInScopeW" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: (C1 ('MetaCons "UnsupportedIndexedMatch" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Doc)) :+: C1 ('MetaCons "AsPatternShadowsConstructorOrPatternSynonym" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ConstructorOrPatternSynonym)))) :+: (C1 ('MetaCons "PatternShadowsConstructor" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Name) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: (C1 ('MetaCons "PlentyInHardCompileTimeMode" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QωOrigin)) :+: C1 ('MetaCons "RecordFieldWarning" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 RecordFieldWarning))))) :+: ((C1 ('MetaCons "MissingTypeSignatureForOpaque" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 IsOpaque)) :+: (C1 ('MetaCons "NotAffectedByOpaque" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "UnfoldingWrongName" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)))) :+: (C1 ('MetaCons "UnfoldTransparentName" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: (C1 ('MetaCons "UselessOpaque" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "HiddenNotInArgumentPosition" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr)))))) :+: (((C1 ('MetaCons "InstanceNotInArgumentPosition" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr)) :+: (C1 ('MetaCons "MacroInLetBindings" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "AbstractInLetBindings" 'PrefixI 'False) (U1 :: Type -> Type))) :+: (C1 ('MetaCons "InvalidDisplayForm" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String)) :+: (C1 ('MetaCons "UnusedVariablesInDisplayForm" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 Name))) :+: C1 ('MetaCons "TooManyArgumentsToSort" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 (NamedArg Expr))))))) :+: ((C1 ('MetaCons "WithClauseProjectionFixityMismatch" 'PrefixI 'True) ((S1 ('MetaSel ('Just "withClausePattern") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (NamedArg Pattern)) :*: S1 ('MetaSel ('Just "withClauseProjectionOrigin") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ProjOrigin)) :*: (S1 ('MetaSel ('Just "parentPattern") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (NamedArg DeBruijnPattern)) :*: S1 ('MetaSel ('Just "parentProjectionOrigin") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ProjOrigin))) :+: (C1 ('MetaCons "TooManyPolarities" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 PragmaPolarities)) :+: C1 ('MetaCons "TopLevelPolarity" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 PolarityModality)))) :+: (C1 ('MetaCons "FaceConstraintCannotBeHidden" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ArgInfo)) :+: (C1 ('MetaCons "FaceConstraintCannotBeNamed" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 NamedName)) :+: C1 ('MetaCons "CustomBackendWarning" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Doc)))))))))

Methods

from :: Warning -> Rep Warning x #

to :: Rep Warning x -> Warning #

Show Warning Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep Warning Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep Warning = D1 ('MetaData "Warning" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) ((((((C1 ('MetaCons "NicifierIssue" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 DeclarationWarning)) :+: C1 ('MetaCons "TerminationIssue" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 TerminationError)))) :+: (C1 ('MetaCons "UnreachableClauses" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 Range))) :+: (C1 ('MetaCons "CoverageIssue" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 (Telescope, [NamedArg DeBruijnPattern])))) :+: C1 ('MetaCons "CoverageNoExactSplit" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 Clause)))))) :+: ((C1 ('MetaCons "InlineNoExactSplit" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Clause)) :+: (C1 ('MetaCons "NotStrictlyPositive" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Seq OccursWhere))) :+: C1 ('MetaCons "ConstructorDoesNotFitInData" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 DataOrRecord_) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Sort) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Sort) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 TCErr)))))) :+: (C1 ('MetaCons "CoinductiveEtaRecord" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: (C1 ('MetaCons "UnsolvedMetaVariables" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Set1 Range))) :+: C1 ('MetaCons "UnsolvedInteractionMetas" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Set1 Range))))))) :+: (((C1 ('MetaCons "UnsolvedConstraints" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 ProblemConstraint))) :+: (C1 ('MetaCons "InteractionMetaBoundaries" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Set1 Range))) :+: C1 ('MetaCons "CantGeneralizeOverSorts" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Set1 MetaId))))) :+: (C1 ('MetaCons "AbsurdPatternRequiresAbsentRHS" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "OldBuiltin" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 BuiltinId) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 BuiltinId)) :+: C1 ('MetaCons "BuiltinDeclaresIdentifier" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 BuiltinId))))) :+: ((C1 ('MetaCons "DuplicateRecordDirective" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 RecordDirective)) :+: (C1 ('MetaCons "EmptyRewritePragma" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "EmptyWhere" 'PrefixI 'False) (U1 :: Type -> Type))) :+: (C1 ('MetaCons "FixingRelevance" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Relevance) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Relevance))) :+: (C1 ('MetaCons "FixingCohesion" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Cohesion) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Cohesion))) :+: C1 ('MetaCons "FixingPolarity" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 PolarityModality) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 PolarityModality)))))))) :+: ((((C1 ('MetaCons "IllformedAsClause" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String)) :+: (C1 ('MetaCons "InvalidCharacterLiteral" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Char)) :+: C1 ('MetaCons "ClashesViaRenaming" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 NameOrModule) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Set1 Name))))) :+: (C1 ('MetaCons "UselessPatternDeclarationForRecord" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String)) :+: (C1 ('MetaCons "UselessPragma" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Range) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Doc)) :+: C1 ('MetaCons "UselessPublic" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 UselessPublicReason))))) :+: ((C1 ('MetaCons "UselessHiding" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 ImportedName))) :+: (C1 ('MetaCons "UselessInline" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "UselessTactic" 'PrefixI 'False) (U1 :: Type -> Type))) :+: (C1 ('MetaCons "WrongInstanceDeclaration" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "InstanceWithExplicitArg" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "InstanceNoOutputTypeName" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Doc)))))) :+: (((C1 ('MetaCons "InstanceArgWithExplicitArg" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Doc)) :+: (C1 ('MetaCons "InversionDepthReached" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: C1 ('MetaCons "SafeFlagPostulate" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)))) :+: (C1 ('MetaCons "SafeFlagPragma" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (Set String))) :+: (C1 ('MetaCons "SafeFlagWithoutKFlagPrimEraseEquality" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "WithoutKFlagPrimEraseEquality" 'PrefixI 'False) (U1 :: Type -> Type)))) :+: ((C1 ('MetaCons "ConflictingPragmaOptions" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String)) :+: (C1 ('MetaCons "OptionWarning" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 OptionWarning)) :+: C1 ('MetaCons "ParseWarning" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ParseWarning)))) :+: (C1 ('MetaCons "LibraryWarning" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 LibWarning)) :+: (C1 ('MetaCons "DeprecationWarning" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String))) :+: C1 ('MetaCons "UserWarning" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Text)))))))) :+: (((((C1 ('MetaCons "DuplicateUsing" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 ImportedName))) :+: C1 ('MetaCons "FixityInRenamingModule" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 Range)))) :+: (C1 ('MetaCons "ModuleDoesntExport" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Name])) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Name]) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 ImportedName)))) :+: (C1 ('MetaCons "InfectiveImport" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Doc)) :+: C1 ('MetaCons "CoInfectiveImport" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Doc))))) :+: ((C1 ('MetaCons "ConfluenceCheckingIncompleteBecauseOfMeta" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: (C1 ('MetaCons "ConfluenceForCubicalNotSupported" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "NotARewriteRule" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 IsAmbiguous)))) :+: (C1 ('MetaCons "IllegalRewriteRule" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 IllegalRewriteRuleReason)) :+: (C1 ('MetaCons "RewriteNonConfluent" 'PrefixI 'False) ((S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term)) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Doc))) :+: C1 ('MetaCons "RewriteMaybeNonConfluent" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Doc]))))))) :+: (((C1 ('MetaCons "RewriteAmbiguousRules" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term))) :+: (C1 ('MetaCons "RewriteMissingRule" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Term))) :+: C1 ('MetaCons "PragmaCompileErased" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 BackendName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)))) :+: (C1 ('MetaCons "PragmaCompileList" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "PragmaCompileMaybe" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "PragmaCompileUnparsable" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String))))) :+: ((C1 ('MetaCons "PragmaCompileWrong" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String)) :+: (C1 ('MetaCons "PragmaCompileWrongName" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 IsAmbiguous)) :+: C1 ('MetaCons "PragmaExpectsDefinedSymbol" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)))) :+: (C1 ('MetaCons "PragmaExpectsUnambiguousConstructorOrFunction" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 IsAmbiguous))) :+: (C1 ('MetaCons "PragmaExpectsUnambiguousProjectionOrFunction" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String) :*: (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 IsAmbiguous))) :+: C1 ('MetaCons "NoMain" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 TopLevelModuleName))))))) :+: ((((C1 ('MetaCons "NotInScopeW" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: (C1 ('MetaCons "UnsupportedIndexedMatch" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Doc)) :+: C1 ('MetaCons "AsPatternShadowsConstructorOrPatternSynonym" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ConstructorOrPatternSynonym)))) :+: (C1 ('MetaCons "PatternShadowsConstructor" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Name) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: (C1 ('MetaCons "PlentyInHardCompileTimeMode" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QωOrigin)) :+: C1 ('MetaCons "RecordFieldWarning" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 RecordFieldWarning))))) :+: ((C1 ('MetaCons "MissingTypeSignatureForOpaque" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 IsOpaque)) :+: (C1 ('MetaCons "NotAffectedByOpaque" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "UnfoldingWrongName" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)))) :+: (C1 ('MetaCons "UnfoldTransparentName" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName)) :+: (C1 ('MetaCons "UselessOpaque" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "HiddenNotInArgumentPosition" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr)))))) :+: (((C1 ('MetaCons "InstanceNotInArgumentPosition" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Expr)) :+: (C1 ('MetaCons "MacroInLetBindings" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "AbstractInLetBindings" 'PrefixI 'False) (U1 :: Type -> Type))) :+: (C1 ('MetaCons "InvalidDisplayForm" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String)) :+: (C1 ('MetaCons "UnusedVariablesInDisplayForm" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 Name))) :+: C1 ('MetaCons "TooManyArgumentsToSort" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (List1 (NamedArg Expr))))))) :+: ((C1 ('MetaCons "WithClauseProjectionFixityMismatch" 'PrefixI 'True) ((S1 ('MetaSel ('Just "withClausePattern") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (NamedArg Pattern)) :*: S1 ('MetaSel ('Just "withClauseProjectionOrigin") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ProjOrigin)) :*: (S1 ('MetaSel ('Just "parentPattern") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (NamedArg DeBruijnPattern)) :*: S1 ('MetaSel ('Just "parentProjectionOrigin") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ProjOrigin))) :+: (C1 ('MetaCons "TooManyPolarities" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 PragmaPolarities)) :+: C1 ('MetaCons "TopLevelPolarity" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 QName) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 PolarityModality)))) :+: (C1 ('MetaCons "FaceConstraintCannotBeHidden" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 ArgInfo)) :+: (C1 ('MetaCons "FaceConstraintCannotBeNamed" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 NamedName)) :+: C1 ('MetaCons "CustomBackendWarning" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Doc)))))))))

data WarningsAndNonFatalErrors Source #

Assorted warnings and errors to be displayed to the user

data WhyCheckModality Source #

Why are we performing a modality check?

Constructors

ConstructorType

Because --without-K is enabled, so the types of data constructors must be usable at the context's modality.

IndexedClause

Because --without-K is enabled, so the result type of clauses must be usable at the context's modality.

IndexedClauseArg Name Name

Because --without-K is enabled, so any argument (second name) which mentions a dotted argument (first name) must have a type which is usable at the context's modality.

GeneratedClause

Because we double-check the --cubical-compatible clauses. This is an internal error!

Instances

Instances details
NFData WhyCheckModality Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: WhyCheckModality -> () #

Generic WhyCheckModality Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep WhyCheckModality 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep WhyCheckModality = D1 ('MetaData "WhyCheckModality" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) ((C1 ('MetaCons "ConstructorType" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "IndexedClause" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "IndexedClauseArg" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Name) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Name)) :+: C1 ('MetaCons "GeneratedClause" 'PrefixI 'False) (U1 :: Type -> Type)))
Show WhyCheckModality Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep WhyCheckModality Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep WhyCheckModality = D1 ('MetaData "WhyCheckModality" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) ((C1 ('MetaCons "ConstructorType" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "IndexedClause" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "IndexedClauseArg" 'PrefixI 'False) (S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Name) :*: S1 ('MetaSel ('Nothing :: Maybe Symbol) 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Name)) :+: C1 ('MetaCons "GeneratedClause" 'PrefixI 'False) (U1 :: Type -> Type)))

data WhyInvalidInstanceType Source #

Reason for why the instance type is invalid.

Constructors

ImproperInstHead

The type isn't headed by a local, a definition, or a postulate (e.g. it's a universe)

ImproperInstTele

The type we're looking for has a visible argument

Instances

Instances details
NFData WhyInvalidInstanceType Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: WhyInvalidInstanceType -> () #

Generic WhyInvalidInstanceType Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Associated Types

type Rep WhyInvalidInstanceType 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep WhyInvalidInstanceType = D1 ('MetaData "WhyInvalidInstanceType" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "ImproperInstHead" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "ImproperInstTele" 'PrefixI 'False) (U1 :: Type -> Type))
Show WhyInvalidInstanceType Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep WhyInvalidInstanceType Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep WhyInvalidInstanceType = D1 ('MetaData "WhyInvalidInstanceType" "Agda.TypeChecking.Monad.Base" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "ImproperInstHead" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "ImproperInstTele" 'PrefixI 'False) (U1 :: Type -> Type))

data WhyNotAHaskellType Source #

Extra information for NotAHaskellType error.

Instances

Instances details
NFData WhyNotAHaskellType Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: WhyNotAHaskellType -> () #

Generic WhyNotAHaskellType Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Show WhyNotAHaskellType Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

type Rep WhyNotAHaskellType Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

data CannotQuoteTerm Source #

Extra information for error CannotQuoteTerm.

Instances

Instances details
NFData CannotQuoteTerm Source # 
Instance details

Defined in Agda.Interaction.Options.Errors

Methods

rnf :: CannotQuoteTerm -> () #

Bounded CannotQuoteTerm Source # 
Instance details

Defined in Agda.Interaction.Options.Errors

Enum CannotQuoteTerm Source # 
Instance details

Defined in Agda.Interaction.Options.Errors

Generic CannotQuoteTerm Source # 
Instance details

Defined in Agda.Interaction.Options.Errors

Associated Types

type Rep CannotQuoteTerm 
Instance details

Defined in Agda.Interaction.Options.Errors

type Rep CannotQuoteTerm = D1 ('MetaData "CannotQuoteTerm" "Agda.Interaction.Options.Errors" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "CannotQuoteTermHidden" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "CannotQuoteTermNothing" 'PrefixI 'False) (U1 :: Type -> Type))
Show CannotQuoteTerm Source # 
Instance details

Defined in Agda.Interaction.Options.Errors

type Rep CannotQuoteTerm Source # 
Instance details

Defined in Agda.Interaction.Options.Errors

type Rep CannotQuoteTerm = D1 ('MetaData "CannotQuoteTerm" "Agda.Interaction.Options.Errors" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "CannotQuoteTermHidden" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "CannotQuoteTermNothing" 'PrefixI 'False) (U1 :: Type -> Type))

data ErasedDatatypeReason Source #

The reason for an ErasedDatatype error.

Constructors

SeveralConstructors

There are several constructors.

NoErasedMatches

The flag --erased-matches is not used.

NoK

The K rule is not activated.

Instances

Instances details
NFData ErasedDatatypeReason Source # 
Instance details

Defined in Agda.Interaction.Options.Errors

Methods

rnf :: ErasedDatatypeReason -> () #

Bounded ErasedDatatypeReason Source # 
Instance details

Defined in Agda.Interaction.Options.Errors

Enum ErasedDatatypeReason Source # 
Instance details

Defined in Agda.Interaction.Options.Errors

Generic ErasedDatatypeReason Source # 
Instance details

Defined in Agda.Interaction.Options.Errors

Associated Types

type Rep ErasedDatatypeReason 
Instance details

Defined in Agda.Interaction.Options.Errors

type Rep ErasedDatatypeReason = D1 ('MetaData "ErasedDatatypeReason" "Agda.Interaction.Options.Errors" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "SeveralConstructors" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "NoErasedMatches" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "NoK" 'PrefixI 'False) (U1 :: Type -> Type)))
Show ErasedDatatypeReason Source # 
Instance details

Defined in Agda.Interaction.Options.Errors

type Rep ErasedDatatypeReason Source # 
Instance details

Defined in Agda.Interaction.Options.Errors

type Rep ErasedDatatypeReason = D1 ('MetaData "ErasedDatatypeReason" "Agda.Interaction.Options.Errors" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "SeveralConstructors" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "NoErasedMatches" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "NoK" 'PrefixI 'False) (U1 :: Type -> Type)))

data NotAValidLetBinding Source #

Reasons for error NotAValidLetBinding.

Instances

Instances details
NFData NotAValidLetBinding Source # 
Instance details

Defined in Agda.Interaction.Options.Errors

Methods

rnf :: NotAValidLetBinding -> () #

Bounded NotAValidLetBinding Source # 
Instance details

Defined in Agda.Interaction.Options.Errors

Enum NotAValidLetBinding Source # 
Instance details

Defined in Agda.Interaction.Options.Errors

Generic NotAValidLetBinding Source # 
Instance details

Defined in Agda.Interaction.Options.Errors

Associated Types

type Rep NotAValidLetBinding 
Instance details

Defined in Agda.Interaction.Options.Errors

type Rep NotAValidLetBinding = D1 ('MetaData "NotAValidLetBinding" "Agda.Interaction.Options.Errors" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "MissingRHS" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "NotAValidLetPattern" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "WhereClausesNotAllowed" 'PrefixI 'False) (U1 :: Type -> Type)))
Show NotAValidLetBinding Source # 
Instance details

Defined in Agda.Interaction.Options.Errors

type Rep NotAValidLetBinding Source # 
Instance details

Defined in Agda.Interaction.Options.Errors

type Rep NotAValidLetBinding = D1 ('MetaData "NotAValidLetBinding" "Agda.Interaction.Options.Errors" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "MissingRHS" 'PrefixI 'False) (U1 :: Type -> Type) :+: (C1 ('MetaCons "NotAValidLetPattern" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "WhereClausesNotAllowed" 'PrefixI 'False) (U1 :: Type -> Type)))

data NotAValidLetExpression Source #

Reasons for error NotAValidLetExpression.

Constructors

MissingBody 

Instances

Instances details
NFData NotAValidLetExpression Source # 
Instance details

Defined in Agda.Interaction.Options.Errors

Methods

rnf :: NotAValidLetExpression -> () #

Bounded NotAValidLetExpression Source # 
Instance details

Defined in Agda.Interaction.Options.Errors

Enum NotAValidLetExpression Source # 
Instance details

Defined in Agda.Interaction.Options.Errors

Generic NotAValidLetExpression Source # 
Instance details

Defined in Agda.Interaction.Options.Errors

Associated Types

type Rep NotAValidLetExpression 
Instance details

Defined in Agda.Interaction.Options.Errors

type Rep NotAValidLetExpression = D1 ('MetaData "NotAValidLetExpression" "Agda.Interaction.Options.Errors" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "MissingBody" 'PrefixI 'False) (U1 :: Type -> Type))
Show NotAValidLetExpression Source # 
Instance details

Defined in Agda.Interaction.Options.Errors

type Rep NotAValidLetExpression Source # 
Instance details

Defined in Agda.Interaction.Options.Errors

type Rep NotAValidLetExpression = D1 ('MetaData "NotAValidLetExpression" "Agda.Interaction.Options.Errors" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "MissingBody" 'PrefixI 'False) (U1 :: Type -> Type))

data NotAllowedInDotPatterns Source #

Things not allowed in dot patterns.

Instances

Instances details
NFData NotAllowedInDotPatterns Source # 
Instance details

Defined in Agda.Interaction.Options.Errors

Methods

rnf :: NotAllowedInDotPatterns -> () #

Bounded NotAllowedInDotPatterns Source # 
Instance details

Defined in Agda.Interaction.Options.Errors

Enum NotAllowedInDotPatterns Source # 
Instance details

Defined in Agda.Interaction.Options.Errors

Generic NotAllowedInDotPatterns Source # 
Instance details

Defined in Agda.Interaction.Options.Errors

Associated Types

type Rep NotAllowedInDotPatterns 
Instance details

Defined in Agda.Interaction.Options.Errors

type Rep NotAllowedInDotPatterns = D1 ('MetaData "NotAllowedInDotPatterns" "Agda.Interaction.Options.Errors" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "LetExpressions" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "PatternLambdas" 'PrefixI 'False) (U1 :: Type -> Type))
Show NotAllowedInDotPatterns Source # 
Instance details

Defined in Agda.Interaction.Options.Errors

type Rep NotAllowedInDotPatterns Source # 
Instance details

Defined in Agda.Interaction.Options.Errors

type Rep NotAllowedInDotPatterns = D1 ('MetaData "NotAllowedInDotPatterns" "Agda.Interaction.Options.Errors" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "LetExpressions" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "PatternLambdas" 'PrefixI 'False) (U1 :: Type -> Type))

data FileId Source #

Unique identifier of a file.

Instances

Instances details
GetFileId FileToId Source # 
Instance details

Defined in Agda.Utils.FileId

GetIdFile IdToFile Source # 
Instance details

Defined in Agda.Utils.FileId

NFData FileId Source # 
Instance details

Defined in Agda.Utils.FileId

Methods

rnf :: FileId -> () #

Enum FileId Source # 
Instance details

Defined in Agda.Utils.FileId

Generic FileId Source # 
Instance details

Defined in Agda.Utils.FileId

Associated Types

type Rep FileId 
Instance details

Defined in Agda.Utils.FileId

type Rep FileId = D1 ('MetaData "FileId" "Agda.Utils.FileId" "Agda-2.8.0-inplace" 'True) (C1 ('MetaCons "FileId" 'PrefixI 'True) (S1 ('MetaSel ('Just "theFileId") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Word32)))

Methods

from :: FileId -> Rep FileId x #

to :: Rep FileId x -> FileId #

Num FileId Source # 
Instance details

Defined in Agda.Utils.FileId

Show FileId Source # 
Instance details

Defined in Agda.Utils.FileId

Eq FileId Source # 
Instance details

Defined in Agda.Utils.FileId

Methods

(==) :: FileId -> FileId -> Bool #

(/=) :: FileId -> FileId -> Bool #

Ord FileId Source # 
Instance details

Defined in Agda.Utils.FileId

type Rep FileId Source # 
Instance details

Defined in Agda.Utils.FileId

type Rep FileId = D1 ('MetaData "FileId" "Agda.Utils.FileId" "Agda-2.8.0-inplace" 'True) (C1 ('MetaCons "FileId" 'PrefixI 'True) (S1 ('MetaSel ('Just "theFileId") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Word32)))

class Monad m => MonadFileId (m :: Type -> Type) where Source #

Minimal complete definition

Nothing

Methods

fileFromId :: FileId -> m File Source #

default fileFromId :: forall (t :: (Type -> Type) -> Type -> Type) (n :: Type -> Type). (MonadTrans t, MonadFileId n, m ~ t n) => FileId -> m File Source #

idFromFile :: File -> m FileId Source #

default idFromFile :: forall (t :: (Type -> Type) -> Type -> Type) (n :: Type -> Type). (MonadTrans t, MonadFileId n, m ~ t n) => File -> m FileId Source #

Instances

Instances details
MonadFileId IM Source # 
Instance details

Defined in Agda.Interaction.Monad

MonadIO m => MonadFileId (TCMT m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.State

MonadFileId m => MonadFileId (ExceptT e m) Source # 
Instance details

Defined in Agda.Utils.FileId

MonadFileId m => MonadFileId (ReaderT r m) Source # 
Instance details

Defined in Agda.Utils.FileId

MonadFileId m => MonadFileId (StateT s m) Source # 
Instance details

Defined in Agda.Utils.FileId

class (Functor m, Applicative m, Monad m) => HasOptions (m :: Type -> Type) where Source #

Minimal complete definition

Nothing

Methods

pragmaOptions :: m PragmaOptions Source #

Returns the pragma options which are currently in effect.

default pragmaOptions :: forall (n :: Type -> Type) (t :: (Type -> Type) -> Type -> Type). (HasOptions n, MonadTrans t, m ~ t n) => m PragmaOptions Source #

commandLineOptions :: m CommandLineOptions Source #

Returns the command line options which are currently in effect.

default commandLineOptions :: forall (n :: Type -> Type) (t :: (Type -> Type) -> Type -> Type). (HasOptions n, MonadTrans t, m ~ t n) => m CommandLineOptions Source #

Instances

Instances details
HasOptions IM Source # 
Instance details

Defined in Agda.Interaction.Monad

HasOptions TerM Source # 
Instance details

Defined in Agda.Termination.Monad

HasOptions ReduceM Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

HasOptions NLM Source # 
Instance details

Defined in Agda.TypeChecking.Rewriting.NonLinMatch

HasOptions m => HasOptions (PureConversionT m) Source # 
Instance details

Defined in Agda.TypeChecking.Conversion.Pure

HasOptions m => HasOptions (BlockT m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

MonadIO m => HasOptions (TCMT m) Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

HasOptions m => HasOptions (NamesT m) Source # 
Instance details

Defined in Agda.TypeChecking.Names

HasOptions m => HasOptions (ListT m) Source # 
Instance details

Defined in Agda.Interaction.Options.HasOptions

HasOptions m => HasOptions (ChangeT m) Source # 
Instance details

Defined in Agda.Interaction.Options.HasOptions

HasOptions m => HasOptions (MaybeT m) Source # 
Instance details

Defined in Agda.Interaction.Options.HasOptions

HasOptions m => HasOptions (ExceptT e m) Source # 
Instance details

Defined in Agda.Interaction.Options.HasOptions

HasOptions m => HasOptions (IdentityT m) Source # 
Instance details

Defined in Agda.Interaction.Options.HasOptions

HasOptions m => HasOptions (ReaderT r m) Source # 
Instance details

Defined in Agda.Interaction.Options.HasOptions

HasOptions m => HasOptions (StateT s m) Source # 
Instance details

Defined in Agda.Interaction.Options.HasOptions

(HasOptions m, Monoid w) => HasOptions (WriterT w m) Source # 
Instance details

Defined in Agda.Interaction.Options.HasOptions

data RecordFieldWarning Source #

Instances

Instances details
EmbPrj RecordFieldWarning Source # 
Instance details

Defined in Agda.TypeChecking.Serialise.Instances.Errors

NFData RecordFieldWarning Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base

Methods

rnf :: RecordFieldWarning -> () #

Generic RecordFieldWarning Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base.Warning

Show RecordFieldWarning Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base.Warning

type Rep RecordFieldWarning Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base.Warning

data UselessPublicReason Source #

Instances

Instances details
EmbPrj UselessPublicReason Source # 
Instance details

Defined in Agda.TypeChecking.Serialise.Instances.Errors

NFData UselessPublicReason Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base.Warning

Methods

rnf :: UselessPublicReason -> () #

Bounded UselessPublicReason Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base.Warning

Enum UselessPublicReason Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base.Warning

Generic UselessPublicReason Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base.Warning

Associated Types

type Rep UselessPublicReason 
Instance details

Defined in Agda.TypeChecking.Monad.Base.Warning

type Rep UselessPublicReason = D1 ('MetaData "UselessPublicReason" "Agda.TypeChecking.Monad.Base.Warning" "Agda-2.8.0-inplace" 'False) ((C1 ('MetaCons "UselessPublicPreamble" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "UselessPublicLet" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "UselessPublicNoOpen" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "UselessPublicAnonymousModule" 'PrefixI 'False) (U1 :: Type -> Type)))
Show UselessPublicReason Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base.Warning

type Rep UselessPublicReason Source # 
Instance details

Defined in Agda.TypeChecking.Monad.Base.Warning

type Rep UselessPublicReason = D1 ('MetaData "UselessPublicReason" "Agda.TypeChecking.Monad.Base.Warning" "Agda-2.8.0-inplace" 'False) ((C1 ('MetaCons "UselessPublicPreamble" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "UselessPublicLet" 'PrefixI 'False) (U1 :: Type -> Type)) :+: (C1 ('MetaCons "UselessPublicNoOpen" 'PrefixI 'False) (U1 :: Type -> Type) :+: C1 ('MetaCons "UselessPublicAnonymousModule" 'PrefixI 'False) (U1 :: Type -> Type)))

Orphan instances

Pretty Comparison Source # 
Instance details

Pretty ModuleToSource Source # 
Instance details

Pretty NamedMeta Source # 
Instance details

Pretty Polarity Source # 
Instance details

KillRange Polarity Source # 
Instance details

GetFileId FileDictWithBuiltins Source # 
Instance details

GetIdFile FileDictWithBuiltins Source # 
Instance details

NFData Comparison Source # 
Instance details

Methods

rnf :: Comparison -> () #

NFData HighlightingLevel Source # 
Instance details

Methods

rnf :: HighlightingLevel -> () #

NFData HighlightingMethod Source # 
Instance details

Methods

rnf :: HighlightingMethod -> () #

NFData Polarity Source # 
Instance details

Methods

rnf :: Polarity -> () #

NFData RecordFieldWarning Source # 
Instance details

Methods

rnf :: RecordFieldWarning -> () #

NFData (BiMap RawTopLevelModuleName ModuleNameHash) Source # 
Instance details