{-# LANGUAGE CPP #-}
module Agda.TypeChecking.Monad.Base
( module Agda.TypeChecking.Monad.Base
, module Agda.TypeChecking.Monad.Base.Types
, module X
, HasOptions (..)
, RecordFieldWarning
) where
import Prelude hiding (null)
import Control.Applicative hiding (empty)
import Control.Arrow ( (&&&) )
import Control.Concurrent ( forkIO )
import Control.DeepSeq
import qualified Control.Exception as E
import Control.Monad.Except ( MonadError(..), ExceptT(..), runExceptT )
import Control.Monad.IO.Class ( MonadIO(..) )
import Control.Monad.State ( MonadState(..), modify, StateT(..), runStateT )
import Control.Monad.Reader ( MonadReader(..), ReaderT(..), runReaderT )
import Control.Monad.Writer ( WriterT(..), runWriterT )
import Control.Monad.Trans ( MonadTrans(..), lift )
import Control.Monad.Trans.Control ( MonadTransControl(..), liftThrough )
import Control.Monad.Trans.Identity ( IdentityT(..), runIdentityT )
import Control.Monad.Trans.Maybe ( MaybeT(..) )
import Control.Parallel ( pseq )
import Data.Array (Ix)
import Data.Function (on)
import Data.Word (Word32)
import Data.IntMap (IntMap)
import qualified Data.IntMap as IntMap
import Data.IntSet (IntSet)
import qualified Data.IntSet as IntSet
import qualified Data.List as List
import Data.Maybe
import Data.Map (Map)
import qualified Data.Map as Map
import Data.Sequence (Seq)
import Data.Set (Set, toList, fromList)
import qualified Data.Set as Set
import Data.HashMap.Strict (HashMap)
import qualified Data.HashMap.Strict as HMap
import qualified Data.HashSet as HashSet
import Data.Hashable
import Data.HashSet (HashSet)
import Data.Semigroup ( Semigroup, (<>))
import Data.Set (Set)
import qualified Data.Set as Set
import Data.String
import Data.Text (Text)
import qualified Data.Text.Lazy as TL
import Data.IORef
import GHC.Generics (Generic)
import System.IO (hFlush, stdout)
import Agda.Benchmarking (Benchmark, Phase)
import {-# SOURCE #-} Agda.Compiler.Treeless.Pretty ()
import Agda.Syntax.Common
import Agda.Syntax.Builtin (SomeBuiltin, BuiltinId, PrimitiveId)
import qualified Agda.Syntax.Concrete as C
import Agda.Syntax.Concrete.Definitions
(NiceDeclaration, DeclarationWarning, declarationWarningName)
import Agda.Syntax.Concrete.Definitions.Errors
(DeclarationException')
import qualified Agda.Syntax.Abstract as A
import Agda.Syntax.Internal as I
import Agda.Syntax.Internal.MetaVars
import Agda.Syntax.Internal.Generic (TermLike(..))
import Agda.Syntax.Parser.Monad (ParseError, ParseWarning, parseWarningName)
import Agda.Syntax.TopLevelModuleName
(RawTopLevelModuleName, TopLevelModuleName)
import Agda.Syntax.Treeless (Compiled)
import Agda.Syntax.Notation
import Agda.Syntax.Position
import Agda.Syntax.Scope.Base
import Agda.Syntax.Info ( MetaKind(InstanceMeta, UnificationMeta), MetaNameSuggestion, MutualInfo )
import Agda.TypeChecking.Monad.Base.Types
import qualified Agda.TypeChecking.Monad.Base.Warning as W
import Agda.TypeChecking.Monad.Base.Warning (RecordFieldWarning)
import Agda.TypeChecking.SizedTypes.Syntax (HypSizeConstraint)
import Agda.TypeChecking.CompiledClause
import Agda.TypeChecking.Coverage.SplitTree
import Agda.TypeChecking.Positivity.Occurrence
import Agda.TypeChecking.Free.Lazy (Free(freeVars'), underBinder', underBinder)
import Agda.TypeChecking.DiscrimTree.Types
import Agda.Compiler.Backend.Base (Backend_boot, Backend'_boot)
import Agda.Interaction.Options
import qualified Agda.Interaction.Options.Errors as ErrorName
import Agda.Interaction.Options.Errors as X
( CannotQuoteTerm(..)
, ErasedDatatypeReason(..)
, NotAValidLetBinding(..)
, NotAValidLetExpression(..)
, NotAllowedInDotPatterns(..)
)
import Agda.Interaction.Options.Warnings
import Agda.Interaction.Response.Base (Response_boot(..))
import Agda.Interaction.Highlighting.Precise
(HighlightingInfo, NameKind)
import Agda.Interaction.Library
import Agda.Interaction.Library.Base ( ExeName, ExeMap, LibErrors )
import Agda.Utils.Benchmark (MonadBench(..))
import Agda.Utils.BiMap (BiMap, HasTag(..))
import qualified Agda.Utils.BiMap as BiMap
import Agda.Utils.Boolean ( fromBool, toBool )
import Agda.Utils.CallStack ( CallStack, HasCallStack, withCallerCallStack )
import Agda.Utils.FileName
import Agda.Utils.Functor
import Agda.Utils.Hash
import Agda.Utils.IO ( CatchIO, catchIO, showIOException )
import Agda.Utils.Lens
import Agda.Utils.List
import Agda.Utils.ListT
import Agda.Utils.List1 (List1, pattern (:|))
import Agda.Utils.List2 (List2, pattern List2)
import qualified Agda.Utils.List1 as List1
import qualified Agda.Utils.Maybe.Strict as Strict
import Agda.Utils.Monad
import Agda.Utils.Null
import Agda.Utils.Permutation
import Agda.Syntax.Common.Pretty
import Agda.Utils.SmallSet (SmallSet, SmallSetElement)
import qualified Agda.Utils.SmallSet as SmallSet
import Agda.Utils.Set1 (Set1)
import Agda.Utils.Singleton
import Agda.Utils.Tuple (Pair)
import Agda.Utils.Update
import Agda.Utils.Impossible
data TCState = TCSt
{ TCState -> PreScopeState
stPreScopeState :: !PreScopeState
, TCState -> PostScopeState
stPostScopeState :: !PostScopeState
, TCState -> PersistentTCState
stPersistentState :: !PersistentTCState
}
deriving (forall x. TCState -> Rep TCState x)
-> (forall x. Rep TCState x -> TCState) -> Generic TCState
forall x. Rep TCState x -> TCState
forall x. TCState -> Rep TCState x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. TCState -> Rep TCState x
from :: forall x. TCState -> Rep TCState x
$cto :: forall x. Rep TCState x -> TCState
to :: forall x. Rep TCState x -> TCState
Generic
class Monad m => ReadTCState m where
getTCState :: m TCState
locallyTCState :: Lens' TCState a -> (a -> a) -> m b -> m b
withTCState :: (TCState -> TCState) -> m a -> m a
withTCState = Lens' TCState TCState -> (TCState -> TCState) -> m a -> m a
forall a b. Lens' TCState a -> (a -> a) -> m b -> m b
forall (m :: * -> *) a b.
ReadTCState m =>
Lens' TCState a -> (a -> a) -> m b -> m b
locallyTCState (TCState -> f TCState) -> TCState -> f TCState
forall a. a -> a
Lens' TCState TCState
id
default getTCState :: (MonadTrans t, ReadTCState n, t n ~ m) => m TCState
getTCState = n TCState -> t n TCState
forall (m :: * -> *) a. Monad m => m a -> t m a
forall (t :: (* -> *) -> * -> *) (m :: * -> *) a.
(MonadTrans t, Monad m) =>
m a -> t m a
lift n TCState
forall (m :: * -> *). ReadTCState m => m TCState
getTCState
default locallyTCState
:: (MonadTransControl t, ReadTCState n, t n ~ m)
=> Lens' TCState a -> (a -> a) -> m b -> m b
locallyTCState Lens' TCState a
l = (n (StT t b) -> n (StT t b)) -> m b -> m b
(n (StT t b) -> n (StT t b)) -> t n b -> t n b
forall (t :: (* -> *) -> * -> *) (m :: * -> *) a b.
(MonadTransControl t, Monad (t m), Monad m) =>
(m (StT t a) -> m (StT t b)) -> t m a -> t m b
liftThrough ((n (StT t b) -> n (StT t b)) -> m b -> m b)
-> ((a -> a) -> n (StT t b) -> n (StT t b))
-> (a -> a)
-> m b
-> m b
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Lens' TCState a -> (a -> a) -> n (StT t b) -> n (StT t b)
forall a b. Lens' TCState a -> (a -> a) -> n b -> n b
forall (m :: * -> *) a b.
ReadTCState m =>
Lens' TCState a -> (a -> a) -> m b -> m b
locallyTCState (a -> f a) -> TCState -> f TCState
Lens' TCState a
l
instance ReadTCState m => ReadTCState (ListT m) where
locallyTCState :: forall a b. Lens' TCState a -> (a -> a) -> ListT m b -> ListT m b
locallyTCState Lens' TCState a
l = (m (Maybe (b, ListT m b)) -> m (Maybe (b, ListT m b)))
-> ListT m b -> ListT m b
forall (m :: * -> *) a (n :: * -> *) b.
(m (Maybe (a, ListT m a)) -> n (Maybe (b, ListT n b)))
-> ListT m a -> ListT n b
mapListT ((m (Maybe (b, ListT m b)) -> m (Maybe (b, ListT m b)))
-> ListT m b -> ListT m b)
-> ((a -> a)
-> m (Maybe (b, ListT m b)) -> m (Maybe (b, ListT m b)))
-> (a -> a)
-> ListT m b
-> ListT m b
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Lens' TCState a
-> (a -> a) -> m (Maybe (b, ListT m b)) -> m (Maybe (b, ListT m b))
forall a b. Lens' TCState a -> (a -> a) -> m b -> m b
forall (m :: * -> *) a b.
ReadTCState m =>
Lens' TCState a -> (a -> a) -> m b -> m b
locallyTCState (a -> f a) -> TCState -> f TCState
Lens' TCState a
l
instance ReadTCState m => ReadTCState (ChangeT m)
instance ReadTCState m => ReadTCState (ExceptT err m)
instance ReadTCState m => ReadTCState (IdentityT m)
instance ReadTCState m => ReadTCState (MaybeT m)
instance ReadTCState m => ReadTCState (ReaderT r m)
instance ReadTCState m => ReadTCState (StateT s m)
instance (Monoid w, ReadTCState m) => ReadTCState (WriterT w m)
instance Show TCState where
show :: TCState -> String
show TCState
_ = String
"TCSt{}"
type Backend = Backend_boot Definition TCM
type Backend' opts env menv mod def = Backend'_boot Definition TCM opts env menv mod def
type BackendForeignCode = Map BackendName ForeignCodeStack
type ImportedModules = Set TopLevelModuleName
type UserWarnings = Map QName Text
data PreScopeState = PreScopeState
{ PreScopeState -> HighlightingInfo
stPreTokens :: !HighlightingInfo
, PreScopeState -> Signature
stPreImports :: !Signature
, PreScopeState -> ImportedModules
stPreImportedModules :: !ImportedModules
, PreScopeState -> ModuleToSource
stPreModuleToSource :: !ModuleToSource
, PreScopeState -> VisitedModules
stPreVisitedModules :: !VisitedModules
, PreScopeState -> ScopeInfo
stPreScope :: !ScopeInfo
, PreScopeState -> PatternSynDefns
stPrePatternSyns :: !A.PatternSynDefns
, PreScopeState -> PatternSynDefns
stPrePatternSynImports :: !A.PatternSynDefns
, PreScopeState -> Maybe (Set QName)
stPreGeneralizedVars :: !(Strict.Maybe (Set QName))
, PreScopeState -> PragmaOptions
stPrePragmaOptions :: !PragmaOptions
, PreScopeState -> BuiltinThings
stPreImportedBuiltins :: !BuiltinThings
, PreScopeState -> DisplayForms
stPreImportedDisplayForms :: !DisplayForms
, PreScopeState -> InteractionId
stPreFreshInteractionId :: !InteractionId
, PreScopeState -> UserWarnings
stPreImportedUserWarnings :: !UserWarnings
, PreScopeState -> UserWarnings
stPreLocalUserWarnings :: !UserWarnings
, PreScopeState -> Maybe BackendName
stPreWarningOnImport :: !(Strict.Maybe Text)
, PreScopeState -> Set QName
stPreImportedPartialDefs :: !(Set QName)
, PreScopeState -> Map String ProjectConfig
stPreProjectConfigs :: !(Map FilePath ProjectConfig)
, PreScopeState -> Map String AgdaLibFile
stPreAgdaLibFiles :: !(Map FilePath AgdaLibFile)
, PreScopeState -> RemoteMetaStore
stPreImportedMetaStore :: !RemoteMetaStore
, PreScopeState -> HashMap QName QName
stPreCopiedNames :: !(HashMap A.QName A.QName)
, PreScopeState -> HashMap QName (HashSet QName)
stPreNameCopies :: !(HashMap A.QName (HashSet A.QName))
}
deriving (forall x. PreScopeState -> Rep PreScopeState x)
-> (forall x. Rep PreScopeState x -> PreScopeState)
-> Generic PreScopeState
forall x. Rep PreScopeState x -> PreScopeState
forall x. PreScopeState -> Rep PreScopeState x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. PreScopeState -> Rep PreScopeState x
from :: forall x. PreScopeState -> Rep PreScopeState x
$cto :: forall x. Rep PreScopeState x -> PreScopeState
to :: forall x. Rep PreScopeState x -> PreScopeState
Generic
data DisambiguatedName = DisambiguatedName NameKind A.QName
deriving (forall x. DisambiguatedName -> Rep DisambiguatedName x)
-> (forall x. Rep DisambiguatedName x -> DisambiguatedName)
-> Generic DisambiguatedName
forall x. Rep DisambiguatedName x -> DisambiguatedName
forall x. DisambiguatedName -> Rep DisambiguatedName x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. DisambiguatedName -> Rep DisambiguatedName x
from :: forall x. DisambiguatedName -> Rep DisambiguatedName x
$cto :: forall x. Rep DisambiguatedName x -> DisambiguatedName
to :: forall x. Rep DisambiguatedName x -> DisambiguatedName
Generic
type DisambiguatedNames = IntMap DisambiguatedName
type ConcreteNames = Map Name (List1 C.Name)
type ShadowingNames = Map Name (Set1 RawName)
type UsedNames = Map RawName (Set1 RawName)
data PostScopeState = PostScopeState
{ PostScopeState -> HighlightingInfo
stPostSyntaxInfo :: !HighlightingInfo
, PostScopeState -> DisambiguatedNames
stPostDisambiguatedNames :: !DisambiguatedNames
, PostScopeState -> LocalMetaStore
stPostOpenMetaStore :: !LocalMetaStore
, PostScopeState -> LocalMetaStore
stPostSolvedMetaStore :: !LocalMetaStore
, PostScopeState -> InteractionPoints
stPostInteractionPoints :: !InteractionPoints
, PostScopeState -> Constraints
stPostAwakeConstraints :: !Constraints
, PostScopeState -> Constraints
stPostSleepingConstraints :: !Constraints
, PostScopeState -> Bool
stPostDirty :: !Bool
, PostScopeState -> Set QName
stPostOccursCheckDefs :: !(Set QName)
, PostScopeState -> Signature
stPostSignature :: !Signature
, PostScopeState -> Map ModuleName CheckpointId
stPostModuleCheckpoints :: !(Map ModuleName CheckpointId)
, PostScopeState -> DisplayForms
stPostImportsDisplayForms :: !DisplayForms
, PostScopeState -> BackendForeignCode
stPostForeignCode :: !BackendForeignCode
, PostScopeState -> Maybe (ModuleName, TopLevelModuleName)
stPostCurrentModule ::
!(Maybe (ModuleName, TopLevelModuleName))
, PostScopeState -> Set QName
stPostPendingInstances :: !(Set QName)
, PostScopeState -> Set QName
stPostTemporaryInstances :: !(Set QName)
, PostScopeState -> ConcreteNames
stPostConcreteNames :: !ConcreteNames
, PostScopeState -> UsedNames
stPostUsedNames :: !UsedNames
, PostScopeState -> ShadowingNames
stPostShadowingNames :: !ShadowingNames
, PostScopeState -> Statistics
stPostStatistics :: !Statistics
, PostScopeState -> Set TCWarning
stPostTCWarnings :: !(Set TCWarning)
, PostScopeState -> MutualBlocks
stPostMutualBlocks :: !MutualBlocks
, PostScopeState -> BuiltinThings
stPostLocalBuiltins :: !BuiltinThings
, PostScopeState -> MetaId
stPostFreshMetaId :: !MetaId
, PostScopeState -> MutualId
stPostFreshMutualId :: !MutualId
, PostScopeState -> ProblemId
stPostFreshProblemId :: !ProblemId
, PostScopeState -> CheckpointId
stPostFreshCheckpointId :: !CheckpointId
, PostScopeState -> Int
stPostFreshInt :: !Int
, PostScopeState -> NameId
stPostFreshNameId :: !NameId
, PostScopeState -> OpaqueId
stPostFreshOpaqueId :: !OpaqueId
, PostScopeState -> Bool
stPostAreWeCaching :: !Bool
, PostScopeState -> Bool
stPostPostponeInstanceSearch :: !Bool
, PostScopeState -> Bool
stPostConsideringInstance :: !Bool
, PostScopeState -> Bool
stPostInstantiateBlocking :: !Bool
, PostScopeState -> Set QName
stPostLocalPartialDefs :: !(Set QName)
, PostScopeState -> Map OpaqueId OpaqueBlock
stPostOpaqueBlocks :: Map OpaqueId OpaqueBlock
, PostScopeState -> Map QName OpaqueId
stPostOpaqueIds :: Map QName OpaqueId
}
deriving ((forall x. PostScopeState -> Rep PostScopeState x)
-> (forall x. Rep PostScopeState x -> PostScopeState)
-> Generic PostScopeState
forall x. Rep PostScopeState x -> PostScopeState
forall x. PostScopeState -> Rep PostScopeState x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. PostScopeState -> Rep PostScopeState x
from :: forall x. PostScopeState -> Rep PostScopeState x
$cto :: forall x. Rep PostScopeState x -> PostScopeState
to :: forall x. Rep PostScopeState x -> PostScopeState
Generic)
data PersistentTCState = PersistentTCSt
{ PersistentTCState -> VisitedModules
stDecodedModules :: !DecodedModules
, PersistentTCState -> BiMap RawTopLevelModuleName ModuleNameHash
stPersistentTopLevelModuleNames ::
!(BiMap RawTopLevelModuleName ModuleNameHash)
, PersistentTCState -> CommandLineOptions
stPersistentOptions :: CommandLineOptions
, PersistentTCState -> InteractionOutputCallback
stInteractionOutputCallback :: InteractionOutputCallback
, PersistentTCState -> Benchmark
stBenchmark :: !Benchmark
, PersistentTCState -> Statistics
stAccumStatistics :: !Statistics
, PersistentTCState -> Maybe LoadedFileCache
stPersistLoadedFileCache :: !(Strict.Maybe LoadedFileCache)
, PersistentTCState -> [Backend]
stPersistBackends :: [Backend]
}
deriving (forall x. PersistentTCState -> Rep PersistentTCState x)
-> (forall x. Rep PersistentTCState x -> PersistentTCState)
-> Generic PersistentTCState
forall x. Rep PersistentTCState x -> PersistentTCState
forall x. PersistentTCState -> Rep PersistentTCState x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. PersistentTCState -> Rep PersistentTCState x
from :: forall x. PersistentTCState -> Rep PersistentTCState x
$cto :: forall x. Rep PersistentTCState x -> PersistentTCState
to :: forall x. Rep PersistentTCState x -> PersistentTCState
Generic
data LoadedFileCache = LoadedFileCache
{ LoadedFileCache -> CachedTypeCheckLog
lfcCached :: !CachedTypeCheckLog
, LoadedFileCache -> CachedTypeCheckLog
lfcCurrent :: !CurrentTypeCheckLog
}
deriving (forall x. LoadedFileCache -> Rep LoadedFileCache x)
-> (forall x. Rep LoadedFileCache x -> LoadedFileCache)
-> Generic LoadedFileCache
forall x. Rep LoadedFileCache x -> LoadedFileCache
forall x. LoadedFileCache -> Rep LoadedFileCache x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. LoadedFileCache -> Rep LoadedFileCache x
from :: forall x. LoadedFileCache -> Rep LoadedFileCache x
$cto :: forall x. Rep LoadedFileCache x -> LoadedFileCache
to :: forall x. Rep LoadedFileCache x -> LoadedFileCache
Generic
type CachedTypeCheckLog = [(TypeCheckAction, PostScopeState)]
type CurrentTypeCheckLog = [(TypeCheckAction, PostScopeState)]
data TypeCheckAction
= EnterSection !Erased !ModuleName !A.Telescope
| LeaveSection !ModuleName
| Decl !A.Declaration
| Pragmas !PragmaOptions
deriving ((forall x. TypeCheckAction -> Rep TypeCheckAction x)
-> (forall x. Rep TypeCheckAction x -> TypeCheckAction)
-> Generic TypeCheckAction
forall x. Rep TypeCheckAction x -> TypeCheckAction
forall x. TypeCheckAction -> Rep TypeCheckAction x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. TypeCheckAction -> Rep TypeCheckAction x
from :: forall x. TypeCheckAction -> Rep TypeCheckAction x
$cto :: forall x. Rep TypeCheckAction x -> TypeCheckAction
to :: forall x. Rep TypeCheckAction x -> TypeCheckAction
Generic)
initPersistentState :: PersistentTCState
initPersistentState :: PersistentTCState
initPersistentState = PersistentTCSt
{ stPersistentOptions :: CommandLineOptions
stPersistentOptions = CommandLineOptions
defaultOptions
, stPersistentTopLevelModuleNames :: BiMap RawTopLevelModuleName ModuleNameHash
stPersistentTopLevelModuleNames = BiMap RawTopLevelModuleName ModuleNameHash
forall a. Null a => a
empty
, stDecodedModules :: VisitedModules
stDecodedModules = VisitedModules
forall k a. Map k a
Map.empty
, stInteractionOutputCallback :: InteractionOutputCallback
stInteractionOutputCallback = InteractionOutputCallback
defaultInteractionOutputCallback
, stBenchmark :: Benchmark
stBenchmark = Benchmark
forall a. Null a => a
empty
, stAccumStatistics :: Statistics
stAccumStatistics = Statistics
forall k a. Map k a
Map.empty
, stPersistLoadedFileCache :: Maybe LoadedFileCache
stPersistLoadedFileCache = Maybe LoadedFileCache
forall a. Null a => a
empty
, stPersistBackends :: [Backend]
stPersistBackends = []
}
initialMetaId :: MetaId
initialMetaId :: MetaId
initialMetaId = MetaId
{ metaId :: Word64
metaId = Word64
0
, metaModule :: ModuleNameHash
metaModule = ModuleNameHash
noModuleNameHash
}
initPreScopeState :: PreScopeState
initPreScopeState :: PreScopeState
initPreScopeState = PreScopeState
{ stPreTokens :: HighlightingInfo
stPreTokens = HighlightingInfo
forall a. Monoid a => a
mempty
, stPreImports :: Signature
stPreImports = Signature
emptySignature
, stPreImportedModules :: ImportedModules
stPreImportedModules = ImportedModules
forall a. Null a => a
empty
, stPreModuleToSource :: ModuleToSource
stPreModuleToSource = ModuleToSource
forall k a. Map k a
Map.empty
, stPreVisitedModules :: VisitedModules
stPreVisitedModules = VisitedModules
forall k a. Map k a
Map.empty
, stPreScope :: ScopeInfo
stPreScope = ScopeInfo
emptyScopeInfo
, stPrePatternSyns :: PatternSynDefns
stPrePatternSyns = PatternSynDefns
forall k a. Map k a
Map.empty
, stPrePatternSynImports :: PatternSynDefns
stPrePatternSynImports = PatternSynDefns
forall k a. Map k a
Map.empty
, stPreGeneralizedVars :: Maybe (Set QName)
stPreGeneralizedVars = Maybe (Set QName)
forall a. Monoid a => a
mempty
, stPrePragmaOptions :: PragmaOptions
stPrePragmaOptions = PragmaOptions
defaultInteractionOptions
, stPreImportedBuiltins :: BuiltinThings
stPreImportedBuiltins = BuiltinThings
forall k a. Map k a
Map.empty
, stPreImportedDisplayForms :: DisplayForms
stPreImportedDisplayForms = DisplayForms
forall k v. HashMap k v
HMap.empty
, stPreFreshInteractionId :: InteractionId
stPreFreshInteractionId = InteractionId
0
, stPreImportedUserWarnings :: UserWarnings
stPreImportedUserWarnings = UserWarnings
forall k a. Map k a
Map.empty
, stPreLocalUserWarnings :: UserWarnings
stPreLocalUserWarnings = UserWarnings
forall k a. Map k a
Map.empty
, stPreWarningOnImport :: Maybe BackendName
stPreWarningOnImport = Maybe BackendName
forall a. Null a => a
empty
, stPreImportedPartialDefs :: Set QName
stPreImportedPartialDefs = Set QName
forall a. Set a
Set.empty
, stPreProjectConfigs :: Map String ProjectConfig
stPreProjectConfigs = Map String ProjectConfig
forall k a. Map k a
Map.empty
, stPreAgdaLibFiles :: Map String AgdaLibFile
stPreAgdaLibFiles = Map String AgdaLibFile
forall k a. Map k a
Map.empty
, stPreImportedMetaStore :: RemoteMetaStore
stPreImportedMetaStore = RemoteMetaStore
forall k v. HashMap k v
HMap.empty
, stPreCopiedNames :: HashMap QName QName
stPreCopiedNames = HashMap QName QName
forall k v. HashMap k v
HMap.empty
, stPreNameCopies :: HashMap QName (HashSet QName)
stPreNameCopies = HashMap QName (HashSet QName)
forall k v. HashMap k v
HMap.empty
}
initPostScopeState :: PostScopeState
initPostScopeState :: PostScopeState
initPostScopeState = PostScopeState
{ stPostSyntaxInfo :: HighlightingInfo
stPostSyntaxInfo = HighlightingInfo
forall a. Monoid a => a
mempty
, stPostDisambiguatedNames :: DisambiguatedNames
stPostDisambiguatedNames = DisambiguatedNames
forall a. IntMap a
IntMap.empty
, stPostOpenMetaStore :: LocalMetaStore
stPostOpenMetaStore = LocalMetaStore
forall k a. Map k a
Map.empty
, stPostSolvedMetaStore :: LocalMetaStore
stPostSolvedMetaStore = LocalMetaStore
forall k a. Map k a
Map.empty
, stPostInteractionPoints :: InteractionPoints
stPostInteractionPoints = InteractionPoints
forall a. Null a => a
empty
, stPostAwakeConstraints :: Constraints
stPostAwakeConstraints = []
, stPostSleepingConstraints :: Constraints
stPostSleepingConstraints = []
, stPostDirty :: Bool
stPostDirty = Bool
False
, stPostOccursCheckDefs :: Set QName
stPostOccursCheckDefs = Set QName
forall a. Set a
Set.empty
, stPostSignature :: Signature
stPostSignature = Signature
emptySignature
, stPostModuleCheckpoints :: Map ModuleName CheckpointId
stPostModuleCheckpoints = Map ModuleName CheckpointId
forall k a. Map k a
Map.empty
, stPostImportsDisplayForms :: DisplayForms
stPostImportsDisplayForms = DisplayForms
forall k v. HashMap k v
HMap.empty
, stPostCurrentModule :: Maybe (ModuleName, TopLevelModuleName)
stPostCurrentModule = Maybe (ModuleName, TopLevelModuleName)
forall a. Null a => a
empty
, stPostPendingInstances :: Set QName
stPostPendingInstances = Set QName
forall a. Set a
Set.empty
, stPostTemporaryInstances :: Set QName
stPostTemporaryInstances = Set QName
forall a. Set a
Set.empty
, stPostConcreteNames :: ConcreteNames
stPostConcreteNames = ConcreteNames
forall k a. Map k a
Map.empty
, stPostUsedNames :: UsedNames
stPostUsedNames = UsedNames
forall k a. Map k a
Map.empty
, stPostShadowingNames :: ShadowingNames
stPostShadowingNames = ShadowingNames
forall k a. Map k a
Map.empty
, stPostStatistics :: Statistics
stPostStatistics = Statistics
forall k a. Map k a
Map.empty
, stPostTCWarnings :: Set TCWarning
stPostTCWarnings = Set TCWarning
forall a. Null a => a
empty
, stPostMutualBlocks :: MutualBlocks
stPostMutualBlocks = MutualBlocks
forall a. Null a => a
empty
, stPostLocalBuiltins :: BuiltinThings
stPostLocalBuiltins = BuiltinThings
forall k a. Map k a
Map.empty
, stPostFreshMetaId :: MetaId
stPostFreshMetaId = MetaId
initialMetaId
, stPostFreshMutualId :: MutualId
stPostFreshMutualId = MutualId
0
, stPostFreshProblemId :: ProblemId
stPostFreshProblemId = ProblemId
1
, stPostFreshCheckpointId :: CheckpointId
stPostFreshCheckpointId = CheckpointId
1
, stPostFreshInt :: Int
stPostFreshInt = Int
0
, stPostFreshNameId :: NameId
stPostFreshNameId = Word64 -> ModuleNameHash -> NameId
NameId Word64
0 ModuleNameHash
noModuleNameHash
, stPostFreshOpaqueId :: OpaqueId
stPostFreshOpaqueId = Word64 -> ModuleNameHash -> OpaqueId
OpaqueId Word64
0 ModuleNameHash
noModuleNameHash
, stPostAreWeCaching :: Bool
stPostAreWeCaching = Bool
False
, stPostPostponeInstanceSearch :: Bool
stPostPostponeInstanceSearch = Bool
False
, stPostConsideringInstance :: Bool
stPostConsideringInstance = Bool
False
, stPostInstantiateBlocking :: Bool
stPostInstantiateBlocking = Bool
False
, stPostLocalPartialDefs :: Set QName
stPostLocalPartialDefs = Set QName
forall a. Set a
Set.empty
, stPostOpaqueBlocks :: Map OpaqueId OpaqueBlock
stPostOpaqueBlocks = Map OpaqueId OpaqueBlock
forall k a. Map k a
Map.empty
, stPostOpaqueIds :: Map QName OpaqueId
stPostOpaqueIds = Map QName OpaqueId
forall k a. Map k a
Map.empty
, stPostForeignCode :: BackendForeignCode
stPostForeignCode = BackendForeignCode
forall k a. Map k a
Map.empty
}
initState :: TCState
initState :: TCState
initState = TCSt
{ stPersistentState :: PersistentTCState
stPersistentState = PersistentTCState
initPersistentState
, stPreScopeState :: PreScopeState
stPreScopeState = PreScopeState
initPreScopeState
, stPostScopeState :: PostScopeState
stPostScopeState = PostScopeState
initPostScopeState
}
lensPersistentState :: Lens' TCState PersistentTCState
lensPersistentState :: Lens' TCState PersistentTCState
lensPersistentState PersistentTCState -> f PersistentTCState
f TCState
s = PersistentTCState -> f PersistentTCState
f (TCState -> PersistentTCState
stPersistentState TCState
s) f PersistentTCState -> (PersistentTCState -> TCState) -> f TCState
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ PersistentTCState
x -> TCState
s { stPersistentState = x }
lensPreScopeState :: Lens' TCState PreScopeState
lensPreScopeState :: Lens' TCState PreScopeState
lensPreScopeState PreScopeState -> f PreScopeState
f TCState
s = PreScopeState -> f PreScopeState
f (TCState -> PreScopeState
stPreScopeState TCState
s) f PreScopeState -> (PreScopeState -> TCState) -> f TCState
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ PreScopeState
x -> TCState
s { stPreScopeState = x }
lensPostScopeState :: Lens' TCState PostScopeState
lensPostScopeState :: Lens' TCState PostScopeState
lensPostScopeState PostScopeState -> f PostScopeState
f TCState
s = PostScopeState -> f PostScopeState
f (TCState -> PostScopeState
stPostScopeState TCState
s) f PostScopeState -> (PostScopeState -> TCState) -> f TCState
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ PostScopeState
x -> TCState
s { stPostScopeState = x }
lensLoadedFileCache :: Lens' PersistentTCState (Strict.Maybe LoadedFileCache)
lensLoadedFileCache :: Lens' PersistentTCState (Maybe LoadedFileCache)
lensLoadedFileCache Maybe LoadedFileCache -> f (Maybe LoadedFileCache)
f PersistentTCState
s = Maybe LoadedFileCache -> f (Maybe LoadedFileCache)
f (PersistentTCState -> Maybe LoadedFileCache
stPersistLoadedFileCache PersistentTCState
s) f (Maybe LoadedFileCache)
-> (Maybe LoadedFileCache -> PersistentTCState)
-> f PersistentTCState
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ Maybe LoadedFileCache
x -> PersistentTCState
s { stPersistLoadedFileCache = x }
lensBackends :: Lens' PersistentTCState [Backend]
lensBackends :: Lens' PersistentTCState [Backend]
lensBackends [Backend] -> f [Backend]
f PersistentTCState
s = [Backend] -> f [Backend]
f (PersistentTCState -> [Backend]
stPersistBackends PersistentTCState
s) f [Backend]
-> ([Backend] -> PersistentTCState) -> f PersistentTCState
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ [Backend]
x -> PersistentTCState
s { stPersistBackends = x }
lensTopLevelModuleNames :: Lens' PersistentTCState (BiMap RawTopLevelModuleName ModuleNameHash)
lensTopLevelModuleNames :: Lens'
PersistentTCState (BiMap RawTopLevelModuleName ModuleNameHash)
lensTopLevelModuleNames BiMap RawTopLevelModuleName ModuleNameHash
-> f (BiMap RawTopLevelModuleName ModuleNameHash)
f PersistentTCState
s =
BiMap RawTopLevelModuleName ModuleNameHash
-> f (BiMap RawTopLevelModuleName ModuleNameHash)
f (PersistentTCState -> BiMap RawTopLevelModuleName ModuleNameHash
stPersistentTopLevelModuleNames PersistentTCState
s) f (BiMap RawTopLevelModuleName ModuleNameHash)
-> (BiMap RawTopLevelModuleName ModuleNameHash
-> PersistentTCState)
-> f PersistentTCState
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ BiMap RawTopLevelModuleName ModuleNameHash
x -> PersistentTCState
s { stPersistentTopLevelModuleNames = x }
lensPreTokens :: Lens' PreScopeState HighlightingInfo
lensPreTokens :: Lens' PreScopeState HighlightingInfo
lensPreTokens HighlightingInfo -> f HighlightingInfo
f PreScopeState
s = HighlightingInfo -> f HighlightingInfo
f (PreScopeState -> HighlightingInfo
stPreTokens PreScopeState
s) f HighlightingInfo
-> (HighlightingInfo -> PreScopeState) -> f PreScopeState
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ HighlightingInfo
x -> PreScopeState
s { stPreTokens = x }
lensImports :: Lens' PreScopeState Signature
lensImports :: Lens' PreScopeState Signature
lensImports Signature -> f Signature
f PreScopeState
s = Signature -> f Signature
f (PreScopeState -> Signature
stPreImports PreScopeState
s) f Signature -> (Signature -> PreScopeState) -> f PreScopeState
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ Signature
x -> PreScopeState
s { stPreImports = x }
lensImportedModules :: Lens' PreScopeState ImportedModules
lensImportedModules :: Lens' PreScopeState ImportedModules
lensImportedModules ImportedModules -> f ImportedModules
f PreScopeState
s = ImportedModules -> f ImportedModules
f (PreScopeState -> ImportedModules
stPreImportedModules PreScopeState
s) f ImportedModules
-> (ImportedModules -> PreScopeState) -> f PreScopeState
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ ImportedModules
x -> PreScopeState
s { stPreImportedModules = x }
lensModuleToSource :: Lens' PreScopeState ModuleToSource
lensModuleToSource :: Lens' PreScopeState ModuleToSource
lensModuleToSource ModuleToSource -> f ModuleToSource
f PreScopeState
s = ModuleToSource -> f ModuleToSource
f (PreScopeState -> ModuleToSource
stPreModuleToSource PreScopeState
s ) f ModuleToSource
-> (ModuleToSource -> PreScopeState) -> f PreScopeState
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ ModuleToSource
x -> PreScopeState
s { stPreModuleToSource = x }
lensVisitedModules :: Lens' PreScopeState VisitedModules
lensVisitedModules :: Lens' PreScopeState VisitedModules
lensVisitedModules VisitedModules -> f VisitedModules
f PreScopeState
s = VisitedModules -> f VisitedModules
f (PreScopeState -> VisitedModules
stPreVisitedModules PreScopeState
s ) f VisitedModules
-> (VisitedModules -> PreScopeState) -> f PreScopeState
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ VisitedModules
x -> PreScopeState
s { stPreVisitedModules = x }
lensScope :: Lens' PreScopeState ScopeInfo
lensScope :: Lens' PreScopeState ScopeInfo
lensScope ScopeInfo -> f ScopeInfo
f PreScopeState
s = ScopeInfo -> f ScopeInfo
f (PreScopeState -> ScopeInfo
stPreScope PreScopeState
s ) f ScopeInfo -> (ScopeInfo -> PreScopeState) -> f PreScopeState
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ ScopeInfo
x -> PreScopeState
s { stPreScope = x }
lensPatternSyns :: Lens' PreScopeState A.PatternSynDefns
lensPatternSyns :: Lens' PreScopeState PatternSynDefns
lensPatternSyns PatternSynDefns -> f PatternSynDefns
f PreScopeState
s = PatternSynDefns -> f PatternSynDefns
f (PreScopeState -> PatternSynDefns
stPrePatternSyns PreScopeState
s ) f PatternSynDefns
-> (PatternSynDefns -> PreScopeState) -> f PreScopeState
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ PatternSynDefns
x -> PreScopeState
s { stPrePatternSyns = x }
lensPatternSynImports :: Lens' PreScopeState A.PatternSynDefns
lensPatternSynImports :: Lens' PreScopeState PatternSynDefns
lensPatternSynImports PatternSynDefns -> f PatternSynDefns
f PreScopeState
s = PatternSynDefns -> f PatternSynDefns
f (PreScopeState -> PatternSynDefns
stPrePatternSynImports PreScopeState
s ) f PatternSynDefns
-> (PatternSynDefns -> PreScopeState) -> f PreScopeState
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ PatternSynDefns
x -> PreScopeState
s { stPrePatternSynImports = x }
lensGeneralizedVars :: Lens' PreScopeState (Strict.Maybe (Set QName))
lensGeneralizedVars :: Lens' PreScopeState (Maybe (Set QName))
lensGeneralizedVars Maybe (Set QName) -> f (Maybe (Set QName))
f PreScopeState
s = Maybe (Set QName) -> f (Maybe (Set QName))
f (PreScopeState -> Maybe (Set QName)
stPreGeneralizedVars PreScopeState
s ) f (Maybe (Set QName))
-> (Maybe (Set QName) -> PreScopeState) -> f PreScopeState
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ Maybe (Set QName)
x -> PreScopeState
s { stPreGeneralizedVars = x }
instance LensPragmaOptions PreScopeState where
lensPragmaOptions :: Lens' PreScopeState PragmaOptions
lensPragmaOptions PragmaOptions -> f PragmaOptions
f PreScopeState
s = PragmaOptions -> f PragmaOptions
f (PreScopeState -> PragmaOptions
stPrePragmaOptions PreScopeState
s ) f PragmaOptions
-> (PragmaOptions -> PreScopeState) -> f PreScopeState
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ PragmaOptions
x -> PreScopeState
s { stPrePragmaOptions = x }
lensImportedBuiltins :: Lens' PreScopeState BuiltinThings
lensImportedBuiltins :: Lens' PreScopeState BuiltinThings
lensImportedBuiltins BuiltinThings -> f BuiltinThings
f PreScopeState
s = BuiltinThings -> f BuiltinThings
f (PreScopeState -> BuiltinThings
stPreImportedBuiltins PreScopeState
s ) f BuiltinThings
-> (BuiltinThings -> PreScopeState) -> f PreScopeState
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ BuiltinThings
x -> PreScopeState
s { stPreImportedBuiltins = x }
lensFreshInteractionId :: Lens' PreScopeState InteractionId
lensFreshInteractionId :: Lens' PreScopeState InteractionId
lensFreshInteractionId InteractionId -> f InteractionId
f PreScopeState
s = InteractionId -> f InteractionId
f (PreScopeState -> InteractionId
stPreFreshInteractionId PreScopeState
s ) f InteractionId
-> (InteractionId -> PreScopeState) -> f PreScopeState
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ InteractionId
x -> PreScopeState
s { stPreFreshInteractionId = x }
lensImportedUserWarnings :: Lens' PreScopeState UserWarnings
lensImportedUserWarnings :: Lens' PreScopeState UserWarnings
lensImportedUserWarnings UserWarnings -> f UserWarnings
f PreScopeState
s = UserWarnings -> f UserWarnings
f (PreScopeState -> UserWarnings
stPreImportedUserWarnings PreScopeState
s ) f UserWarnings
-> (UserWarnings -> PreScopeState) -> f PreScopeState
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ UserWarnings
x -> PreScopeState
s { stPreImportedUserWarnings = x }
lensLocalUserWarnings :: Lens' PreScopeState UserWarnings
lensLocalUserWarnings :: Lens' PreScopeState UserWarnings
lensLocalUserWarnings UserWarnings -> f UserWarnings
f PreScopeState
s = UserWarnings -> f UserWarnings
f (PreScopeState -> UserWarnings
stPreLocalUserWarnings PreScopeState
s ) f UserWarnings
-> (UserWarnings -> PreScopeState) -> f PreScopeState
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ UserWarnings
x -> PreScopeState
s { stPreLocalUserWarnings = x }
lensWarningOnImport :: Lens' PreScopeState (Strict.Maybe Text)
lensWarningOnImport :: Lens' PreScopeState (Maybe BackendName)
lensWarningOnImport Maybe BackendName -> f (Maybe BackendName)
f PreScopeState
s = Maybe BackendName -> f (Maybe BackendName)
f (PreScopeState -> Maybe BackendName
stPreWarningOnImport PreScopeState
s) f (Maybe BackendName)
-> (Maybe BackendName -> PreScopeState) -> f PreScopeState
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ Maybe BackendName
x -> PreScopeState
s { stPreWarningOnImport = x }
lensImportedPartialDefs :: Lens' PreScopeState (Set QName)
lensImportedPartialDefs :: Lens' PreScopeState (Set QName)
lensImportedPartialDefs Set QName -> f (Set QName)
f PreScopeState
s = Set QName -> f (Set QName)
f (PreScopeState -> Set QName
stPreImportedPartialDefs PreScopeState
s) f (Set QName) -> (Set QName -> PreScopeState) -> f PreScopeState
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ Set QName
x -> PreScopeState
s { stPreImportedPartialDefs = x }
lensProjectConfigs :: Lens' PreScopeState (Map FilePath ProjectConfig)
lensProjectConfigs :: Lens' PreScopeState (Map String ProjectConfig)
lensProjectConfigs Map String ProjectConfig -> f (Map String ProjectConfig)
f PreScopeState
s = Map String ProjectConfig -> f (Map String ProjectConfig)
f (PreScopeState -> Map String ProjectConfig
stPreProjectConfigs PreScopeState
s) f (Map String ProjectConfig)
-> (Map String ProjectConfig -> PreScopeState) -> f PreScopeState
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ Map String ProjectConfig
x -> PreScopeState
s { stPreProjectConfigs = x }
lensAgdaLibFiles :: Lens' PreScopeState (Map FilePath AgdaLibFile)
lensAgdaLibFiles :: Lens' PreScopeState (Map String AgdaLibFile)
lensAgdaLibFiles Map String AgdaLibFile -> f (Map String AgdaLibFile)
f PreScopeState
s = Map String AgdaLibFile -> f (Map String AgdaLibFile)
f (PreScopeState -> Map String AgdaLibFile
stPreAgdaLibFiles PreScopeState
s) f (Map String AgdaLibFile)
-> (Map String AgdaLibFile -> PreScopeState) -> f PreScopeState
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ Map String AgdaLibFile
x -> PreScopeState
s { stPreAgdaLibFiles = x }
lensImportedMetaStore :: Lens' PreScopeState RemoteMetaStore
lensImportedMetaStore :: Lens' PreScopeState RemoteMetaStore
lensImportedMetaStore RemoteMetaStore -> f RemoteMetaStore
f PreScopeState
s = RemoteMetaStore -> f RemoteMetaStore
f (PreScopeState -> RemoteMetaStore
stPreImportedMetaStore PreScopeState
s) f RemoteMetaStore
-> (RemoteMetaStore -> PreScopeState) -> f PreScopeState
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \RemoteMetaStore
x -> PreScopeState
s { stPreImportedMetaStore = x }
lensCopiedNames :: Lens' PreScopeState (HashMap QName QName)
lensCopiedNames :: Lens' PreScopeState (HashMap QName QName)
lensCopiedNames HashMap QName QName -> f (HashMap QName QName)
f PreScopeState
s = HashMap QName QName -> f (HashMap QName QName)
f (PreScopeState -> HashMap QName QName
stPreCopiedNames PreScopeState
s) f (HashMap QName QName)
-> (HashMap QName QName -> PreScopeState) -> f PreScopeState
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ HashMap QName QName
x -> PreScopeState
s { stPreCopiedNames = x }
lensNameCopies :: Lens' PreScopeState (HashMap QName (HashSet QName))
lensNameCopies :: Lens' PreScopeState (HashMap QName (HashSet QName))
lensNameCopies HashMap QName (HashSet QName) -> f (HashMap QName (HashSet QName))
f PreScopeState
s = HashMap QName (HashSet QName) -> f (HashMap QName (HashSet QName))
f (PreScopeState -> HashMap QName (HashSet QName)
stPreNameCopies PreScopeState
s) f (HashMap QName (HashSet QName))
-> (HashMap QName (HashSet QName) -> PreScopeState)
-> f PreScopeState
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ HashMap QName (HashSet QName)
x -> PreScopeState
s { stPreNameCopies = x }
lensForeignCode :: Lens' PostScopeState BackendForeignCode
lensForeignCode :: Lens' PostScopeState BackendForeignCode
lensForeignCode BackendForeignCode -> f BackendForeignCode
f PostScopeState
s = BackendForeignCode -> f BackendForeignCode
f (PostScopeState -> BackendForeignCode
stPostForeignCode PostScopeState
s ) f BackendForeignCode
-> (BackendForeignCode -> PostScopeState) -> f PostScopeState
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ BackendForeignCode
x -> PostScopeState
s { stPostForeignCode = x }
lensLocalPartialDefs :: Lens' PostScopeState (Set QName)
lensLocalPartialDefs :: Lens' PostScopeState (Set QName)
lensLocalPartialDefs Set QName -> f (Set QName)
f PostScopeState
s = Set QName -> f (Set QName)
f (PostScopeState -> Set QName
stPostLocalPartialDefs PostScopeState
s) f (Set QName) -> (Set QName -> PostScopeState) -> f PostScopeState
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ Set QName
x -> PostScopeState
s { stPostLocalPartialDefs = x }
lensFreshNameId :: Lens' PostScopeState NameId
lensFreshNameId :: Lens' PostScopeState NameId
lensFreshNameId NameId -> f NameId
f PostScopeState
s = NameId -> f NameId
f (PostScopeState -> NameId
stPostFreshNameId PostScopeState
s) f NameId -> (NameId -> PostScopeState) -> f PostScopeState
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ NameId
x -> PostScopeState
s { stPostFreshNameId = x }
lensFreshOpaqueId :: Lens' PostScopeState OpaqueId
lensFreshOpaqueId :: Lens' PostScopeState OpaqueId
lensFreshOpaqueId OpaqueId -> f OpaqueId
f PostScopeState
s = OpaqueId -> f OpaqueId
f (PostScopeState -> OpaqueId
stPostFreshOpaqueId PostScopeState
s) f OpaqueId -> (OpaqueId -> PostScopeState) -> f PostScopeState
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ OpaqueId
x -> PostScopeState
s { stPostFreshOpaqueId = x }
lensOpaqueBlocks :: Lens' PostScopeState (Map OpaqueId OpaqueBlock)
lensOpaqueBlocks :: Lens' PostScopeState (Map OpaqueId OpaqueBlock)
lensOpaqueBlocks Map OpaqueId OpaqueBlock -> f (Map OpaqueId OpaqueBlock)
f PostScopeState
s = Map OpaqueId OpaqueBlock -> f (Map OpaqueId OpaqueBlock)
f (PostScopeState -> Map OpaqueId OpaqueBlock
stPostOpaqueBlocks PostScopeState
s) f (Map OpaqueId OpaqueBlock)
-> (Map OpaqueId OpaqueBlock -> PostScopeState) -> f PostScopeState
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ Map OpaqueId OpaqueBlock
x -> PostScopeState
s { stPostOpaqueBlocks = x }
lensOpaqueIds :: Lens' PostScopeState (Map QName OpaqueId)
lensOpaqueIds :: Lens' PostScopeState (Map QName OpaqueId)
lensOpaqueIds Map QName OpaqueId -> f (Map QName OpaqueId)
f PostScopeState
s = Map QName OpaqueId -> f (Map QName OpaqueId)
f (PostScopeState -> Map QName OpaqueId
stPostOpaqueIds PostScopeState
s) f (Map QName OpaqueId)
-> (Map QName OpaqueId -> PostScopeState) -> f PostScopeState
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ Map QName OpaqueId
x -> PostScopeState
s { stPostOpaqueIds = x }
lensSyntaxInfo :: Lens' PostScopeState HighlightingInfo
lensSyntaxInfo :: Lens' PostScopeState HighlightingInfo
lensSyntaxInfo HighlightingInfo -> f HighlightingInfo
f PostScopeState
s = HighlightingInfo -> f HighlightingInfo
f (PostScopeState -> HighlightingInfo
stPostSyntaxInfo PostScopeState
s) f HighlightingInfo
-> (HighlightingInfo -> PostScopeState) -> f PostScopeState
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ HighlightingInfo
x -> PostScopeState
s { stPostSyntaxInfo = x }
lensDisambiguatedNames :: Lens' PostScopeState DisambiguatedNames
lensDisambiguatedNames :: Lens' PostScopeState DisambiguatedNames
lensDisambiguatedNames DisambiguatedNames -> f DisambiguatedNames
f PostScopeState
s = DisambiguatedNames -> f DisambiguatedNames
f (PostScopeState -> DisambiguatedNames
stPostDisambiguatedNames PostScopeState
s) f DisambiguatedNames
-> (DisambiguatedNames -> PostScopeState) -> f PostScopeState
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ DisambiguatedNames
x -> PostScopeState
s { stPostDisambiguatedNames = x }
lensOpenMetaStore :: Lens' PostScopeState LocalMetaStore
lensOpenMetaStore :: Lens' PostScopeState LocalMetaStore
lensOpenMetaStore LocalMetaStore -> f LocalMetaStore
f PostScopeState
s = LocalMetaStore -> f LocalMetaStore
f (PostScopeState -> LocalMetaStore
stPostOpenMetaStore PostScopeState
s) f LocalMetaStore
-> (LocalMetaStore -> PostScopeState) -> f PostScopeState
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ LocalMetaStore
x -> PostScopeState
s { stPostOpenMetaStore = x }
lensSolvedMetaStore :: Lens' PostScopeState LocalMetaStore
lensSolvedMetaStore :: Lens' PostScopeState LocalMetaStore
lensSolvedMetaStore LocalMetaStore -> f LocalMetaStore
f PostScopeState
s = LocalMetaStore -> f LocalMetaStore
f (PostScopeState -> LocalMetaStore
stPostSolvedMetaStore PostScopeState
s) f LocalMetaStore
-> (LocalMetaStore -> PostScopeState) -> f PostScopeState
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ LocalMetaStore
x -> PostScopeState
s { stPostSolvedMetaStore = x }
lensInteractionPoints :: Lens' PostScopeState InteractionPoints
lensInteractionPoints :: Lens' PostScopeState InteractionPoints
lensInteractionPoints InteractionPoints -> f InteractionPoints
f PostScopeState
s = InteractionPoints -> f InteractionPoints
f (PostScopeState -> InteractionPoints
stPostInteractionPoints PostScopeState
s) f InteractionPoints
-> (InteractionPoints -> PostScopeState) -> f PostScopeState
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ InteractionPoints
x -> PostScopeState
s { stPostInteractionPoints = x }
lensAwakeConstraints :: Lens' PostScopeState Constraints
lensAwakeConstraints :: Lens' PostScopeState Constraints
lensAwakeConstraints Constraints -> f Constraints
f PostScopeState
s = Constraints -> f Constraints
f (PostScopeState -> Constraints
stPostAwakeConstraints PostScopeState
s) f Constraints
-> (Constraints -> PostScopeState) -> f PostScopeState
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ Constraints
x -> PostScopeState
s { stPostAwakeConstraints = x }
lensSleepingConstraints :: Lens' PostScopeState Constraints
lensSleepingConstraints :: Lens' PostScopeState Constraints
lensSleepingConstraints Constraints -> f Constraints
f PostScopeState
s = Constraints -> f Constraints
f (PostScopeState -> Constraints
stPostSleepingConstraints PostScopeState
s) f Constraints
-> (Constraints -> PostScopeState) -> f PostScopeState
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ Constraints
x -> PostScopeState
s { stPostSleepingConstraints = x }
lensDirty :: Lens' PostScopeState Bool
lensDirty :: Lens' PostScopeState Bool
lensDirty Bool -> f Bool
f PostScopeState
s = Bool -> f Bool
f (PostScopeState -> Bool
stPostDirty PostScopeState
s) f Bool -> (Bool -> PostScopeState) -> f PostScopeState
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ Bool
x -> PostScopeState
s { stPostDirty = x }
lensOccursCheckDefs :: Lens' PostScopeState (Set QName)
lensOccursCheckDefs :: Lens' PostScopeState (Set QName)
lensOccursCheckDefs Set QName -> f (Set QName)
f PostScopeState
s = Set QName -> f (Set QName)
f (PostScopeState -> Set QName
stPostOccursCheckDefs PostScopeState
s) f (Set QName) -> (Set QName -> PostScopeState) -> f PostScopeState
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ Set QName
x -> PostScopeState
s { stPostOccursCheckDefs = x }
lensSignature :: Lens' PostScopeState Signature
lensSignature :: Lens' PostScopeState Signature
lensSignature Signature -> f Signature
f PostScopeState
s = Signature -> f Signature
f (PostScopeState -> Signature
stPostSignature PostScopeState
s) f Signature -> (Signature -> PostScopeState) -> f PostScopeState
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ Signature
x -> PostScopeState
s { stPostSignature = x }
lensModuleCheckpoints :: Lens' PostScopeState (Map ModuleName CheckpointId)
lensModuleCheckpoints :: Lens' PostScopeState (Map ModuleName CheckpointId)
lensModuleCheckpoints Map ModuleName CheckpointId -> f (Map ModuleName CheckpointId)
f PostScopeState
s = Map ModuleName CheckpointId -> f (Map ModuleName CheckpointId)
f (PostScopeState -> Map ModuleName CheckpointId
stPostModuleCheckpoints PostScopeState
s) f (Map ModuleName CheckpointId)
-> (Map ModuleName CheckpointId -> PostScopeState)
-> f PostScopeState
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ Map ModuleName CheckpointId
x -> PostScopeState
s { stPostModuleCheckpoints = x }
lensImportsDisplayForms :: Lens' PostScopeState DisplayForms
lensImportsDisplayForms :: Lens' PostScopeState DisplayForms
lensImportsDisplayForms DisplayForms -> f DisplayForms
f PostScopeState
s = DisplayForms -> f DisplayForms
f (PostScopeState -> DisplayForms
stPostImportsDisplayForms PostScopeState
s) f DisplayForms
-> (DisplayForms -> PostScopeState) -> f PostScopeState
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ DisplayForms
x -> PostScopeState
s { stPostImportsDisplayForms = x }
lensImportedDisplayForms :: Lens' PreScopeState DisplayForms
lensImportedDisplayForms :: Lens' PreScopeState DisplayForms
lensImportedDisplayForms DisplayForms -> f DisplayForms
f PreScopeState
s = DisplayForms -> f DisplayForms
f (PreScopeState -> DisplayForms
stPreImportedDisplayForms PreScopeState
s) f DisplayForms
-> (DisplayForms -> PreScopeState) -> f PreScopeState
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ DisplayForms
x -> PreScopeState
s { stPreImportedDisplayForms = x }
lensTemporaryInstances :: Lens' PostScopeState (Set QName)
lensTemporaryInstances :: Lens' PostScopeState (Set QName)
lensTemporaryInstances Set QName -> f (Set QName)
f PostScopeState
s = Set QName -> f (Set QName)
f (PostScopeState -> Set QName
stPostTemporaryInstances PostScopeState
s) f (Set QName) -> (Set QName -> PostScopeState) -> f PostScopeState
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ Set QName
x -> PostScopeState
s { stPostTemporaryInstances = x }
lensConcreteNames :: Lens' PostScopeState ConcreteNames
lensConcreteNames :: Lens' PostScopeState ConcreteNames
lensConcreteNames ConcreteNames -> f ConcreteNames
f PostScopeState
s = ConcreteNames -> f ConcreteNames
f (PostScopeState -> ConcreteNames
stPostConcreteNames PostScopeState
s) f ConcreteNames
-> (ConcreteNames -> PostScopeState) -> f PostScopeState
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ ConcreteNames
x -> PostScopeState
s { stPostConcreteNames = x }
lensUsedNames :: Lens' PostScopeState UsedNames
lensUsedNames :: Lens' PostScopeState UsedNames
lensUsedNames UsedNames -> f UsedNames
f PostScopeState
s = UsedNames -> f UsedNames
f (PostScopeState -> UsedNames
stPostUsedNames PostScopeState
s) f UsedNames -> (UsedNames -> PostScopeState) -> f PostScopeState
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ UsedNames
x -> PostScopeState
s { stPostUsedNames = x }
lensShadowingNames :: Lens' PostScopeState ShadowingNames
lensShadowingNames :: Lens' PostScopeState ShadowingNames
lensShadowingNames ShadowingNames -> f ShadowingNames
f PostScopeState
s = ShadowingNames -> f ShadowingNames
f (PostScopeState -> ShadowingNames
stPostShadowingNames PostScopeState
s) f ShadowingNames
-> (ShadowingNames -> PostScopeState) -> f PostScopeState
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ ShadowingNames
x -> PostScopeState
s { stPostShadowingNames = x }
lensStatistics :: Lens' PostScopeState Statistics
lensStatistics :: Lens' PostScopeState Statistics
lensStatistics Statistics -> f Statistics
f PostScopeState
s = Statistics -> f Statistics
f (PostScopeState -> Statistics
stPostStatistics PostScopeState
s) f Statistics -> (Statistics -> PostScopeState) -> f PostScopeState
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ Statistics
x -> PostScopeState
s { stPostStatistics = x }
lensTCWarnings :: Lens' PostScopeState (Set TCWarning)
lensTCWarnings :: Lens' PostScopeState (Set TCWarning)
lensTCWarnings Set TCWarning -> f (Set TCWarning)
f PostScopeState
s = Set TCWarning -> f (Set TCWarning)
f (PostScopeState -> Set TCWarning
stPostTCWarnings PostScopeState
s) f (Set TCWarning)
-> (Set TCWarning -> PostScopeState) -> f PostScopeState
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ Set TCWarning
x -> PostScopeState
s { stPostTCWarnings = x }
lensMutualBlocks :: Lens' PostScopeState MutualBlocks
lensMutualBlocks :: Lens' PostScopeState MutualBlocks
lensMutualBlocks MutualBlocks -> f MutualBlocks
f PostScopeState
s = MutualBlocks -> f MutualBlocks
f (PostScopeState -> MutualBlocks
stPostMutualBlocks PostScopeState
s) f MutualBlocks
-> (MutualBlocks -> PostScopeState) -> f PostScopeState
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ MutualBlocks
x -> PostScopeState
s { stPostMutualBlocks = x }
lensLocalBuiltins :: Lens' PostScopeState BuiltinThings
lensLocalBuiltins :: Lens' PostScopeState BuiltinThings
lensLocalBuiltins BuiltinThings -> f BuiltinThings
f PostScopeState
s = BuiltinThings -> f BuiltinThings
f (PostScopeState -> BuiltinThings
stPostLocalBuiltins PostScopeState
s) f BuiltinThings
-> (BuiltinThings -> PostScopeState) -> f PostScopeState
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ BuiltinThings
x -> PostScopeState
s { stPostLocalBuiltins = x }
lensFreshMetaId :: Lens' PostScopeState MetaId
lensFreshMetaId :: Lens' PostScopeState MetaId
lensFreshMetaId MetaId -> f MetaId
f PostScopeState
s = MetaId -> f MetaId
f (PostScopeState -> MetaId
stPostFreshMetaId PostScopeState
s) f MetaId -> (MetaId -> PostScopeState) -> f PostScopeState
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ MetaId
x -> PostScopeState
s { stPostFreshMetaId = x }
lensFreshMutualId :: Lens' PostScopeState MutualId
lensFreshMutualId :: Lens' PostScopeState MutualId
lensFreshMutualId MutualId -> f MutualId
f PostScopeState
s = MutualId -> f MutualId
f (PostScopeState -> MutualId
stPostFreshMutualId PostScopeState
s) f MutualId -> (MutualId -> PostScopeState) -> f PostScopeState
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ MutualId
x -> PostScopeState
s { stPostFreshMutualId = x }
lensFreshProblemId :: Lens' PostScopeState ProblemId
lensFreshProblemId :: Lens' PostScopeState ProblemId
lensFreshProblemId ProblemId -> f ProblemId
f PostScopeState
s = ProblemId -> f ProblemId
f (PostScopeState -> ProblemId
stPostFreshProblemId PostScopeState
s) f ProblemId -> (ProblemId -> PostScopeState) -> f PostScopeState
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ ProblemId
x -> PostScopeState
s { stPostFreshProblemId = x }
lensFreshCheckpointId :: Lens' PostScopeState CheckpointId
lensFreshCheckpointId :: Lens' PostScopeState CheckpointId
lensFreshCheckpointId CheckpointId -> f CheckpointId
f PostScopeState
s = CheckpointId -> f CheckpointId
f (PostScopeState -> CheckpointId
stPostFreshCheckpointId PostScopeState
s) f CheckpointId
-> (CheckpointId -> PostScopeState) -> f PostScopeState
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ CheckpointId
x -> PostScopeState
s { stPostFreshCheckpointId = x }
lensFreshInt :: Lens' PostScopeState Int
lensFreshInt :: Lens' PostScopeState Int
lensFreshInt Int -> f Int
f PostScopeState
s = Int -> f Int
f (PostScopeState -> Int
stPostFreshInt PostScopeState
s) f Int -> (Int -> PostScopeState) -> f PostScopeState
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ Int
x -> PostScopeState
s { stPostFreshInt = x }
lensAreWeCaching :: Lens' PostScopeState Bool
lensAreWeCaching :: Lens' PostScopeState Bool
lensAreWeCaching Bool -> f Bool
f PostScopeState
s = Bool -> f Bool
f (PostScopeState -> Bool
stPostAreWeCaching PostScopeState
s) f Bool -> (Bool -> PostScopeState) -> f PostScopeState
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \Bool
x -> PostScopeState
s { stPostAreWeCaching = x }
lensPostponeInstanceSearch :: Lens' PostScopeState Bool
lensPostponeInstanceSearch :: Lens' PostScopeState Bool
lensPostponeInstanceSearch Bool -> f Bool
f PostScopeState
s = Bool -> f Bool
f (PostScopeState -> Bool
stPostPostponeInstanceSearch PostScopeState
s) f Bool -> (Bool -> PostScopeState) -> f PostScopeState
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ Bool
x -> PostScopeState
s { stPostPostponeInstanceSearch = x }
lensConsideringInstance :: Lens' PostScopeState Bool
lensConsideringInstance :: Lens' PostScopeState Bool
lensConsideringInstance Bool -> f Bool
f PostScopeState
s = Bool -> f Bool
f (PostScopeState -> Bool
stPostConsideringInstance PostScopeState
s) f Bool -> (Bool -> PostScopeState) -> f PostScopeState
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ Bool
x -> PostScopeState
s { stPostConsideringInstance = x }
lensInstantiateBlocking :: Lens' PostScopeState Bool
lensInstantiateBlocking :: Lens' PostScopeState Bool
lensInstantiateBlocking Bool -> f Bool
f PostScopeState
s = Bool -> f Bool
f (PostScopeState -> Bool
stPostInstantiateBlocking PostScopeState
s) f Bool -> (Bool -> PostScopeState) -> f PostScopeState
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ Bool
x -> PostScopeState
s { stPostInstantiateBlocking = x }
stLoadedFileCache :: Lens' TCState (Maybe LoadedFileCache)
stLoadedFileCache :: Lens' TCState (Maybe LoadedFileCache)
stLoadedFileCache = (PersistentTCState -> f PersistentTCState) -> TCState -> f TCState
Lens' TCState PersistentTCState
lensPersistentState ((PersistentTCState -> f PersistentTCState)
-> TCState -> f TCState)
-> ((Maybe LoadedFileCache -> f (Maybe LoadedFileCache))
-> PersistentTCState -> f PersistentTCState)
-> (Maybe LoadedFileCache -> f (Maybe LoadedFileCache))
-> TCState
-> f TCState
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (Maybe LoadedFileCache -> f (Maybe LoadedFileCache))
-> PersistentTCState -> f PersistentTCState
Lens' PersistentTCState (Maybe LoadedFileCache)
lensLoadedFileCache ((Maybe LoadedFileCache -> f (Maybe LoadedFileCache))
-> PersistentTCState -> f PersistentTCState)
-> ((Maybe LoadedFileCache -> f (Maybe LoadedFileCache))
-> Maybe LoadedFileCache -> f (Maybe LoadedFileCache))
-> (Maybe LoadedFileCache -> f (Maybe LoadedFileCache))
-> PersistentTCState
-> f PersistentTCState
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (Maybe LoadedFileCache -> f (Maybe LoadedFileCache))
-> Maybe LoadedFileCache -> f (Maybe LoadedFileCache)
forall a (f :: * -> *).
Functor f =>
(Maybe a -> f (Maybe a)) -> Maybe a -> f (Maybe a)
Strict.lensMaybeLazy
stBackends :: Lens' TCState [Backend]
stBackends :: Lens' TCState [Backend]
stBackends = (PersistentTCState -> f PersistentTCState) -> TCState -> f TCState
Lens' TCState PersistentTCState
lensPersistentState ((PersistentTCState -> f PersistentTCState)
-> TCState -> f TCState)
-> (([Backend] -> f [Backend])
-> PersistentTCState -> f PersistentTCState)
-> ([Backend] -> f [Backend])
-> TCState
-> f TCState
forall b c a. (b -> c) -> (a -> b) -> a -> c
. ([Backend] -> f [Backend])
-> PersistentTCState -> f PersistentTCState
Lens' PersistentTCState [Backend]
lensBackends
stTopLevelModuleNames :: Lens' TCState (BiMap RawTopLevelModuleName ModuleNameHash)
stTopLevelModuleNames :: Lens' TCState (BiMap RawTopLevelModuleName ModuleNameHash)
stTopLevelModuleNames = (PersistentTCState -> f PersistentTCState) -> TCState -> f TCState
Lens' TCState PersistentTCState
lensPersistentState ((PersistentTCState -> f PersistentTCState)
-> TCState -> f TCState)
-> ((BiMap RawTopLevelModuleName ModuleNameHash
-> f (BiMap RawTopLevelModuleName ModuleNameHash))
-> PersistentTCState -> f PersistentTCState)
-> (BiMap RawTopLevelModuleName ModuleNameHash
-> f (BiMap RawTopLevelModuleName ModuleNameHash))
-> TCState
-> f TCState
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (BiMap RawTopLevelModuleName ModuleNameHash
-> f (BiMap RawTopLevelModuleName ModuleNameHash))
-> PersistentTCState -> f PersistentTCState
Lens'
PersistentTCState (BiMap RawTopLevelModuleName ModuleNameHash)
lensTopLevelModuleNames
stTokens :: Lens' TCState HighlightingInfo
stTokens :: Lens' TCState HighlightingInfo
stTokens = (PreScopeState -> f PreScopeState) -> TCState -> f TCState
Lens' TCState PreScopeState
lensPreScopeState ((PreScopeState -> f PreScopeState) -> TCState -> f TCState)
-> ((HighlightingInfo -> f HighlightingInfo)
-> PreScopeState -> f PreScopeState)
-> (HighlightingInfo -> f HighlightingInfo)
-> TCState
-> f TCState
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (HighlightingInfo -> f HighlightingInfo)
-> PreScopeState -> f PreScopeState
Lens' PreScopeState HighlightingInfo
lensPreTokens
stImports :: Lens' TCState Signature
stImports :: Lens' TCState Signature
stImports = (PreScopeState -> f PreScopeState) -> TCState -> f TCState
Lens' TCState PreScopeState
lensPreScopeState ((PreScopeState -> f PreScopeState) -> TCState -> f TCState)
-> ((Signature -> f Signature) -> PreScopeState -> f PreScopeState)
-> (Signature -> f Signature)
-> TCState
-> f TCState
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (Signature -> f Signature) -> PreScopeState -> f PreScopeState
Lens' PreScopeState Signature
lensImports
stImportedModules :: Lens' TCState ImportedModules
stImportedModules :: Lens' TCState ImportedModules
stImportedModules = (PreScopeState -> f PreScopeState) -> TCState -> f TCState
Lens' TCState PreScopeState
lensPreScopeState ((PreScopeState -> f PreScopeState) -> TCState -> f TCState)
-> ((ImportedModules -> f ImportedModules)
-> PreScopeState -> f PreScopeState)
-> (ImportedModules -> f ImportedModules)
-> TCState
-> f TCState
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (ImportedModules -> f ImportedModules)
-> PreScopeState -> f PreScopeState
Lens' PreScopeState ImportedModules
lensImportedModules
stModuleToSource :: Lens' TCState ModuleToSource
stModuleToSource :: Lens' TCState ModuleToSource
stModuleToSource = (PreScopeState -> f PreScopeState) -> TCState -> f TCState
Lens' TCState PreScopeState
lensPreScopeState ((PreScopeState -> f PreScopeState) -> TCState -> f TCState)
-> ((ModuleToSource -> f ModuleToSource)
-> PreScopeState -> f PreScopeState)
-> (ModuleToSource -> f ModuleToSource)
-> TCState
-> f TCState
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (ModuleToSource -> f ModuleToSource)
-> PreScopeState -> f PreScopeState
Lens' PreScopeState ModuleToSource
lensModuleToSource
stVisitedModules :: Lens' TCState VisitedModules
stVisitedModules :: Lens' TCState VisitedModules
stVisitedModules = (PreScopeState -> f PreScopeState) -> TCState -> f TCState
Lens' TCState PreScopeState
lensPreScopeState ((PreScopeState -> f PreScopeState) -> TCState -> f TCState)
-> ((VisitedModules -> f VisitedModules)
-> PreScopeState -> f PreScopeState)
-> (VisitedModules -> f VisitedModules)
-> TCState
-> f TCState
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (VisitedModules -> f VisitedModules)
-> PreScopeState -> f PreScopeState
Lens' PreScopeState VisitedModules
lensVisitedModules
stScope :: Lens' TCState ScopeInfo
stScope :: Lens' TCState ScopeInfo
stScope = (PreScopeState -> f PreScopeState) -> TCState -> f TCState
Lens' TCState PreScopeState
lensPreScopeState ((PreScopeState -> f PreScopeState) -> TCState -> f TCState)
-> ((ScopeInfo -> f ScopeInfo) -> PreScopeState -> f PreScopeState)
-> (ScopeInfo -> f ScopeInfo)
-> TCState
-> f TCState
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (ScopeInfo -> f ScopeInfo) -> PreScopeState -> f PreScopeState
Lens' PreScopeState ScopeInfo
lensScope
stPatternSyns :: Lens' TCState A.PatternSynDefns
stPatternSyns :: Lens' TCState PatternSynDefns
stPatternSyns = (PreScopeState -> f PreScopeState) -> TCState -> f TCState
Lens' TCState PreScopeState
lensPreScopeState ((PreScopeState -> f PreScopeState) -> TCState -> f TCState)
-> ((PatternSynDefns -> f PatternSynDefns)
-> PreScopeState -> f PreScopeState)
-> (PatternSynDefns -> f PatternSynDefns)
-> TCState
-> f TCState
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (PatternSynDefns -> f PatternSynDefns)
-> PreScopeState -> f PreScopeState
Lens' PreScopeState PatternSynDefns
lensPatternSyns
stPatternSynImports :: Lens' TCState A.PatternSynDefns
stPatternSynImports :: Lens' TCState PatternSynDefns
stPatternSynImports = (PreScopeState -> f PreScopeState) -> TCState -> f TCState
Lens' TCState PreScopeState
lensPreScopeState ((PreScopeState -> f PreScopeState) -> TCState -> f TCState)
-> ((PatternSynDefns -> f PatternSynDefns)
-> PreScopeState -> f PreScopeState)
-> (PatternSynDefns -> f PatternSynDefns)
-> TCState
-> f TCState
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (PatternSynDefns -> f PatternSynDefns)
-> PreScopeState -> f PreScopeState
Lens' PreScopeState PatternSynDefns
lensPatternSynImports
stGeneralizedVars :: Lens' TCState (Maybe (Set QName))
stGeneralizedVars :: Lens' TCState (Maybe (Set QName))
stGeneralizedVars = (PreScopeState -> f PreScopeState) -> TCState -> f TCState
Lens' TCState PreScopeState
lensPreScopeState ((PreScopeState -> f PreScopeState) -> TCState -> f TCState)
-> ((Maybe (Set QName) -> f (Maybe (Set QName)))
-> PreScopeState -> f PreScopeState)
-> (Maybe (Set QName) -> f (Maybe (Set QName)))
-> TCState
-> f TCState
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (Maybe (Set QName) -> f (Maybe (Set QName)))
-> PreScopeState -> f PreScopeState
Lens' PreScopeState (Maybe (Set QName))
lensGeneralizedVars ((Maybe (Set QName) -> f (Maybe (Set QName)))
-> PreScopeState -> f PreScopeState)
-> ((Maybe (Set QName) -> f (Maybe (Set QName)))
-> Maybe (Set QName) -> f (Maybe (Set QName)))
-> (Maybe (Set QName) -> f (Maybe (Set QName)))
-> PreScopeState
-> f PreScopeState
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (Maybe (Set QName) -> f (Maybe (Set QName)))
-> Maybe (Set QName) -> f (Maybe (Set QName))
forall a (f :: * -> *).
Functor f =>
(Maybe a -> f (Maybe a)) -> Maybe a -> f (Maybe a)
Strict.lensMaybeLazy
instance LensPragmaOptions TCState where
lensPragmaOptions :: Lens' TCState PragmaOptions
lensPragmaOptions = (PreScopeState -> f PreScopeState) -> TCState -> f TCState
Lens' TCState PreScopeState
lensPreScopeState ((PreScopeState -> f PreScopeState) -> TCState -> f TCState)
-> ((PragmaOptions -> f PragmaOptions)
-> PreScopeState -> f PreScopeState)
-> (PragmaOptions -> f PragmaOptions)
-> TCState
-> f TCState
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (PragmaOptions -> f PragmaOptions)
-> PreScopeState -> f PreScopeState
forall a. LensPragmaOptions a => Lens' a PragmaOptions
Lens' PreScopeState PragmaOptions
lensPragmaOptions
stPragmaOptions :: Lens' TCState PragmaOptions
stPragmaOptions :: Lens' TCState PragmaOptions
stPragmaOptions = (PragmaOptions -> f PragmaOptions) -> TCState -> f TCState
forall a. LensPragmaOptions a => Lens' a PragmaOptions
Lens' TCState PragmaOptions
lensPragmaOptions
stImportedBuiltins :: Lens' TCState BuiltinThings
stImportedBuiltins :: Lens' TCState BuiltinThings
stImportedBuiltins = (PreScopeState -> f PreScopeState) -> TCState -> f TCState
Lens' TCState PreScopeState
lensPreScopeState ((PreScopeState -> f PreScopeState) -> TCState -> f TCState)
-> ((BuiltinThings -> f BuiltinThings)
-> PreScopeState -> f PreScopeState)
-> (BuiltinThings -> f BuiltinThings)
-> TCState
-> f TCState
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (BuiltinThings -> f BuiltinThings)
-> PreScopeState -> f PreScopeState
Lens' PreScopeState BuiltinThings
lensImportedBuiltins
stForeignCode :: Lens' TCState BackendForeignCode
stForeignCode :: Lens' TCState BackendForeignCode
stForeignCode = (PostScopeState -> f PostScopeState) -> TCState -> f TCState
Lens' TCState PostScopeState
lensPostScopeState ((PostScopeState -> f PostScopeState) -> TCState -> f TCState)
-> ((BackendForeignCode -> f BackendForeignCode)
-> PostScopeState -> f PostScopeState)
-> (BackendForeignCode -> f BackendForeignCode)
-> TCState
-> f TCState
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (BackendForeignCode -> f BackendForeignCode)
-> PostScopeState -> f PostScopeState
Lens' PostScopeState BackendForeignCode
lensForeignCode
stFreshInteractionId :: Lens' TCState InteractionId
stFreshInteractionId :: Lens' TCState InteractionId
stFreshInteractionId = (PreScopeState -> f PreScopeState) -> TCState -> f TCState
Lens' TCState PreScopeState
lensPreScopeState ((PreScopeState -> f PreScopeState) -> TCState -> f TCState)
-> ((InteractionId -> f InteractionId)
-> PreScopeState -> f PreScopeState)
-> (InteractionId -> f InteractionId)
-> TCState
-> f TCState
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (InteractionId -> f InteractionId)
-> PreScopeState -> f PreScopeState
Lens' PreScopeState InteractionId
lensFreshInteractionId
stImportedUserWarnings :: Lens' TCState UserWarnings
stImportedUserWarnings :: Lens' TCState UserWarnings
stImportedUserWarnings = (PreScopeState -> f PreScopeState) -> TCState -> f TCState
Lens' TCState PreScopeState
lensPreScopeState ((PreScopeState -> f PreScopeState) -> TCState -> f TCState)
-> ((UserWarnings -> f UserWarnings)
-> PreScopeState -> f PreScopeState)
-> (UserWarnings -> f UserWarnings)
-> TCState
-> f TCState
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (UserWarnings -> f UserWarnings)
-> PreScopeState -> f PreScopeState
Lens' PreScopeState UserWarnings
lensImportedUserWarnings
stLocalUserWarnings :: Lens' TCState UserWarnings
stLocalUserWarnings :: Lens' TCState UserWarnings
stLocalUserWarnings = (PreScopeState -> f PreScopeState) -> TCState -> f TCState
Lens' TCState PreScopeState
lensPreScopeState ((PreScopeState -> f PreScopeState) -> TCState -> f TCState)
-> ((UserWarnings -> f UserWarnings)
-> PreScopeState -> f PreScopeState)
-> (UserWarnings -> f UserWarnings)
-> TCState
-> f TCState
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (UserWarnings -> f UserWarnings)
-> PreScopeState -> f PreScopeState
Lens' PreScopeState UserWarnings
lensLocalUserWarnings
getUserWarnings :: ReadTCState m => m UserWarnings
getUserWarnings :: forall (m :: * -> *). ReadTCState m => m UserWarnings
getUserWarnings = do
iuw <- Lens' TCState UserWarnings -> m UserWarnings
forall (m :: * -> *) a. ReadTCState m => Lens' TCState a -> m a
useR (UserWarnings -> f UserWarnings) -> TCState -> f TCState
Lens' TCState UserWarnings
stImportedUserWarnings
luw <- useR stLocalUserWarnings
return $ iuw `Map.union` luw
stWarningOnImport :: Lens' TCState (Maybe Text)
stWarningOnImport :: Lens' TCState (Maybe BackendName)
stWarningOnImport = (PreScopeState -> f PreScopeState) -> TCState -> f TCState
Lens' TCState PreScopeState
lensPreScopeState ((PreScopeState -> f PreScopeState) -> TCState -> f TCState)
-> ((Maybe BackendName -> f (Maybe BackendName))
-> PreScopeState -> f PreScopeState)
-> (Maybe BackendName -> f (Maybe BackendName))
-> TCState
-> f TCState
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (Maybe BackendName -> f (Maybe BackendName))
-> PreScopeState -> f PreScopeState
Lens' PreScopeState (Maybe BackendName)
lensWarningOnImport ((Maybe BackendName -> f (Maybe BackendName))
-> PreScopeState -> f PreScopeState)
-> ((Maybe BackendName -> f (Maybe BackendName))
-> Maybe BackendName -> f (Maybe BackendName))
-> (Maybe BackendName -> f (Maybe BackendName))
-> PreScopeState
-> f PreScopeState
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (Maybe BackendName -> f (Maybe BackendName))
-> Maybe BackendName -> f (Maybe BackendName)
forall a (f :: * -> *).
Functor f =>
(Maybe a -> f (Maybe a)) -> Maybe a -> f (Maybe a)
Strict.lensMaybeLazy
stImportedPartialDefs :: Lens' TCState (Set QName)
stImportedPartialDefs :: Lens' TCState (Set QName)
stImportedPartialDefs = (PreScopeState -> f PreScopeState) -> TCState -> f TCState
Lens' TCState PreScopeState
lensPreScopeState ((PreScopeState -> f PreScopeState) -> TCState -> f TCState)
-> ((Set QName -> f (Set QName))
-> PreScopeState -> f PreScopeState)
-> (Set QName -> f (Set QName))
-> TCState
-> f TCState
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (Set QName -> f (Set QName)) -> PreScopeState -> f PreScopeState
Lens' PreScopeState (Set QName)
lensImportedPartialDefs
stProjectConfigs :: Lens' TCState (Map FilePath ProjectConfig)
stProjectConfigs :: Lens' TCState (Map String ProjectConfig)
stProjectConfigs = (PreScopeState -> f PreScopeState) -> TCState -> f TCState
Lens' TCState PreScopeState
lensPreScopeState ((PreScopeState -> f PreScopeState) -> TCState -> f TCState)
-> ((Map String ProjectConfig -> f (Map String ProjectConfig))
-> PreScopeState -> f PreScopeState)
-> (Map String ProjectConfig -> f (Map String ProjectConfig))
-> TCState
-> f TCState
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (Map String ProjectConfig -> f (Map String ProjectConfig))
-> PreScopeState -> f PreScopeState
Lens' PreScopeState (Map String ProjectConfig)
lensProjectConfigs
stAgdaLibFiles :: Lens' TCState (Map FilePath AgdaLibFile)
stAgdaLibFiles :: Lens' TCState (Map String AgdaLibFile)
stAgdaLibFiles = (PreScopeState -> f PreScopeState) -> TCState -> f TCState
Lens' TCState PreScopeState
lensPreScopeState ((PreScopeState -> f PreScopeState) -> TCState -> f TCState)
-> ((Map String AgdaLibFile -> f (Map String AgdaLibFile))
-> PreScopeState -> f PreScopeState)
-> (Map String AgdaLibFile -> f (Map String AgdaLibFile))
-> TCState
-> f TCState
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (Map String AgdaLibFile -> f (Map String AgdaLibFile))
-> PreScopeState -> f PreScopeState
Lens' PreScopeState (Map String AgdaLibFile)
lensAgdaLibFiles
stImportedMetaStore :: Lens' TCState RemoteMetaStore
stImportedMetaStore :: Lens' TCState RemoteMetaStore
stImportedMetaStore = (PreScopeState -> f PreScopeState) -> TCState -> f TCState
Lens' TCState PreScopeState
lensPreScopeState ((PreScopeState -> f PreScopeState) -> TCState -> f TCState)
-> ((RemoteMetaStore -> f RemoteMetaStore)
-> PreScopeState -> f PreScopeState)
-> (RemoteMetaStore -> f RemoteMetaStore)
-> TCState
-> f TCState
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (RemoteMetaStore -> f RemoteMetaStore)
-> PreScopeState -> f PreScopeState
Lens' PreScopeState RemoteMetaStore
lensImportedMetaStore
stCopiedNames :: Lens' TCState (HashMap QName QName)
stCopiedNames :: Lens' TCState (HashMap QName QName)
stCopiedNames = (PreScopeState -> f PreScopeState) -> TCState -> f TCState
Lens' TCState PreScopeState
lensPreScopeState ((PreScopeState -> f PreScopeState) -> TCState -> f TCState)
-> ((HashMap QName QName -> f (HashMap QName QName))
-> PreScopeState -> f PreScopeState)
-> (HashMap QName QName -> f (HashMap QName QName))
-> TCState
-> f TCState
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (HashMap QName QName -> f (HashMap QName QName))
-> PreScopeState -> f PreScopeState
Lens' PreScopeState (HashMap QName QName)
lensCopiedNames
stNameCopies :: Lens' TCState (HashMap QName (HashSet QName))
stNameCopies :: Lens' TCState (HashMap QName (HashSet QName))
stNameCopies = (PreScopeState -> f PreScopeState) -> TCState -> f TCState
Lens' TCState PreScopeState
lensPreScopeState ((PreScopeState -> f PreScopeState) -> TCState -> f TCState)
-> ((HashMap QName (HashSet QName)
-> f (HashMap QName (HashSet QName)))
-> PreScopeState -> f PreScopeState)
-> (HashMap QName (HashSet QName)
-> f (HashMap QName (HashSet QName)))
-> TCState
-> f TCState
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (HashMap QName (HashSet QName)
-> f (HashMap QName (HashSet QName)))
-> PreScopeState -> f PreScopeState
Lens' PreScopeState (HashMap QName (HashSet QName))
lensNameCopies
stImportedDisplayForms :: Lens' TCState DisplayForms
stImportedDisplayForms :: Lens' TCState DisplayForms
stImportedDisplayForms = (PreScopeState -> f PreScopeState) -> TCState -> f TCState
Lens' TCState PreScopeState
lensPreScopeState ((PreScopeState -> f PreScopeState) -> TCState -> f TCState)
-> ((DisplayForms -> f DisplayForms)
-> PreScopeState -> f PreScopeState)
-> (DisplayForms -> f DisplayForms)
-> TCState
-> f TCState
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (DisplayForms -> f DisplayForms)
-> PreScopeState -> f PreScopeState
Lens' PreScopeState DisplayForms
lensImportedDisplayForms
stLocalPartialDefs :: Lens' TCState (Set QName)
stLocalPartialDefs :: Lens' TCState (Set QName)
stLocalPartialDefs = (PostScopeState -> f PostScopeState) -> TCState -> f TCState
Lens' TCState PostScopeState
lensPostScopeState ((PostScopeState -> f PostScopeState) -> TCState -> f TCState)
-> ((Set QName -> f (Set QName))
-> PostScopeState -> f PostScopeState)
-> (Set QName -> f (Set QName))
-> TCState
-> f TCState
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (Set QName -> f (Set QName)) -> PostScopeState -> f PostScopeState
Lens' PostScopeState (Set QName)
lensLocalPartialDefs
getPartialDefs :: ReadTCState m => m (Set QName)
getPartialDefs :: forall (m :: * -> *). ReadTCState m => m (Set QName)
getPartialDefs = do
ipd <- Lens' TCState (Set QName) -> m (Set QName)
forall (m :: * -> *) a. ReadTCState m => Lens' TCState a -> m a
useR (Set QName -> f (Set QName)) -> TCState -> f TCState
Lens' TCState (Set QName)
stImportedPartialDefs
lpd <- useR stLocalPartialDefs
return $ ipd `Set.union` lpd
stFreshNameId :: Lens' TCState NameId
stFreshNameId :: Lens' TCState NameId
stFreshNameId = (PostScopeState -> f PostScopeState) -> TCState -> f TCState
Lens' TCState PostScopeState
lensPostScopeState ((PostScopeState -> f PostScopeState) -> TCState -> f TCState)
-> ((NameId -> f NameId) -> PostScopeState -> f PostScopeState)
-> (NameId -> f NameId)
-> TCState
-> f TCState
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (NameId -> f NameId) -> PostScopeState -> f PostScopeState
Lens' PostScopeState NameId
lensFreshNameId
stFreshOpaqueId :: Lens' TCState OpaqueId
stFreshOpaqueId :: Lens' TCState OpaqueId
stFreshOpaqueId = (PostScopeState -> f PostScopeState) -> TCState -> f TCState
Lens' TCState PostScopeState
lensPostScopeState ((PostScopeState -> f PostScopeState) -> TCState -> f TCState)
-> ((OpaqueId -> f OpaqueId) -> PostScopeState -> f PostScopeState)
-> (OpaqueId -> f OpaqueId)
-> TCState
-> f TCState
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (OpaqueId -> f OpaqueId) -> PostScopeState -> f PostScopeState
Lens' PostScopeState OpaqueId
lensFreshOpaqueId
stOpaqueBlocks :: Lens' TCState (Map OpaqueId OpaqueBlock)
stOpaqueBlocks :: Lens' TCState (Map OpaqueId OpaqueBlock)
stOpaqueBlocks = (PostScopeState -> f PostScopeState) -> TCState -> f TCState
Lens' TCState PostScopeState
lensPostScopeState ((PostScopeState -> f PostScopeState) -> TCState -> f TCState)
-> ((Map OpaqueId OpaqueBlock -> f (Map OpaqueId OpaqueBlock))
-> PostScopeState -> f PostScopeState)
-> (Map OpaqueId OpaqueBlock -> f (Map OpaqueId OpaqueBlock))
-> TCState
-> f TCState
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (Map OpaqueId OpaqueBlock -> f (Map OpaqueId OpaqueBlock))
-> PostScopeState -> f PostScopeState
Lens' PostScopeState (Map OpaqueId OpaqueBlock)
lensOpaqueBlocks
stOpaqueIds :: Lens' TCState (Map QName OpaqueId)
stOpaqueIds :: Lens' TCState (Map QName OpaqueId)
stOpaqueIds = (PostScopeState -> f PostScopeState) -> TCState -> f TCState
Lens' TCState PostScopeState
lensPostScopeState ((PostScopeState -> f PostScopeState) -> TCState -> f TCState)
-> ((Map QName OpaqueId -> f (Map QName OpaqueId))
-> PostScopeState -> f PostScopeState)
-> (Map QName OpaqueId -> f (Map QName OpaqueId))
-> TCState
-> f TCState
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (Map QName OpaqueId -> f (Map QName OpaqueId))
-> PostScopeState -> f PostScopeState
Lens' PostScopeState (Map QName OpaqueId)
lensOpaqueIds
stSyntaxInfo :: Lens' TCState HighlightingInfo
stSyntaxInfo :: Lens' TCState HighlightingInfo
stSyntaxInfo = (PostScopeState -> f PostScopeState) -> TCState -> f TCState
Lens' TCState PostScopeState
lensPostScopeState ((PostScopeState -> f PostScopeState) -> TCState -> f TCState)
-> ((HighlightingInfo -> f HighlightingInfo)
-> PostScopeState -> f PostScopeState)
-> (HighlightingInfo -> f HighlightingInfo)
-> TCState
-> f TCState
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (HighlightingInfo -> f HighlightingInfo)
-> PostScopeState -> f PostScopeState
Lens' PostScopeState HighlightingInfo
lensSyntaxInfo
stDisambiguatedNames :: Lens' TCState DisambiguatedNames
stDisambiguatedNames :: Lens' TCState DisambiguatedNames
stDisambiguatedNames = (PostScopeState -> f PostScopeState) -> TCState -> f TCState
Lens' TCState PostScopeState
lensPostScopeState ((PostScopeState -> f PostScopeState) -> TCState -> f TCState)
-> ((DisambiguatedNames -> f DisambiguatedNames)
-> PostScopeState -> f PostScopeState)
-> (DisambiguatedNames -> f DisambiguatedNames)
-> TCState
-> f TCState
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (DisambiguatedNames -> f DisambiguatedNames)
-> PostScopeState -> f PostScopeState
Lens' PostScopeState DisambiguatedNames
lensDisambiguatedNames
stOpenMetaStore :: Lens' TCState LocalMetaStore
stOpenMetaStore :: Lens' TCState LocalMetaStore
stOpenMetaStore = (PostScopeState -> f PostScopeState) -> TCState -> f TCState
Lens' TCState PostScopeState
lensPostScopeState ((PostScopeState -> f PostScopeState) -> TCState -> f TCState)
-> ((LocalMetaStore -> f LocalMetaStore)
-> PostScopeState -> f PostScopeState)
-> (LocalMetaStore -> f LocalMetaStore)
-> TCState
-> f TCState
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (LocalMetaStore -> f LocalMetaStore)
-> PostScopeState -> f PostScopeState
Lens' PostScopeState LocalMetaStore
lensOpenMetaStore
stSolvedMetaStore :: Lens' TCState LocalMetaStore
stSolvedMetaStore :: Lens' TCState LocalMetaStore
stSolvedMetaStore = (PostScopeState -> f PostScopeState) -> TCState -> f TCState
Lens' TCState PostScopeState
lensPostScopeState ((PostScopeState -> f PostScopeState) -> TCState -> f TCState)
-> ((LocalMetaStore -> f LocalMetaStore)
-> PostScopeState -> f PostScopeState)
-> (LocalMetaStore -> f LocalMetaStore)
-> TCState
-> f TCState
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (LocalMetaStore -> f LocalMetaStore)
-> PostScopeState -> f PostScopeState
Lens' PostScopeState LocalMetaStore
lensSolvedMetaStore
stInteractionPoints :: Lens' TCState InteractionPoints
stInteractionPoints :: Lens' TCState InteractionPoints
stInteractionPoints = (PostScopeState -> f PostScopeState) -> TCState -> f TCState
Lens' TCState PostScopeState
lensPostScopeState ((PostScopeState -> f PostScopeState) -> TCState -> f TCState)
-> ((InteractionPoints -> f InteractionPoints)
-> PostScopeState -> f PostScopeState)
-> (InteractionPoints -> f InteractionPoints)
-> TCState
-> f TCState
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (InteractionPoints -> f InteractionPoints)
-> PostScopeState -> f PostScopeState
Lens' PostScopeState InteractionPoints
lensInteractionPoints
stAwakeConstraints :: Lens' TCState Constraints
stAwakeConstraints :: Lens' TCState Constraints
stAwakeConstraints = (PostScopeState -> f PostScopeState) -> TCState -> f TCState
Lens' TCState PostScopeState
lensPostScopeState ((PostScopeState -> f PostScopeState) -> TCState -> f TCState)
-> ((Constraints -> f Constraints)
-> PostScopeState -> f PostScopeState)
-> (Constraints -> f Constraints)
-> TCState
-> f TCState
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (Constraints -> f Constraints)
-> PostScopeState -> f PostScopeState
Lens' PostScopeState Constraints
lensAwakeConstraints
stSleepingConstraints :: Lens' TCState Constraints
stSleepingConstraints :: Lens' TCState Constraints
stSleepingConstraints = (PostScopeState -> f PostScopeState) -> TCState -> f TCState
Lens' TCState PostScopeState
lensPostScopeState ((PostScopeState -> f PostScopeState) -> TCState -> f TCState)
-> ((Constraints -> f Constraints)
-> PostScopeState -> f PostScopeState)
-> (Constraints -> f Constraints)
-> TCState
-> f TCState
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (Constraints -> f Constraints)
-> PostScopeState -> f PostScopeState
Lens' PostScopeState Constraints
lensSleepingConstraints
stDirty :: Lens' TCState Bool
stDirty :: Lens' TCState Bool
stDirty = (PostScopeState -> f PostScopeState) -> TCState -> f TCState
Lens' TCState PostScopeState
lensPostScopeState ((PostScopeState -> f PostScopeState) -> TCState -> f TCState)
-> ((Bool -> f Bool) -> PostScopeState -> f PostScopeState)
-> (Bool -> f Bool)
-> TCState
-> f TCState
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (Bool -> f Bool) -> PostScopeState -> f PostScopeState
Lens' PostScopeState Bool
lensDirty
stOccursCheckDefs :: Lens' TCState (Set QName)
stOccursCheckDefs :: Lens' TCState (Set QName)
stOccursCheckDefs = (PostScopeState -> f PostScopeState) -> TCState -> f TCState
Lens' TCState PostScopeState
lensPostScopeState ((PostScopeState -> f PostScopeState) -> TCState -> f TCState)
-> ((Set QName -> f (Set QName))
-> PostScopeState -> f PostScopeState)
-> (Set QName -> f (Set QName))
-> TCState
-> f TCState
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (Set QName -> f (Set QName)) -> PostScopeState -> f PostScopeState
Lens' PostScopeState (Set QName)
lensOccursCheckDefs
stSignature :: Lens' TCState Signature
stSignature :: Lens' TCState Signature
stSignature = (PostScopeState -> f PostScopeState) -> TCState -> f TCState
Lens' TCState PostScopeState
lensPostScopeState ((PostScopeState -> f PostScopeState) -> TCState -> f TCState)
-> ((Signature -> f Signature)
-> PostScopeState -> f PostScopeState)
-> (Signature -> f Signature)
-> TCState
-> f TCState
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (Signature -> f Signature) -> PostScopeState -> f PostScopeState
Lens' PostScopeState Signature
lensSignature
stModuleCheckpoints :: Lens' TCState (Map ModuleName CheckpointId)
stModuleCheckpoints :: Lens' TCState (Map ModuleName CheckpointId)
stModuleCheckpoints = (PostScopeState -> f PostScopeState) -> TCState -> f TCState
Lens' TCState PostScopeState
lensPostScopeState ((PostScopeState -> f PostScopeState) -> TCState -> f TCState)
-> ((Map ModuleName CheckpointId
-> f (Map ModuleName CheckpointId))
-> PostScopeState -> f PostScopeState)
-> (Map ModuleName CheckpointId -> f (Map ModuleName CheckpointId))
-> TCState
-> f TCState
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (Map ModuleName CheckpointId -> f (Map ModuleName CheckpointId))
-> PostScopeState -> f PostScopeState
Lens' PostScopeState (Map ModuleName CheckpointId)
lensModuleCheckpoints
stImportsDisplayForms :: Lens' TCState DisplayForms
stImportsDisplayForms :: Lens' TCState DisplayForms
stImportsDisplayForms = (PostScopeState -> f PostScopeState) -> TCState -> f TCState
Lens' TCState PostScopeState
lensPostScopeState ((PostScopeState -> f PostScopeState) -> TCState -> f TCState)
-> ((DisplayForms -> f DisplayForms)
-> PostScopeState -> f PostScopeState)
-> (DisplayForms -> f DisplayForms)
-> TCState
-> f TCState
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (DisplayForms -> f DisplayForms)
-> PostScopeState -> f PostScopeState
Lens' PostScopeState DisplayForms
lensImportsDisplayForms
stCurrentModule ::
Lens' TCState (Maybe (ModuleName, TopLevelModuleName))
stCurrentModule :: Lens' TCState (Maybe (ModuleName, TopLevelModuleName))
stCurrentModule Maybe (ModuleName, TopLevelModuleName)
-> f (Maybe (ModuleName, TopLevelModuleName))
f TCState
s =
Maybe (ModuleName, TopLevelModuleName)
-> f (Maybe (ModuleName, TopLevelModuleName))
f (PostScopeState -> Maybe (ModuleName, TopLevelModuleName)
stPostCurrentModule (TCState -> PostScopeState
stPostScopeState TCState
s)) f (Maybe (ModuleName, TopLevelModuleName))
-> (Maybe (ModuleName, TopLevelModuleName) -> TCState) -> f TCState
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&>
\Maybe (ModuleName, TopLevelModuleName)
x -> TCState
s {stPostScopeState =
(stPostScopeState s)
{stPostCurrentModule = case x of
Maybe (ModuleName, TopLevelModuleName)
Nothing -> Maybe (ModuleName, TopLevelModuleName)
forall a. Maybe a
Nothing
Just (!ModuleName
m, !TopLevelModuleName
top) -> (ModuleName, TopLevelModuleName)
-> Maybe (ModuleName, TopLevelModuleName)
forall a. a -> Maybe a
Just (ModuleName
m, TopLevelModuleName
top)}}
stInstanceDefs :: Lens' TCState TempInstanceTable
stInstanceDefs :: Lens' TCState TempInstanceTable
stInstanceDefs TempInstanceTable -> f TempInstanceTable
f TCState
s =
TempInstanceTable -> f TempInstanceTable
f ( TCState
s TCState -> Lens' TCState InstanceTable -> InstanceTable
forall o i. o -> Lens' o i -> i
^. (Signature -> f Signature) -> TCState -> f TCState
Lens' TCState Signature
stSignature ((Signature -> f Signature) -> TCState -> f TCState)
-> ((InstanceTable -> f InstanceTable) -> Signature -> f Signature)
-> (InstanceTable -> f InstanceTable)
-> TCState
-> f TCState
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (InstanceTable -> f InstanceTable) -> Signature -> f Signature
Lens' Signature InstanceTable
sigInstances
, PostScopeState -> Set QName
stPostPendingInstances (TCState -> PostScopeState
stPostScopeState TCState
s)
)
f TempInstanceTable -> (TempInstanceTable -> TCState) -> f TCState
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \(InstanceTable
t, Set QName
x) ->
Lens' TCState InstanceTable -> LensSet TCState InstanceTable
forall o i. Lens' o i -> LensSet o i
set ((Signature -> f Signature) -> TCState -> f TCState
Lens' TCState Signature
stSignature ((Signature -> f Signature) -> TCState -> f TCState)
-> ((InstanceTable -> f InstanceTable) -> Signature -> f Signature)
-> (InstanceTable -> f InstanceTable)
-> TCState
-> f TCState
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (InstanceTable -> f InstanceTable) -> Signature -> f Signature
Lens' Signature InstanceTable
sigInstances) InstanceTable
t
(TCState
s { stPostScopeState = (stPostScopeState s) { stPostPendingInstances = x }})
stTemporaryInstances :: Lens' TCState (Set QName)
stTemporaryInstances :: Lens' TCState (Set QName)
stTemporaryInstances = (PostScopeState -> f PostScopeState) -> TCState -> f TCState
Lens' TCState PostScopeState
lensPostScopeState ((PostScopeState -> f PostScopeState) -> TCState -> f TCState)
-> ((Set QName -> f (Set QName))
-> PostScopeState -> f PostScopeState)
-> (Set QName -> f (Set QName))
-> TCState
-> f TCState
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (Set QName -> f (Set QName)) -> PostScopeState -> f PostScopeState
Lens' PostScopeState (Set QName)
lensTemporaryInstances
stConcreteNames :: Lens' TCState ConcreteNames
stConcreteNames :: Lens' TCState ConcreteNames
stConcreteNames = (PostScopeState -> f PostScopeState) -> TCState -> f TCState
Lens' TCState PostScopeState
lensPostScopeState ((PostScopeState -> f PostScopeState) -> TCState -> f TCState)
-> ((ConcreteNames -> f ConcreteNames)
-> PostScopeState -> f PostScopeState)
-> (ConcreteNames -> f ConcreteNames)
-> TCState
-> f TCState
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (ConcreteNames -> f ConcreteNames)
-> PostScopeState -> f PostScopeState
Lens' PostScopeState ConcreteNames
lensConcreteNames
stUsedNames :: Lens' TCState UsedNames
stUsedNames :: Lens' TCState UsedNames
stUsedNames = (PostScopeState -> f PostScopeState) -> TCState -> f TCState
Lens' TCState PostScopeState
lensPostScopeState ((PostScopeState -> f PostScopeState) -> TCState -> f TCState)
-> ((UsedNames -> f UsedNames)
-> PostScopeState -> f PostScopeState)
-> (UsedNames -> f UsedNames)
-> TCState
-> f TCState
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (UsedNames -> f UsedNames) -> PostScopeState -> f PostScopeState
Lens' PostScopeState UsedNames
lensUsedNames
stShadowingNames :: Lens' TCState ShadowingNames
stShadowingNames :: Lens' TCState ShadowingNames
stShadowingNames = (PostScopeState -> f PostScopeState) -> TCState -> f TCState
Lens' TCState PostScopeState
lensPostScopeState ((PostScopeState -> f PostScopeState) -> TCState -> f TCState)
-> ((ShadowingNames -> f ShadowingNames)
-> PostScopeState -> f PostScopeState)
-> (ShadowingNames -> f ShadowingNames)
-> TCState
-> f TCState
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (ShadowingNames -> f ShadowingNames)
-> PostScopeState -> f PostScopeState
Lens' PostScopeState ShadowingNames
lensShadowingNames
stStatistics :: Lens' TCState Statistics
stStatistics :: Lens' TCState Statistics
stStatistics = (PostScopeState -> f PostScopeState) -> TCState -> f TCState
Lens' TCState PostScopeState
lensPostScopeState ((PostScopeState -> f PostScopeState) -> TCState -> f TCState)
-> ((Statistics -> f Statistics)
-> PostScopeState -> f PostScopeState)
-> (Statistics -> f Statistics)
-> TCState
-> f TCState
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (Statistics -> f Statistics) -> PostScopeState -> f PostScopeState
Lens' PostScopeState Statistics
lensStatistics
stTCWarnings :: Lens' TCState (Set TCWarning)
stTCWarnings :: Lens' TCState (Set TCWarning)
stTCWarnings = (PostScopeState -> f PostScopeState) -> TCState -> f TCState
Lens' TCState PostScopeState
lensPostScopeState ((PostScopeState -> f PostScopeState) -> TCState -> f TCState)
-> ((Set TCWarning -> f (Set TCWarning))
-> PostScopeState -> f PostScopeState)
-> (Set TCWarning -> f (Set TCWarning))
-> TCState
-> f TCState
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (Set TCWarning -> f (Set TCWarning))
-> PostScopeState -> f PostScopeState
Lens' PostScopeState (Set TCWarning)
lensTCWarnings
stMutualBlocks :: Lens' TCState MutualBlocks
stMutualBlocks :: Lens' TCState MutualBlocks
stMutualBlocks = (PostScopeState -> f PostScopeState) -> TCState -> f TCState
Lens' TCState PostScopeState
lensPostScopeState ((PostScopeState -> f PostScopeState) -> TCState -> f TCState)
-> ((MutualBlocks -> f MutualBlocks)
-> PostScopeState -> f PostScopeState)
-> (MutualBlocks -> f MutualBlocks)
-> TCState
-> f TCState
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (MutualBlocks -> f MutualBlocks)
-> PostScopeState -> f PostScopeState
Lens' PostScopeState MutualBlocks
lensMutualBlocks
stLocalBuiltins :: Lens' TCState BuiltinThings
stLocalBuiltins :: Lens' TCState BuiltinThings
stLocalBuiltins = (PostScopeState -> f PostScopeState) -> TCState -> f TCState
Lens' TCState PostScopeState
lensPostScopeState ((PostScopeState -> f PostScopeState) -> TCState -> f TCState)
-> ((BuiltinThings -> f BuiltinThings)
-> PostScopeState -> f PostScopeState)
-> (BuiltinThings -> f BuiltinThings)
-> TCState
-> f TCState
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (BuiltinThings -> f BuiltinThings)
-> PostScopeState -> f PostScopeState
Lens' PostScopeState BuiltinThings
lensLocalBuiltins
stFreshMetaId :: Lens' TCState MetaId
stFreshMetaId :: Lens' TCState MetaId
stFreshMetaId = (PostScopeState -> f PostScopeState) -> TCState -> f TCState
Lens' TCState PostScopeState
lensPostScopeState ((PostScopeState -> f PostScopeState) -> TCState -> f TCState)
-> ((MetaId -> f MetaId) -> PostScopeState -> f PostScopeState)
-> (MetaId -> f MetaId)
-> TCState
-> f TCState
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (MetaId -> f MetaId) -> PostScopeState -> f PostScopeState
Lens' PostScopeState MetaId
lensFreshMetaId
stFreshMutualId :: Lens' TCState MutualId
stFreshMutualId :: Lens' TCState MutualId
stFreshMutualId = (PostScopeState -> f PostScopeState) -> TCState -> f TCState
Lens' TCState PostScopeState
lensPostScopeState ((PostScopeState -> f PostScopeState) -> TCState -> f TCState)
-> ((MutualId -> f MutualId) -> PostScopeState -> f PostScopeState)
-> (MutualId -> f MutualId)
-> TCState
-> f TCState
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (MutualId -> f MutualId) -> PostScopeState -> f PostScopeState
Lens' PostScopeState MutualId
lensFreshMutualId
stFreshProblemId :: Lens' TCState ProblemId
stFreshProblemId :: Lens' TCState ProblemId
stFreshProblemId = (PostScopeState -> f PostScopeState) -> TCState -> f TCState
Lens' TCState PostScopeState
lensPostScopeState ((PostScopeState -> f PostScopeState) -> TCState -> f TCState)
-> ((ProblemId -> f ProblemId)
-> PostScopeState -> f PostScopeState)
-> (ProblemId -> f ProblemId)
-> TCState
-> f TCState
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (ProblemId -> f ProblemId) -> PostScopeState -> f PostScopeState
Lens' PostScopeState ProblemId
lensFreshProblemId
stFreshCheckpointId :: Lens' TCState CheckpointId
stFreshCheckpointId :: Lens' TCState CheckpointId
stFreshCheckpointId = (PostScopeState -> f PostScopeState) -> TCState -> f TCState
Lens' TCState PostScopeState
lensPostScopeState ((PostScopeState -> f PostScopeState) -> TCState -> f TCState)
-> ((CheckpointId -> f CheckpointId)
-> PostScopeState -> f PostScopeState)
-> (CheckpointId -> f CheckpointId)
-> TCState
-> f TCState
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (CheckpointId -> f CheckpointId)
-> PostScopeState -> f PostScopeState
Lens' PostScopeState CheckpointId
lensFreshCheckpointId
stFreshInt :: Lens' TCState Int
stFreshInt :: Lens' TCState Int
stFreshInt = (PostScopeState -> f PostScopeState) -> TCState -> f TCState
Lens' TCState PostScopeState
lensPostScopeState ((PostScopeState -> f PostScopeState) -> TCState -> f TCState)
-> ((Int -> f Int) -> PostScopeState -> f PostScopeState)
-> (Int -> f Int)
-> TCState
-> f TCState
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (Int -> f Int) -> PostScopeState -> f PostScopeState
Lens' PostScopeState Int
lensFreshInt
stAreWeCaching :: Lens' TCState Bool
stAreWeCaching :: Lens' TCState Bool
stAreWeCaching = (PostScopeState -> f PostScopeState) -> TCState -> f TCState
Lens' TCState PostScopeState
lensPostScopeState ((PostScopeState -> f PostScopeState) -> TCState -> f TCState)
-> ((Bool -> f Bool) -> PostScopeState -> f PostScopeState)
-> (Bool -> f Bool)
-> TCState
-> f TCState
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (Bool -> f Bool) -> PostScopeState -> f PostScopeState
Lens' PostScopeState Bool
lensAreWeCaching
stPostponeInstanceSearch :: Lens' TCState Bool
stPostponeInstanceSearch :: Lens' TCState Bool
stPostponeInstanceSearch = (PostScopeState -> f PostScopeState) -> TCState -> f TCState
Lens' TCState PostScopeState
lensPostScopeState ((PostScopeState -> f PostScopeState) -> TCState -> f TCState)
-> ((Bool -> f Bool) -> PostScopeState -> f PostScopeState)
-> (Bool -> f Bool)
-> TCState
-> f TCState
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (Bool -> f Bool) -> PostScopeState -> f PostScopeState
Lens' PostScopeState Bool
lensPostponeInstanceSearch
stConsideringInstance :: Lens' TCState Bool
stConsideringInstance :: Lens' TCState Bool
stConsideringInstance = (PostScopeState -> f PostScopeState) -> TCState -> f TCState
Lens' TCState PostScopeState
lensPostScopeState ((PostScopeState -> f PostScopeState) -> TCState -> f TCState)
-> ((Bool -> f Bool) -> PostScopeState -> f PostScopeState)
-> (Bool -> f Bool)
-> TCState
-> f TCState
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (Bool -> f Bool) -> PostScopeState -> f PostScopeState
Lens' PostScopeState Bool
lensConsideringInstance
stInstantiateBlocking :: Lens' TCState Bool
stInstantiateBlocking :: Lens' TCState Bool
stInstantiateBlocking = (PostScopeState -> f PostScopeState) -> TCState -> f TCState
Lens' TCState PostScopeState
lensPostScopeState ((PostScopeState -> f PostScopeState) -> TCState -> f TCState)
-> ((Bool -> f Bool) -> PostScopeState -> f PostScopeState)
-> (Bool -> f Bool)
-> TCState
-> f TCState
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (Bool -> f Bool) -> PostScopeState -> f PostScopeState
Lens' PostScopeState Bool
lensInstantiateBlocking
stInstanceTree :: Lens' TCState (DiscrimTree QName)
stInstanceTree :: Lens' TCState (DiscrimTree QName)
stInstanceTree = (Signature -> f Signature) -> TCState -> f TCState
Lens' TCState Signature
stSignature ((Signature -> f Signature) -> TCState -> f TCState)
-> ((DiscrimTree QName -> f (DiscrimTree QName))
-> Signature -> f Signature)
-> (DiscrimTree QName -> f (DiscrimTree QName))
-> TCState
-> f TCState
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (InstanceTable -> f InstanceTable) -> Signature -> f Signature
Lens' Signature InstanceTable
sigInstances ((InstanceTable -> f InstanceTable) -> Signature -> f Signature)
-> ((DiscrimTree QName -> f (DiscrimTree QName))
-> InstanceTable -> f InstanceTable)
-> (DiscrimTree QName -> f (DiscrimTree QName))
-> Signature
-> f Signature
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (DiscrimTree QName -> f (DiscrimTree QName))
-> InstanceTable -> f InstanceTable
Lens' InstanceTable (DiscrimTree QName)
itableTree
stBuiltinThings :: TCState -> BuiltinThings
stBuiltinThings :: TCState -> BuiltinThings
stBuiltinThings TCState
s = (Builtin PrimFun -> Builtin PrimFun -> Builtin PrimFun)
-> BuiltinThings -> BuiltinThings -> BuiltinThings
forall k a. Ord k => (a -> a -> a) -> Map k a -> Map k a -> Map k a
Map.unionWith Builtin PrimFun -> Builtin PrimFun -> Builtin PrimFun
forall a. Builtin a -> Builtin a -> Builtin a
unionBuiltin (TCState
s TCState -> Lens' TCState BuiltinThings -> BuiltinThings
forall o i. o -> Lens' o i -> i
^. (BuiltinThings -> f BuiltinThings) -> TCState -> f TCState
Lens' TCState BuiltinThings
stLocalBuiltins) (TCState
s TCState -> Lens' TCState BuiltinThings -> BuiltinThings
forall o i. o -> Lens' o i -> i
^. (BuiltinThings -> f BuiltinThings) -> TCState -> f TCState
Lens' TCState BuiltinThings
stImportedBuiltins)
unionBuiltin :: Builtin a -> Builtin a -> Builtin a
unionBuiltin :: forall a. Builtin a -> Builtin a -> Builtin a
unionBuiltin = ((Builtin a, Builtin a) -> Builtin a)
-> Builtin a -> Builtin a -> Builtin a
forall a b c. ((a, b) -> c) -> a -> b -> c
curry (((Builtin a, Builtin a) -> Builtin a)
-> Builtin a -> Builtin a -> Builtin a)
-> ((Builtin a, Builtin a) -> Builtin a)
-> Builtin a
-> Builtin a
-> Builtin a
forall a b. (a -> b) -> a -> b
$ \case
(BuiltinRewriteRelations Set QName
xs, BuiltinRewriteRelations Set QName
ys) -> Set QName -> Builtin a
forall pf. Set QName -> Builtin pf
BuiltinRewriteRelations (Set QName -> Builtin a) -> Set QName -> Builtin a
forall a b. (a -> b) -> a -> b
$ Set QName
xs Set QName -> Set QName -> Set QName
forall a. Semigroup a => a -> a -> a
<> Set QName
ys
(Builtin a, Builtin a)
_ -> Builtin a
forall a. HasCallStack => a
__IMPOSSIBLE__
class Enum i => HasFresh i where
freshLens :: Lens' TCState i
nextFresh' :: i -> i
nextFresh' = i -> i
forall a. Enum a => a -> a
succ
{-# INLINE nextFresh #-}
nextFresh :: HasFresh i => TCState -> (i, TCState)
nextFresh :: forall i. HasFresh i => TCState -> (i, TCState)
nextFresh TCState
s =
let !c :: i
c = TCState
s TCState -> Lens' TCState i -> i
forall o i. o -> Lens' o i -> i
^. (i -> f i) -> TCState -> f TCState
forall i. HasFresh i => Lens' TCState i
Lens' TCState i
freshLens
!next :: TCState
next = Lens' TCState i -> LensSet TCState i
forall o i. Lens' o i -> LensSet o i
set (i -> f i) -> TCState -> f TCState
forall i. HasFresh i => Lens' TCState i
Lens' TCState i
freshLens (i -> i
forall i. HasFresh i => i -> i
nextFresh' i
c) TCState
s
in (i
c, TCState
next)
class Monad m => MonadFresh i m where
fresh :: m i
default fresh :: (MonadTrans t, MonadFresh i n, t n ~ m) => m i
fresh = n i -> t n i
forall (m :: * -> *) a. Monad m => m a -> t m a
forall (t :: (* -> *) -> * -> *) (m :: * -> *) a.
(MonadTrans t, Monad m) =>
m a -> t m a
lift n i
forall i (m :: * -> *). MonadFresh i m => m i
fresh
instance MonadFresh i m => MonadFresh i (ExceptT e m)
instance MonadFresh i m => MonadFresh i (MaybeT m)
instance MonadFresh i m => MonadFresh i (ReaderT r m)
instance MonadFresh i m => MonadFresh i (StateT s m)
instance (MonadFresh i m, Monoid w) => MonadFresh i (WriterT w m)
instance MonadFresh i m => MonadFresh i (ListT m)
instance MonadFresh i m => MonadFresh i (IdentityT m)
instance HasFresh i => MonadFresh i TCM where
fresh :: TCM i
fresh = do
!s <- TCMT IO TCState
forall (m :: * -> *). MonadTCState m => m TCState
getTC
let (!c , !s') = nextFresh s
putTC s'
return c
{-# INLINE fresh #-}
instance HasFresh MetaId where
freshLens :: Lens' TCState MetaId
freshLens = (MetaId -> f MetaId) -> TCState -> f TCState
Lens' TCState MetaId
stFreshMetaId
instance HasFresh MutualId where
freshLens :: Lens' TCState MutualId
freshLens = (MutualId -> f MutualId) -> TCState -> f TCState
Lens' TCState MutualId
stFreshMutualId
instance HasFresh InteractionId where
freshLens :: Lens' TCState InteractionId
freshLens = (InteractionId -> f InteractionId) -> TCState -> f TCState
Lens' TCState InteractionId
stFreshInteractionId
instance HasFresh NameId where
freshLens :: Lens' TCState NameId
freshLens = (NameId -> f NameId) -> TCState -> f TCState
Lens' TCState NameId
stFreshNameId
nextFresh' :: NameId -> NameId
nextFresh' = NameId -> NameId
forall a. Enum a => a -> a
succ (NameId -> NameId) -> (NameId -> NameId) -> NameId -> NameId
forall b c a. (b -> c) -> (a -> b) -> a -> c
. NameId -> NameId
forall a. Enum a => a -> a
succ
instance HasFresh OpaqueId where
freshLens :: Lens' TCState OpaqueId
freshLens = (OpaqueId -> f OpaqueId) -> TCState -> f TCState
Lens' TCState OpaqueId
stFreshOpaqueId
instance HasFresh Int where
freshLens :: Lens' TCState Int
freshLens = (Int -> f Int) -> TCState -> f TCState
Lens' TCState Int
stFreshInt
instance HasFresh ProblemId where
freshLens :: Lens' TCState ProblemId
freshLens = (ProblemId -> f ProblemId) -> TCState -> f TCState
Lens' TCState ProblemId
stFreshProblemId
newtype CheckpointId = CheckpointId Int
deriving (CheckpointId -> CheckpointId -> Bool
(CheckpointId -> CheckpointId -> Bool)
-> (CheckpointId -> CheckpointId -> Bool) -> Eq CheckpointId
forall a. (a -> a -> Bool) -> (a -> a -> Bool) -> Eq a
$c== :: CheckpointId -> CheckpointId -> Bool
== :: CheckpointId -> CheckpointId -> Bool
$c/= :: CheckpointId -> CheckpointId -> Bool
/= :: CheckpointId -> CheckpointId -> Bool
Eq, Eq CheckpointId
Eq CheckpointId =>
(CheckpointId -> CheckpointId -> Ordering)
-> (CheckpointId -> CheckpointId -> Bool)
-> (CheckpointId -> CheckpointId -> Bool)
-> (CheckpointId -> CheckpointId -> Bool)
-> (CheckpointId -> CheckpointId -> Bool)
-> (CheckpointId -> CheckpointId -> CheckpointId)
-> (CheckpointId -> CheckpointId -> CheckpointId)
-> Ord CheckpointId
CheckpointId -> CheckpointId -> Bool
CheckpointId -> CheckpointId -> Ordering
CheckpointId -> CheckpointId -> CheckpointId
forall a.
Eq a =>
(a -> a -> Ordering)
-> (a -> a -> Bool)
-> (a -> a -> Bool)
-> (a -> a -> Bool)
-> (a -> a -> Bool)
-> (a -> a -> a)
-> (a -> a -> a)
-> Ord a
$ccompare :: CheckpointId -> CheckpointId -> Ordering
compare :: CheckpointId -> CheckpointId -> Ordering
$c< :: CheckpointId -> CheckpointId -> Bool
< :: CheckpointId -> CheckpointId -> Bool
$c<= :: CheckpointId -> CheckpointId -> Bool
<= :: CheckpointId -> CheckpointId -> Bool
$c> :: CheckpointId -> CheckpointId -> Bool
> :: CheckpointId -> CheckpointId -> Bool
$c>= :: CheckpointId -> CheckpointId -> Bool
>= :: CheckpointId -> CheckpointId -> Bool
$cmax :: CheckpointId -> CheckpointId -> CheckpointId
max :: CheckpointId -> CheckpointId -> CheckpointId
$cmin :: CheckpointId -> CheckpointId -> CheckpointId
min :: CheckpointId -> CheckpointId -> CheckpointId
Ord, Int -> CheckpointId
CheckpointId -> Int
CheckpointId -> [CheckpointId]
CheckpointId -> CheckpointId
CheckpointId -> CheckpointId -> [CheckpointId]
CheckpointId -> CheckpointId -> CheckpointId -> [CheckpointId]
(CheckpointId -> CheckpointId)
-> (CheckpointId -> CheckpointId)
-> (Int -> CheckpointId)
-> (CheckpointId -> Int)
-> (CheckpointId -> [CheckpointId])
-> (CheckpointId -> CheckpointId -> [CheckpointId])
-> (CheckpointId -> CheckpointId -> [CheckpointId])
-> (CheckpointId -> CheckpointId -> CheckpointId -> [CheckpointId])
-> Enum CheckpointId
forall a.
(a -> a)
-> (a -> a)
-> (Int -> a)
-> (a -> Int)
-> (a -> [a])
-> (a -> a -> [a])
-> (a -> a -> [a])
-> (a -> a -> a -> [a])
-> Enum a
$csucc :: CheckpointId -> CheckpointId
succ :: CheckpointId -> CheckpointId
$cpred :: CheckpointId -> CheckpointId
pred :: CheckpointId -> CheckpointId
$ctoEnum :: Int -> CheckpointId
toEnum :: Int -> CheckpointId
$cfromEnum :: CheckpointId -> Int
fromEnum :: CheckpointId -> Int
$cenumFrom :: CheckpointId -> [CheckpointId]
enumFrom :: CheckpointId -> [CheckpointId]
$cenumFromThen :: CheckpointId -> CheckpointId -> [CheckpointId]
enumFromThen :: CheckpointId -> CheckpointId -> [CheckpointId]
$cenumFromTo :: CheckpointId -> CheckpointId -> [CheckpointId]
enumFromTo :: CheckpointId -> CheckpointId -> [CheckpointId]
$cenumFromThenTo :: CheckpointId -> CheckpointId -> CheckpointId -> [CheckpointId]
enumFromThenTo :: CheckpointId -> CheckpointId -> CheckpointId -> [CheckpointId]
Enum, Num CheckpointId
Ord CheckpointId
(Num CheckpointId, Ord CheckpointId) =>
(CheckpointId -> Rational) -> Real CheckpointId
CheckpointId -> Rational
forall a. (Num a, Ord a) => (a -> Rational) -> Real a
$ctoRational :: CheckpointId -> Rational
toRational :: CheckpointId -> Rational
Real, Enum CheckpointId
Real CheckpointId
(Real CheckpointId, Enum CheckpointId) =>
(CheckpointId -> CheckpointId -> CheckpointId)
-> (CheckpointId -> CheckpointId -> CheckpointId)
-> (CheckpointId -> CheckpointId -> CheckpointId)
-> (CheckpointId -> CheckpointId -> CheckpointId)
-> (CheckpointId -> CheckpointId -> (CheckpointId, CheckpointId))
-> (CheckpointId -> CheckpointId -> (CheckpointId, CheckpointId))
-> (CheckpointId -> Integer)
-> Integral CheckpointId
CheckpointId -> Integer
CheckpointId -> CheckpointId -> (CheckpointId, CheckpointId)
CheckpointId -> CheckpointId -> CheckpointId
forall a.
(Real a, Enum a) =>
(a -> a -> a)
-> (a -> a -> a)
-> (a -> a -> a)
-> (a -> a -> a)
-> (a -> a -> (a, a))
-> (a -> a -> (a, a))
-> (a -> Integer)
-> Integral a
$cquot :: CheckpointId -> CheckpointId -> CheckpointId
quot :: CheckpointId -> CheckpointId -> CheckpointId
$crem :: CheckpointId -> CheckpointId -> CheckpointId
rem :: CheckpointId -> CheckpointId -> CheckpointId
$cdiv :: CheckpointId -> CheckpointId -> CheckpointId
div :: CheckpointId -> CheckpointId -> CheckpointId
$cmod :: CheckpointId -> CheckpointId -> CheckpointId
mod :: CheckpointId -> CheckpointId -> CheckpointId
$cquotRem :: CheckpointId -> CheckpointId -> (CheckpointId, CheckpointId)
quotRem :: CheckpointId -> CheckpointId -> (CheckpointId, CheckpointId)
$cdivMod :: CheckpointId -> CheckpointId -> (CheckpointId, CheckpointId)
divMod :: CheckpointId -> CheckpointId -> (CheckpointId, CheckpointId)
$ctoInteger :: CheckpointId -> Integer
toInteger :: CheckpointId -> Integer
Integral, Integer -> CheckpointId
CheckpointId -> CheckpointId
CheckpointId -> CheckpointId -> CheckpointId
(CheckpointId -> CheckpointId -> CheckpointId)
-> (CheckpointId -> CheckpointId -> CheckpointId)
-> (CheckpointId -> CheckpointId -> CheckpointId)
-> (CheckpointId -> CheckpointId)
-> (CheckpointId -> CheckpointId)
-> (CheckpointId -> CheckpointId)
-> (Integer -> CheckpointId)
-> Num CheckpointId
forall a.
(a -> a -> a)
-> (a -> a -> a)
-> (a -> a -> a)
-> (a -> a)
-> (a -> a)
-> (a -> a)
-> (Integer -> a)
-> Num a
$c+ :: CheckpointId -> CheckpointId -> CheckpointId
+ :: CheckpointId -> CheckpointId -> CheckpointId
$c- :: CheckpointId -> CheckpointId -> CheckpointId
- :: CheckpointId -> CheckpointId -> CheckpointId
$c* :: CheckpointId -> CheckpointId -> CheckpointId
* :: CheckpointId -> CheckpointId -> CheckpointId
$cnegate :: CheckpointId -> CheckpointId
negate :: CheckpointId -> CheckpointId
$cabs :: CheckpointId -> CheckpointId
abs :: CheckpointId -> CheckpointId
$csignum :: CheckpointId -> CheckpointId
signum :: CheckpointId -> CheckpointId
$cfromInteger :: Integer -> CheckpointId
fromInteger :: Integer -> CheckpointId
Num, CheckpointId -> ()
(CheckpointId -> ()) -> NFData CheckpointId
forall a. (a -> ()) -> NFData a
$crnf :: CheckpointId -> ()
rnf :: CheckpointId -> ()
NFData)
instance Show CheckpointId where
show :: CheckpointId -> String
show (CheckpointId Int
n) = Int -> String
forall a. Show a => a -> String
show Int
n
instance Pretty CheckpointId where
pretty :: CheckpointId -> Doc
pretty (CheckpointId Int
n) = Int -> Doc
forall a. Pretty a => a -> Doc
pretty Int
n
instance HasFresh CheckpointId where
freshLens :: Lens' TCState CheckpointId
freshLens = (CheckpointId -> f CheckpointId) -> TCState -> f TCState
Lens' TCState CheckpointId
stFreshCheckpointId
freshName :: MonadFresh NameId m => Range -> String -> m Name
freshName :: forall (m :: * -> *).
MonadFresh NameId m =>
Range -> String -> m Name
freshName Range
r String
s = do
i <- m NameId
forall i (m :: * -> *). MonadFresh i m => m i
fresh
return $ mkName r i s
freshNoName :: MonadFresh NameId m => Range -> m Name
freshNoName :: forall (m :: * -> *). MonadFresh NameId m => Range -> m Name
freshNoName Range
r =
do i <- m NameId
forall i (m :: * -> *). MonadFresh i m => m i
fresh
return $ makeName i (C.NoName noRange i) r noFixity' False
freshNoName_ :: MonadFresh NameId m => m Name
freshNoName_ :: forall (m :: * -> *). MonadFresh NameId m => m Name
freshNoName_ = Range -> m Name
forall (m :: * -> *). MonadFresh NameId m => Range -> m Name
freshNoName Range
forall a. Range' a
noRange
freshRecordName :: MonadFresh NameId m => m Name
freshRecordName :: forall (m :: * -> *). MonadFresh NameId m => m Name
freshRecordName = do
i <- m NameId
forall i (m :: * -> *). MonadFresh i m => m i
fresh
return $ makeName i (C.setNotInScope $ C.simpleName "r") noRange noFixity' True
class FreshName a where
freshName_ :: MonadFresh NameId m => a -> m Name
instance FreshName (Range, String) where
freshName_ :: forall (m :: * -> *).
MonadFresh NameId m =>
(Range, String) -> m Name
freshName_ = (Range -> String -> m Name) -> (Range, String) -> m Name
forall a b c. (a -> b -> c) -> (a, b) -> c
uncurry Range -> String -> m Name
forall (m :: * -> *).
MonadFresh NameId m =>
Range -> String -> m Name
freshName
instance FreshName String where
freshName_ :: forall (m :: * -> *). MonadFresh NameId m => String -> m Name
freshName_ = Range -> String -> m Name
forall (m :: * -> *).
MonadFresh NameId m =>
Range -> String -> m Name
freshName Range
forall a. Range' a
noRange
instance FreshName Range where
freshName_ :: forall (m :: * -> *). MonadFresh NameId m => Range -> m Name
freshName_ = Range -> m Name
forall (m :: * -> *). MonadFresh NameId m => Range -> m Name
freshNoName
instance FreshName () where
freshName_ :: forall (m :: * -> *). MonadFresh NameId m => () -> m Name
freshName_ () = m Name
forall (m :: * -> *). MonadFresh NameId m => m Name
freshNoName_
instance FreshName Name where
freshName_ :: forall (m :: * -> *). MonadFresh NameId m => Name -> m Name
freshName_ (Name NameId
_ Name
con Name
can Range
bs Fixity'
fix Bool
rn) = do
i <- m NameId
forall i (m :: * -> *). MonadFresh i m => m i
fresh
pure $ Name i con can bs fix rn
class Monad m => MonadStConcreteNames m where
runStConcreteNames :: StateT ConcreteNames m a -> m a
useConcreteNames :: m ConcreteNames
useConcreteNames = StateT ConcreteNames m ConcreteNames -> m ConcreteNames
forall a. StateT ConcreteNames m a -> m a
forall (m :: * -> *) a.
MonadStConcreteNames m =>
StateT ConcreteNames m a -> m a
runStConcreteNames StateT ConcreteNames m ConcreteNames
forall s (m :: * -> *). MonadState s m => m s
get
modifyConcreteNames :: (ConcreteNames -> ConcreteNames) -> m ()
modifyConcreteNames = StateT ConcreteNames m () -> m ()
forall a. StateT ConcreteNames m a -> m a
forall (m :: * -> *) a.
MonadStConcreteNames m =>
StateT ConcreteNames m a -> m a
runStConcreteNames (StateT ConcreteNames m () -> m ())
-> ((ConcreteNames -> ConcreteNames) -> StateT ConcreteNames m ())
-> (ConcreteNames -> ConcreteNames)
-> m ()
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (ConcreteNames -> ConcreteNames) -> StateT ConcreteNames m ()
forall s (m :: * -> *). MonadState s m => (s -> s) -> m ()
modify
instance MonadStConcreteNames TCM where
runStConcreteNames :: forall a. StateT ConcreteNames (TCMT IO) a -> TCM a
runStConcreteNames StateT ConcreteNames (TCMT IO) a
m = Lens' TCState ConcreteNames
-> (ConcreteNames -> TCMT IO (a, ConcreteNames)) -> TCMT IO a
forall (m :: * -> *) a r.
(MonadTCState m, ReadTCState m) =>
Lens' TCState a -> (a -> m (r, a)) -> m r
stateTCLensM (ConcreteNames -> f ConcreteNames) -> TCState -> f TCState
Lens' TCState ConcreteNames
stConcreteNames ((ConcreteNames -> TCMT IO (a, ConcreteNames)) -> TCMT IO a)
-> (ConcreteNames -> TCMT IO (a, ConcreteNames)) -> TCMT IO a
forall a b. (a -> b) -> a -> b
$ StateT ConcreteNames (TCMT IO) a
-> ConcreteNames -> TCMT IO (a, ConcreteNames)
forall s (m :: * -> *) a. StateT s m a -> s -> m (a, s)
runStateT StateT ConcreteNames (TCMT IO) a
m
instance MonadStConcreteNames m => MonadStConcreteNames (ExceptT e m) where
runStConcreteNames :: forall a. StateT ConcreteNames (ExceptT e m) a -> ExceptT e m a
runStConcreteNames StateT ConcreteNames (ExceptT e m) a
m = m (Either e a) -> ExceptT e m a
forall e (m :: * -> *) a. m (Either e a) -> ExceptT e m a
ExceptT (m (Either e a) -> ExceptT e m a)
-> m (Either e a) -> ExceptT e m a
forall a b. (a -> b) -> a -> b
$ StateT ConcreteNames m (Either e a) -> m (Either e a)
forall a. StateT ConcreteNames m a -> m a
forall (m :: * -> *) a.
MonadStConcreteNames m =>
StateT ConcreteNames m a -> m a
runStConcreteNames (StateT ConcreteNames m (Either e a) -> m (Either e a))
-> StateT ConcreteNames m (Either e a) -> m (Either e a)
forall a b. (a -> b) -> a -> b
$ (ConcreteNames -> m (Either e a, ConcreteNames))
-> StateT ConcreteNames m (Either e a)
forall s (m :: * -> *) a. (s -> m (a, s)) -> StateT s m a
StateT ((ConcreteNames -> m (Either e a, ConcreteNames))
-> StateT ConcreteNames m (Either e a))
-> (ConcreteNames -> m (Either e a, ConcreteNames))
-> StateT ConcreteNames m (Either e a)
forall a b. (a -> b) -> a -> b
$ \ ConcreteNames
ns -> do
ExceptT e m (a, ConcreteNames) -> m (Either e (a, ConcreteNames))
forall e (m :: * -> *) a. ExceptT e m a -> m (Either e a)
runExceptT (StateT ConcreteNames (ExceptT e m) a
-> ConcreteNames -> ExceptT e m (a, ConcreteNames)
forall s (m :: * -> *) a. StateT s m a -> s -> m (a, s)
runStateT StateT ConcreteNames (ExceptT e m) a
m ConcreteNames
ns) m (Either e (a, ConcreteNames))
-> (Either e (a, ConcreteNames) -> (Either e a, ConcreteNames))
-> m (Either e a, ConcreteNames)
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \case
Left e
e -> (e -> Either e a
forall a b. a -> Either a b
Left e
e, ConcreteNames
forall a. Monoid a => a
mempty)
Right (a
x, ConcreteNames
ns') -> (a -> Either e a
forall a b. b -> Either a b
Right a
x, ConcreteNames
ns')
instance MonadStConcreteNames m => MonadStConcreteNames (IdentityT m) where
runStConcreteNames :: forall a. StateT ConcreteNames (IdentityT m) a -> IdentityT m a
runStConcreteNames StateT ConcreteNames (IdentityT m) a
m = m a -> IdentityT m a
forall {k} (f :: k -> *) (a :: k). f a -> IdentityT f a
IdentityT (m a -> IdentityT m a) -> m a -> IdentityT m a
forall a b. (a -> b) -> a -> b
$ StateT ConcreteNames m a -> m a
forall a. StateT ConcreteNames m a -> m a
forall (m :: * -> *) a.
MonadStConcreteNames m =>
StateT ConcreteNames m a -> m a
runStConcreteNames (StateT ConcreteNames m a -> m a)
-> StateT ConcreteNames m a -> m a
forall a b. (a -> b) -> a -> b
$ (ConcreteNames -> m (a, ConcreteNames)) -> StateT ConcreteNames m a
forall s (m :: * -> *) a. (s -> m (a, s)) -> StateT s m a
StateT ((ConcreteNames -> m (a, ConcreteNames))
-> StateT ConcreteNames m a)
-> (ConcreteNames -> m (a, ConcreteNames))
-> StateT ConcreteNames m a
forall a b. (a -> b) -> a -> b
$ IdentityT m (a, ConcreteNames) -> m (a, ConcreteNames)
forall {k} (f :: k -> *) (a :: k). IdentityT f a -> f a
runIdentityT (IdentityT m (a, ConcreteNames) -> m (a, ConcreteNames))
-> (ConcreteNames -> IdentityT m (a, ConcreteNames))
-> ConcreteNames
-> m (a, ConcreteNames)
forall b c a. (b -> c) -> (a -> b) -> a -> c
. StateT ConcreteNames (IdentityT m) a
-> ConcreteNames -> IdentityT m (a, ConcreteNames)
forall s (m :: * -> *) a. StateT s m a -> s -> m (a, s)
runStateT StateT ConcreteNames (IdentityT m) a
m
instance MonadStConcreteNames m => MonadStConcreteNames (MaybeT m) where
runStConcreteNames :: forall a. StateT ConcreteNames (MaybeT m) a -> MaybeT m a
runStConcreteNames StateT ConcreteNames (MaybeT m) a
m = m (Maybe a) -> MaybeT m a
forall (m :: * -> *) a. m (Maybe a) -> MaybeT m a
MaybeT (m (Maybe a) -> MaybeT m a) -> m (Maybe a) -> MaybeT m a
forall a b. (a -> b) -> a -> b
$ StateT ConcreteNames m (Maybe a) -> m (Maybe a)
forall a. StateT ConcreteNames m a -> m a
forall (m :: * -> *) a.
MonadStConcreteNames m =>
StateT ConcreteNames m a -> m a
runStConcreteNames (StateT ConcreteNames m (Maybe a) -> m (Maybe a))
-> StateT ConcreteNames m (Maybe a) -> m (Maybe a)
forall a b. (a -> b) -> a -> b
$ (ConcreteNames -> m (Maybe a, ConcreteNames))
-> StateT ConcreteNames m (Maybe a)
forall s (m :: * -> *) a. (s -> m (a, s)) -> StateT s m a
StateT ((ConcreteNames -> m (Maybe a, ConcreteNames))
-> StateT ConcreteNames m (Maybe a))
-> (ConcreteNames -> m (Maybe a, ConcreteNames))
-> StateT ConcreteNames m (Maybe a)
forall a b. (a -> b) -> a -> b
$ \ ConcreteNames
ns -> do
MaybeT m (a, ConcreteNames) -> m (Maybe (a, ConcreteNames))
forall (m :: * -> *) a. MaybeT m a -> m (Maybe a)
runMaybeT (StateT ConcreteNames (MaybeT m) a
-> ConcreteNames -> MaybeT m (a, ConcreteNames)
forall s (m :: * -> *) a. StateT s m a -> s -> m (a, s)
runStateT StateT ConcreteNames (MaybeT m) a
m ConcreteNames
ns) m (Maybe (a, ConcreteNames))
-> (Maybe (a, ConcreteNames) -> (Maybe a, ConcreteNames))
-> m (Maybe a, ConcreteNames)
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \case
Maybe (a, ConcreteNames)
Nothing -> (Maybe a
forall a. Maybe a
Nothing, ConcreteNames
forall a. Monoid a => a
mempty)
Just (a
x, ConcreteNames
ns') -> (a -> Maybe a
forall a. a -> Maybe a
Just a
x, ConcreteNames
ns')
instance MonadStConcreteNames m => MonadStConcreteNames (ReaderT r m) where
runStConcreteNames :: forall a. StateT ConcreteNames (ReaderT r m) a -> ReaderT r m a
runStConcreteNames StateT ConcreteNames (ReaderT r m) a
m = (r -> m a) -> ReaderT r m a
forall r (m :: * -> *) a. (r -> m a) -> ReaderT r m a
ReaderT ((r -> m a) -> ReaderT r m a) -> (r -> m a) -> ReaderT r m a
forall a b. (a -> b) -> a -> b
$ StateT ConcreteNames m a -> m a
forall a. StateT ConcreteNames m a -> m a
forall (m :: * -> *) a.
MonadStConcreteNames m =>
StateT ConcreteNames m a -> m a
runStConcreteNames (StateT ConcreteNames m a -> m a)
-> (r -> StateT ConcreteNames m a) -> r -> m a
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (ConcreteNames -> m (a, ConcreteNames)) -> StateT ConcreteNames m a
forall s (m :: * -> *) a. (s -> m (a, s)) -> StateT s m a
StateT ((ConcreteNames -> m (a, ConcreteNames))
-> StateT ConcreteNames m a)
-> (r -> ConcreteNames -> m (a, ConcreteNames))
-> r
-> StateT ConcreteNames m a
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (ConcreteNames -> r -> m (a, ConcreteNames))
-> r -> ConcreteNames -> m (a, ConcreteNames)
forall a b c. (a -> b -> c) -> b -> a -> c
flip (ReaderT r m (a, ConcreteNames) -> r -> m (a, ConcreteNames)
forall r (m :: * -> *) a. ReaderT r m a -> r -> m a
runReaderT (ReaderT r m (a, ConcreteNames) -> r -> m (a, ConcreteNames))
-> (ConcreteNames -> ReaderT r m (a, ConcreteNames))
-> ConcreteNames
-> r
-> m (a, ConcreteNames)
forall b c a. (b -> c) -> (a -> b) -> a -> c
. StateT ConcreteNames (ReaderT r m) a
-> ConcreteNames -> ReaderT r m (a, ConcreteNames)
forall s (m :: * -> *) a. StateT s m a -> s -> m (a, s)
runStateT StateT ConcreteNames (ReaderT r m) a
m)
instance MonadStConcreteNames m => MonadStConcreteNames (StateT s m) where
runStConcreteNames :: forall a. StateT ConcreteNames (StateT s m) a -> StateT s m a
runStConcreteNames StateT ConcreteNames (StateT s m) a
m = (s -> m (a, s)) -> StateT s m a
forall s (m :: * -> *) a. (s -> m (a, s)) -> StateT s m a
StateT ((s -> m (a, s)) -> StateT s m a)
-> (s -> m (a, s)) -> StateT s m a
forall a b. (a -> b) -> a -> b
$ \s
s -> StateT ConcreteNames m (a, s) -> m (a, s)
forall a. StateT ConcreteNames m a -> m a
forall (m :: * -> *) a.
MonadStConcreteNames m =>
StateT ConcreteNames m a -> m a
runStConcreteNames (StateT ConcreteNames m (a, s) -> m (a, s))
-> StateT ConcreteNames m (a, s) -> m (a, s)
forall a b. (a -> b) -> a -> b
$ (ConcreteNames -> m ((a, s), ConcreteNames))
-> StateT ConcreteNames m (a, s)
forall s (m :: * -> *) a. (s -> m (a, s)) -> StateT s m a
StateT ((ConcreteNames -> m ((a, s), ConcreteNames))
-> StateT ConcreteNames m (a, s))
-> (ConcreteNames -> m ((a, s), ConcreteNames))
-> StateT ConcreteNames m (a, s)
forall a b. (a -> b) -> a -> b
$ \ConcreteNames
ns -> do
((x,ns'),s') <- StateT s m (a, ConcreteNames) -> s -> m ((a, ConcreteNames), s)
forall s (m :: * -> *) a. StateT s m a -> s -> m (a, s)
runStateT (StateT ConcreteNames (StateT s m) a
-> ConcreteNames -> StateT s m (a, ConcreteNames)
forall s (m :: * -> *) a. StateT s m a -> s -> m (a, s)
runStateT StateT ConcreteNames (StateT s m) a
m ConcreteNames
ns) s
s
return ((x,s'),ns')
instance (MonadStConcreteNames m, Monoid w) => MonadStConcreteNames (WriterT w m) where
runStConcreteNames :: forall a. StateT ConcreteNames (WriterT w m) a -> WriterT w m a
runStConcreteNames StateT ConcreteNames (WriterT w m) a
m = m (a, w) -> WriterT w m a
forall w (m :: * -> *) a. m (a, w) -> WriterT w m a
WriterT (m (a, w) -> WriterT w m a) -> m (a, w) -> WriterT w m a
forall a b. (a -> b) -> a -> b
$ StateT ConcreteNames m (a, w) -> m (a, w)
forall a. StateT ConcreteNames m a -> m a
forall (m :: * -> *) a.
MonadStConcreteNames m =>
StateT ConcreteNames m a -> m a
runStConcreteNames (StateT ConcreteNames m (a, w) -> m (a, w))
-> StateT ConcreteNames m (a, w) -> m (a, w)
forall a b. (a -> b) -> a -> b
$ (ConcreteNames -> m ((a, w), ConcreteNames))
-> StateT ConcreteNames m (a, w)
forall s (m :: * -> *) a. (s -> m (a, s)) -> StateT s m a
StateT ((ConcreteNames -> m ((a, w), ConcreteNames))
-> StateT ConcreteNames m (a, w))
-> (ConcreteNames -> m ((a, w), ConcreteNames))
-> StateT ConcreteNames m (a, w)
forall a b. (a -> b) -> a -> b
$ \ ConcreteNames
ns -> do
((x,ns'),w) <- WriterT w m (a, ConcreteNames) -> m ((a, ConcreteNames), w)
forall w (m :: * -> *) a. WriterT w m a -> m (a, w)
runWriterT (WriterT w m (a, ConcreteNames) -> m ((a, ConcreteNames), w))
-> WriterT w m (a, ConcreteNames) -> m ((a, ConcreteNames), w)
forall a b. (a -> b) -> a -> b
$ StateT ConcreteNames (WriterT w m) a
-> ConcreteNames -> WriterT w m (a, ConcreteNames)
forall s (m :: * -> *) a. StateT s m a -> s -> m (a, s)
runStateT StateT ConcreteNames (WriterT w m) a
m ConcreteNames
ns
return ((x,w),ns')
data ModuleCheckMode
= ModuleScopeChecked
| ModuleTypeChecked
deriving (ModuleCheckMode -> ModuleCheckMode -> Bool
(ModuleCheckMode -> ModuleCheckMode -> Bool)
-> (ModuleCheckMode -> ModuleCheckMode -> Bool)
-> Eq ModuleCheckMode
forall a. (a -> a -> Bool) -> (a -> a -> Bool) -> Eq a
$c== :: ModuleCheckMode -> ModuleCheckMode -> Bool
== :: ModuleCheckMode -> ModuleCheckMode -> Bool
$c/= :: ModuleCheckMode -> ModuleCheckMode -> Bool
/= :: ModuleCheckMode -> ModuleCheckMode -> Bool
Eq, Eq ModuleCheckMode
Eq ModuleCheckMode =>
(ModuleCheckMode -> ModuleCheckMode -> Ordering)
-> (ModuleCheckMode -> ModuleCheckMode -> Bool)
-> (ModuleCheckMode -> ModuleCheckMode -> Bool)
-> (ModuleCheckMode -> ModuleCheckMode -> Bool)
-> (ModuleCheckMode -> ModuleCheckMode -> Bool)
-> (ModuleCheckMode -> ModuleCheckMode -> ModuleCheckMode)
-> (ModuleCheckMode -> ModuleCheckMode -> ModuleCheckMode)
-> Ord ModuleCheckMode
ModuleCheckMode -> ModuleCheckMode -> Bool
ModuleCheckMode -> ModuleCheckMode -> Ordering
ModuleCheckMode -> ModuleCheckMode -> ModuleCheckMode
forall a.
Eq a =>
(a -> a -> Ordering)
-> (a -> a -> Bool)
-> (a -> a -> Bool)
-> (a -> a -> Bool)
-> (a -> a -> Bool)
-> (a -> a -> a)
-> (a -> a -> a)
-> Ord a
$ccompare :: ModuleCheckMode -> ModuleCheckMode -> Ordering
compare :: ModuleCheckMode -> ModuleCheckMode -> Ordering
$c< :: ModuleCheckMode -> ModuleCheckMode -> Bool
< :: ModuleCheckMode -> ModuleCheckMode -> Bool
$c<= :: ModuleCheckMode -> ModuleCheckMode -> Bool
<= :: ModuleCheckMode -> ModuleCheckMode -> Bool
$c> :: ModuleCheckMode -> ModuleCheckMode -> Bool
> :: ModuleCheckMode -> ModuleCheckMode -> Bool
$c>= :: ModuleCheckMode -> ModuleCheckMode -> Bool
>= :: ModuleCheckMode -> ModuleCheckMode -> Bool
$cmax :: ModuleCheckMode -> ModuleCheckMode -> ModuleCheckMode
max :: ModuleCheckMode -> ModuleCheckMode -> ModuleCheckMode
$cmin :: ModuleCheckMode -> ModuleCheckMode -> ModuleCheckMode
min :: ModuleCheckMode -> ModuleCheckMode -> ModuleCheckMode
Ord, ModuleCheckMode
ModuleCheckMode -> ModuleCheckMode -> Bounded ModuleCheckMode
forall a. a -> a -> Bounded a
$cminBound :: ModuleCheckMode
minBound :: ModuleCheckMode
$cmaxBound :: ModuleCheckMode
maxBound :: ModuleCheckMode
Bounded, Int -> ModuleCheckMode
ModuleCheckMode -> Int
ModuleCheckMode -> [ModuleCheckMode]
ModuleCheckMode -> ModuleCheckMode
ModuleCheckMode -> ModuleCheckMode -> [ModuleCheckMode]
ModuleCheckMode
-> ModuleCheckMode -> ModuleCheckMode -> [ModuleCheckMode]
(ModuleCheckMode -> ModuleCheckMode)
-> (ModuleCheckMode -> ModuleCheckMode)
-> (Int -> ModuleCheckMode)
-> (ModuleCheckMode -> Int)
-> (ModuleCheckMode -> [ModuleCheckMode])
-> (ModuleCheckMode -> ModuleCheckMode -> [ModuleCheckMode])
-> (ModuleCheckMode -> ModuleCheckMode -> [ModuleCheckMode])
-> (ModuleCheckMode
-> ModuleCheckMode -> ModuleCheckMode -> [ModuleCheckMode])
-> Enum ModuleCheckMode
forall a.
(a -> a)
-> (a -> a)
-> (Int -> a)
-> (a -> Int)
-> (a -> [a])
-> (a -> a -> [a])
-> (a -> a -> [a])
-> (a -> a -> a -> [a])
-> Enum a
$csucc :: ModuleCheckMode -> ModuleCheckMode
succ :: ModuleCheckMode -> ModuleCheckMode
$cpred :: ModuleCheckMode -> ModuleCheckMode
pred :: ModuleCheckMode -> ModuleCheckMode
$ctoEnum :: Int -> ModuleCheckMode
toEnum :: Int -> ModuleCheckMode
$cfromEnum :: ModuleCheckMode -> Int
fromEnum :: ModuleCheckMode -> Int
$cenumFrom :: ModuleCheckMode -> [ModuleCheckMode]
enumFrom :: ModuleCheckMode -> [ModuleCheckMode]
$cenumFromThen :: ModuleCheckMode -> ModuleCheckMode -> [ModuleCheckMode]
enumFromThen :: ModuleCheckMode -> ModuleCheckMode -> [ModuleCheckMode]
$cenumFromTo :: ModuleCheckMode -> ModuleCheckMode -> [ModuleCheckMode]
enumFromTo :: ModuleCheckMode -> ModuleCheckMode -> [ModuleCheckMode]
$cenumFromThenTo :: ModuleCheckMode
-> ModuleCheckMode -> ModuleCheckMode -> [ModuleCheckMode]
enumFromThenTo :: ModuleCheckMode
-> ModuleCheckMode -> ModuleCheckMode -> [ModuleCheckMode]
Enum, Int -> ModuleCheckMode -> ShowS
[ModuleCheckMode] -> ShowS
ModuleCheckMode -> String
(Int -> ModuleCheckMode -> ShowS)
-> (ModuleCheckMode -> String)
-> ([ModuleCheckMode] -> ShowS)
-> Show ModuleCheckMode
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> ModuleCheckMode -> ShowS
showsPrec :: Int -> ModuleCheckMode -> ShowS
$cshow :: ModuleCheckMode -> String
show :: ModuleCheckMode -> String
$cshowList :: [ModuleCheckMode] -> ShowS
showList :: [ModuleCheckMode] -> ShowS
Show, (forall x. ModuleCheckMode -> Rep ModuleCheckMode x)
-> (forall x. Rep ModuleCheckMode x -> ModuleCheckMode)
-> Generic ModuleCheckMode
forall x. Rep ModuleCheckMode x -> ModuleCheckMode
forall x. ModuleCheckMode -> Rep ModuleCheckMode x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. ModuleCheckMode -> Rep ModuleCheckMode x
from :: forall x. ModuleCheckMode -> Rep ModuleCheckMode x
$cto :: forall x. Rep ModuleCheckMode x -> ModuleCheckMode
to :: forall x. Rep ModuleCheckMode x -> ModuleCheckMode
Generic)
data ModuleInfo = ModuleInfo
{ ModuleInfo -> Interface
miInterface :: Interface
, ModuleInfo -> Set TCWarning
miWarnings :: Set TCWarning
, ModuleInfo -> Bool
miPrimitive :: Bool
, ModuleInfo -> ModuleCheckMode
miMode :: ModuleCheckMode
}
deriving (forall x. ModuleInfo -> Rep ModuleInfo x)
-> (forall x. Rep ModuleInfo x -> ModuleInfo) -> Generic ModuleInfo
forall x. Rep ModuleInfo x -> ModuleInfo
forall x. ModuleInfo -> Rep ModuleInfo x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. ModuleInfo -> Rep ModuleInfo x
from :: forall x. ModuleInfo -> Rep ModuleInfo x
$cto :: forall x. Rep ModuleInfo x -> ModuleInfo
to :: forall x. Rep ModuleInfo x -> ModuleInfo
Generic
type VisitedModules = Map TopLevelModuleName ModuleInfo
type DecodedModules = Map TopLevelModuleName ModuleInfo
data ForeignCode = ForeignCode Range String
deriving (Int -> ForeignCode -> ShowS
[ForeignCode] -> ShowS
ForeignCode -> String
(Int -> ForeignCode -> ShowS)
-> (ForeignCode -> String)
-> ([ForeignCode] -> ShowS)
-> Show ForeignCode
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> ForeignCode -> ShowS
showsPrec :: Int -> ForeignCode -> ShowS
$cshow :: ForeignCode -> String
show :: ForeignCode -> String
$cshowList :: [ForeignCode] -> ShowS
showList :: [ForeignCode] -> ShowS
Show, (forall x. ForeignCode -> Rep ForeignCode x)
-> (forall x. Rep ForeignCode x -> ForeignCode)
-> Generic ForeignCode
forall x. Rep ForeignCode x -> ForeignCode
forall x. ForeignCode -> Rep ForeignCode x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. ForeignCode -> Rep ForeignCode x
from :: forall x. ForeignCode -> Rep ForeignCode x
$cto :: forall x. Rep ForeignCode x -> ForeignCode
to :: forall x. Rep ForeignCode x -> ForeignCode
Generic)
newtype ForeignCodeStack = ForeignCodeStack
{ ForeignCodeStack -> [ForeignCode]
getForeignCodeStack :: [ForeignCode]
} deriving (Int -> ForeignCodeStack -> ShowS
[ForeignCodeStack] -> ShowS
ForeignCodeStack -> String
(Int -> ForeignCodeStack -> ShowS)
-> (ForeignCodeStack -> String)
-> ([ForeignCodeStack] -> ShowS)
-> Show ForeignCodeStack
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> ForeignCodeStack -> ShowS
showsPrec :: Int -> ForeignCodeStack -> ShowS
$cshow :: ForeignCodeStack -> String
show :: ForeignCodeStack -> String
$cshowList :: [ForeignCodeStack] -> ShowS
showList :: [ForeignCodeStack] -> ShowS
Show, (forall x. ForeignCodeStack -> Rep ForeignCodeStack x)
-> (forall x. Rep ForeignCodeStack x -> ForeignCodeStack)
-> Generic ForeignCodeStack
forall x. Rep ForeignCodeStack x -> ForeignCodeStack
forall x. ForeignCodeStack -> Rep ForeignCodeStack x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. ForeignCodeStack -> Rep ForeignCodeStack x
from :: forall x. ForeignCodeStack -> Rep ForeignCodeStack x
$cto :: forall x. Rep ForeignCodeStack x -> ForeignCodeStack
to :: forall x. Rep ForeignCodeStack x -> ForeignCodeStack
Generic, ForeignCodeStack -> ()
(ForeignCodeStack -> ()) -> NFData ForeignCodeStack
forall a. (a -> ()) -> NFData a
$crnf :: ForeignCodeStack -> ()
rnf :: ForeignCodeStack -> ()
NFData)
data Interface = Interface
{ Interface -> Word64
iSourceHash :: !Hash
, Interface -> Text
iSource :: TL.Text
, Interface -> FileType
iFileType :: FileType
, Interface -> [(TopLevelModuleName, Word64)]
iImportedModules :: [(TopLevelModuleName, Hash)]
, Interface -> ModuleName
iModuleName :: ModuleName
, Interface -> TopLevelModuleName
iTopLevelModuleName :: TopLevelModuleName
, Interface -> Map ModuleName Scope
iScope :: Map ModuleName Scope
, Interface -> ScopeInfo
iInsideScope :: ScopeInfo
, Interface -> Signature
iSignature :: Signature
, Interface -> RemoteMetaStore
iMetaBindings :: RemoteMetaStore
, Interface -> DisplayForms
iDisplayForms :: DisplayForms
, Interface -> UserWarnings
iUserWarnings :: UserWarnings
, Interface -> Maybe BackendName
iImportWarning :: Maybe Text
, Interface -> BuiltinThings' (PrimitiveId, QName)
iBuiltin :: BuiltinThings' (PrimitiveId, QName)
, Interface -> BackendForeignCode
iForeignCode :: Map BackendName ForeignCodeStack
, Interface -> HighlightingInfo
iHighlighting :: HighlightingInfo
, Interface -> [OptionsPragma]
iDefaultPragmaOptions :: [OptionsPragma]
, Interface -> [OptionsPragma]
iFilePragmaOptions :: [OptionsPragma]
, Interface -> PragmaOptions
iOptionsUsed :: PragmaOptions
, Interface -> PatternSynDefns
iPatternSyns :: A.PatternSynDefns
, Interface -> Set TCWarning
iWarnings :: Set TCWarning
, Interface -> Set QName
iPartialDefs :: Set QName
, Interface -> Map OpaqueId OpaqueBlock
iOpaqueBlocks :: Map OpaqueId OpaqueBlock
, Interface -> Map QName OpaqueId
iOpaqueNames :: Map QName OpaqueId
}
deriving (Int -> Interface -> ShowS
[Interface] -> ShowS
Interface -> String
(Int -> Interface -> ShowS)
-> (Interface -> String)
-> ([Interface] -> ShowS)
-> Show Interface
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> Interface -> ShowS
showsPrec :: Int -> Interface -> ShowS
$cshow :: Interface -> String
show :: Interface -> String
$cshowList :: [Interface] -> ShowS
showList :: [Interface] -> ShowS
Show, (forall x. Interface -> Rep Interface x)
-> (forall x. Rep Interface x -> Interface) -> Generic Interface
forall x. Rep Interface x -> Interface
forall x. Interface -> Rep Interface x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. Interface -> Rep Interface x
from :: forall x. Interface -> Rep Interface x
$cto :: forall x. Rep Interface x -> Interface
to :: forall x. Rep Interface x -> Interface
Generic)
instance Pretty Interface where
pretty :: Interface -> Doc
pretty (Interface
Word64
sourceH Text
source FileType
fileT [(TopLevelModuleName, Word64)]
importedM ModuleName
moduleN TopLevelModuleName
topModN Map ModuleName Scope
scope ScopeInfo
insideS
Signature
signature RemoteMetaStore
metas DisplayForms
display UserWarnings
userwarn Maybe BackendName
importwarn BuiltinThings' (PrimitiveId, QName)
builtin
BackendForeignCode
foreignCode HighlightingInfo
highlighting [OptionsPragma]
libPragmaO [OptionsPragma]
filePragmaO PragmaOptions
oUsed
PatternSynDefns
patternS Set TCWarning
warnings Set QName
partialdefs Map OpaqueId OpaqueBlock
oblocks Map QName OpaqueId
onames) =
Doc -> Int -> Doc -> Doc
forall a. Doc a -> Int -> Doc a -> Doc a
hang Doc
"Interface" Int
2 (Doc -> Doc) -> Doc -> Doc
forall a b. (a -> b) -> a -> b
$ [Doc] -> Doc
forall (t :: * -> *). Foldable t => t Doc -> Doc
vcat
[ Doc
"source hash:" Doc -> Doc -> Doc
forall a. Doc a -> Doc a -> Doc a
<+> (String -> Doc
forall a. Pretty a => a -> Doc
pretty (String -> Doc) -> (Word64 -> String) -> Word64 -> Doc
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Word64 -> String
forall a. Show a => a -> String
show) Word64
sourceH
, Doc
"source:" Doc -> Doc -> Doc
forall a. Doc a -> Doc a -> Doc a
$$ Int -> Doc -> Doc
forall a. Int -> Doc a -> Doc a
nest Int
2 (String -> Doc
forall a. String -> Doc a
text (String -> Doc) -> String -> Doc
forall a b. (a -> b) -> a -> b
$ Text -> String
TL.unpack Text
source)
, Doc
"file type:" Doc -> Doc -> Doc
forall a. Doc a -> Doc a -> Doc a
<+> (String -> Doc
forall a. Pretty a => a -> Doc
pretty (String -> Doc) -> (FileType -> String) -> FileType -> Doc
forall b c a. (b -> c) -> (a -> b) -> a -> c
. FileType -> String
forall a. Show a => a -> String
show) FileType
fileT
, Doc
"imported modules:" Doc -> Doc -> Doc
forall a. Doc a -> Doc a -> Doc a
<+> (String -> Doc
forall a. Pretty a => a -> Doc
pretty (String -> Doc)
-> ([(TopLevelModuleName, Word64)] -> String)
-> [(TopLevelModuleName, Word64)]
-> Doc
forall b c a. (b -> c) -> (a -> b) -> a -> c
. [(TopLevelModuleName, Word64)] -> String
forall a. Show a => a -> String
show) [(TopLevelModuleName, Word64)]
importedM
, Doc
"module name:" Doc -> Doc -> Doc
forall a. Doc a -> Doc a -> Doc a
<+> ModuleName -> Doc
forall a. Pretty a => a -> Doc
pretty ModuleName
moduleN
, Doc
"top-level module name:" Doc -> Doc -> Doc
forall a. Doc a -> Doc a -> Doc a
<+> TopLevelModuleName -> Doc
forall a. Pretty a => a -> Doc
pretty TopLevelModuleName
topModN
, Doc
"scope:" Doc -> Doc -> Doc
forall a. Doc a -> Doc a -> Doc a
<+> (String -> Doc
forall a. Pretty a => a -> Doc
pretty (String -> Doc)
-> (Map ModuleName Scope -> String) -> Map ModuleName Scope -> Doc
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Map ModuleName Scope -> String
forall a. Show a => a -> String
show) Map ModuleName Scope
scope
, Doc
"inside scope:" Doc -> Doc -> Doc
forall a. Doc a -> Doc a -> Doc a
<+> (String -> Doc
forall a. Pretty a => a -> Doc
pretty (String -> Doc) -> (ScopeInfo -> String) -> ScopeInfo -> Doc
forall b c a. (b -> c) -> (a -> b) -> a -> c
. ScopeInfo -> String
forall a. Show a => a -> String
show) ScopeInfo
insideS
, Doc
"signature:" Doc -> Doc -> Doc
forall a. Doc a -> Doc a -> Doc a
<+> (String -> Doc
forall a. Pretty a => a -> Doc
pretty (String -> Doc) -> (Signature -> String) -> Signature -> Doc
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Signature -> String
forall a. Show a => a -> String
show) Signature
signature
, Doc
"meta-variables:" Doc -> Doc -> Doc
forall a. Doc a -> Doc a -> Doc a
<+> (String -> Doc
forall a. Pretty a => a -> Doc
pretty (String -> Doc)
-> (RemoteMetaStore -> String) -> RemoteMetaStore -> Doc
forall b c a. (b -> c) -> (a -> b) -> a -> c
. RemoteMetaStore -> String
forall a. Show a => a -> String
show) RemoteMetaStore
metas
, Doc
"display:" Doc -> Doc -> Doc
forall a. Doc a -> Doc a -> Doc a
<+> (String -> Doc
forall a. Pretty a => a -> Doc
pretty (String -> Doc) -> (DisplayForms -> String) -> DisplayForms -> Doc
forall b c a. (b -> c) -> (a -> b) -> a -> c
. DisplayForms -> String
forall a. Show a => a -> String
show) DisplayForms
display
, Doc
"user warnings:" Doc -> Doc -> Doc
forall a. Doc a -> Doc a -> Doc a
<+> (String -> Doc
forall a. Pretty a => a -> Doc
pretty (String -> Doc) -> (UserWarnings -> String) -> UserWarnings -> Doc
forall b c a. (b -> c) -> (a -> b) -> a -> c
. UserWarnings -> String
forall a. Show a => a -> String
show) UserWarnings
userwarn
, Doc
"import warning:" Doc -> Doc -> Doc
forall a. Doc a -> Doc a -> Doc a
<+> (String -> Doc
forall a. Pretty a => a -> Doc
pretty (String -> Doc)
-> (Maybe BackendName -> String) -> Maybe BackendName -> Doc
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Maybe BackendName -> String
forall a. Show a => a -> String
show) Maybe BackendName
importwarn
, Doc
"builtin:" Doc -> Doc -> Doc
forall a. Doc a -> Doc a -> Doc a
<+> (String -> Doc
forall a. Pretty a => a -> Doc
pretty (String -> Doc)
-> (BuiltinThings' (PrimitiveId, QName) -> String)
-> BuiltinThings' (PrimitiveId, QName)
-> Doc
forall b c a. (b -> c) -> (a -> b) -> a -> c
. BuiltinThings' (PrimitiveId, QName) -> String
forall a. Show a => a -> String
show) BuiltinThings' (PrimitiveId, QName)
builtin
, Doc
"Foreign code:" Doc -> Doc -> Doc
forall a. Doc a -> Doc a -> Doc a
<+> (String -> Doc
forall a. Pretty a => a -> Doc
pretty (String -> Doc)
-> (BackendForeignCode -> String) -> BackendForeignCode -> Doc
forall b c a. (b -> c) -> (a -> b) -> a -> c
. BackendForeignCode -> String
forall a. Show a => a -> String
show) BackendForeignCode
foreignCode
, Doc
"highlighting:" Doc -> Doc -> Doc
forall a. Doc a -> Doc a -> Doc a
<+> (String -> Doc
forall a. Pretty a => a -> Doc
pretty (String -> Doc)
-> (HighlightingInfo -> String) -> HighlightingInfo -> Doc
forall b c a. (b -> c) -> (a -> b) -> a -> c
. HighlightingInfo -> String
forall a. Show a => a -> String
show) HighlightingInfo
highlighting
, Doc
"library pragma options:" Doc -> Doc -> Doc
forall a. Doc a -> Doc a -> Doc a
<+> (String -> Doc
forall a. Pretty a => a -> Doc
pretty (String -> Doc)
-> ([OptionsPragma] -> String) -> [OptionsPragma] -> Doc
forall b c a. (b -> c) -> (a -> b) -> a -> c
. [OptionsPragma] -> String
forall a. Show a => a -> String
show) [OptionsPragma]
libPragmaO
, Doc
"file pragma options:" Doc -> Doc -> Doc
forall a. Doc a -> Doc a -> Doc a
<+> (String -> Doc
forall a. Pretty a => a -> Doc
pretty (String -> Doc)
-> ([OptionsPragma] -> String) -> [OptionsPragma] -> Doc
forall b c a. (b -> c) -> (a -> b) -> a -> c
. [OptionsPragma] -> String
forall a. Show a => a -> String
show) [OptionsPragma]
filePragmaO
, Doc
"options used:" Doc -> Doc -> Doc
forall a. Doc a -> Doc a -> Doc a
<+> (String -> Doc
forall a. Pretty a => a -> Doc
pretty (String -> Doc)
-> (PragmaOptions -> String) -> PragmaOptions -> Doc
forall b c a. (b -> c) -> (a -> b) -> a -> c
. PragmaOptions -> String
forall a. Show a => a -> String
show) PragmaOptions
oUsed
, Doc
"pattern syns:" Doc -> Doc -> Doc
forall a. Doc a -> Doc a -> Doc a
<+> (String -> Doc
forall a. Pretty a => a -> Doc
pretty (String -> Doc)
-> (PatternSynDefns -> String) -> PatternSynDefns -> Doc
forall b c a. (b -> c) -> (a -> b) -> a -> c
. PatternSynDefns -> String
forall a. Show a => a -> String
show) PatternSynDefns
patternS
, Doc
"warnings:" Doc -> Doc -> Doc
forall a. Doc a -> Doc a -> Doc a
<+> (String -> Doc
forall a. Pretty a => a -> Doc
pretty (String -> Doc)
-> (Set TCWarning -> String) -> Set TCWarning -> Doc
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Set TCWarning -> String
forall a. Show a => a -> String
show) Set TCWarning
warnings
, Doc
"partial definitions:" Doc -> Doc -> Doc
forall a. Doc a -> Doc a -> Doc a
<+> (String -> Doc
forall a. Pretty a => a -> Doc
pretty (String -> Doc) -> (Set QName -> String) -> Set QName -> Doc
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Set QName -> String
forall a. Show a => a -> String
show) Set QName
partialdefs
, Doc
"opaque blocks:" Doc -> Doc -> Doc
forall a. Doc a -> Doc a -> Doc a
<+> Map OpaqueId OpaqueBlock -> Doc
forall a. Pretty a => a -> Doc
pretty Map OpaqueId OpaqueBlock
oblocks
, Doc
"opaque names" Doc -> Doc -> Doc
forall a. Doc a -> Doc a -> Doc a
<+> Map QName OpaqueId -> Doc
forall a. Pretty a => a -> Doc
pretty Map QName OpaqueId
onames
]
iFullHash :: Interface -> Hash
iFullHash :: Interface -> Word64
iFullHash Interface
i = [Word64] -> Word64
combineHashes ([Word64] -> Word64) -> [Word64] -> Word64
forall a b. (a -> b) -> a -> b
$ Interface -> Word64
iSourceHash Interface
i Word64 -> [Word64] -> [Word64]
forall a. a -> [a] -> [a]
: ((TopLevelModuleName, Word64) -> Word64)
-> [(TopLevelModuleName, Word64)] -> [Word64]
forall a b. (a -> b) -> [a] -> [b]
List.map (TopLevelModuleName, Word64) -> Word64
forall a b. (a, b) -> b
snd (Interface -> [(TopLevelModuleName, Word64)]
iImportedModules Interface
i)
intSignature :: Lens' Interface Signature
intSignature :: Lens' Interface Signature
intSignature Signature -> f Signature
f Interface
i = Signature -> f Signature
f (Interface -> Signature
iSignature Interface
i) f Signature -> (Signature -> Interface) -> f Interface
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \Signature
s -> Interface
i { iSignature = s }
data Closure a = Closure
{ forall a. Closure a -> Signature
clSignature :: Signature
, forall a. Closure a -> TCEnv
clEnv :: TCEnv
, forall a. Closure a -> ScopeInfo
clScope :: ScopeInfo
, forall a. Closure a -> Map ModuleName CheckpointId
clModuleCheckpoints :: Map ModuleName CheckpointId
, forall a. Closure a -> a
clValue :: a
}
deriving ((forall a b. (a -> b) -> Closure a -> Closure b)
-> (forall a b. a -> Closure b -> Closure a) -> Functor Closure
forall a b. a -> Closure b -> Closure a
forall a b. (a -> b) -> Closure a -> Closure b
forall (f :: * -> *).
(forall a b. (a -> b) -> f a -> f b)
-> (forall a b. a -> f b -> f a) -> Functor f
$cfmap :: forall a b. (a -> b) -> Closure a -> Closure b
fmap :: forall a b. (a -> b) -> Closure a -> Closure b
$c<$ :: forall a b. a -> Closure b -> Closure a
<$ :: forall a b. a -> Closure b -> Closure a
Functor, (forall m. Monoid m => Closure m -> m)
-> (forall m a. Monoid m => (a -> m) -> Closure a -> m)
-> (forall m a. Monoid m => (a -> m) -> Closure a -> m)
-> (forall a b. (a -> b -> b) -> b -> Closure a -> b)
-> (forall a b. (a -> b -> b) -> b -> Closure a -> b)
-> (forall b a. (b -> a -> b) -> b -> Closure a -> b)
-> (forall b a. (b -> a -> b) -> b -> Closure a -> b)
-> (forall a. (a -> a -> a) -> Closure a -> a)
-> (forall a. (a -> a -> a) -> Closure a -> a)
-> (forall a. Closure a -> [a])
-> (forall a. Closure a -> Bool)
-> (forall a. Closure a -> Int)
-> (forall a. Eq a => a -> Closure a -> Bool)
-> (forall a. Ord a => Closure a -> a)
-> (forall a. Ord a => Closure a -> a)
-> (forall a. Num a => Closure a -> a)
-> (forall a. Num a => Closure a -> a)
-> Foldable Closure
forall a. Eq a => a -> Closure a -> Bool
forall a. Num a => Closure a -> a
forall a. Ord a => Closure a -> a
forall m. Monoid m => Closure m -> m
forall a. Closure a -> Bool
forall a. Closure a -> Int
forall a. Closure a -> [a]
forall a. (a -> a -> a) -> Closure a -> a
forall m a. Monoid m => (a -> m) -> Closure a -> m
forall b a. (b -> a -> b) -> b -> Closure a -> b
forall a b. (a -> b -> b) -> b -> Closure a -> b
forall (t :: * -> *).
(forall m. Monoid m => t m -> m)
-> (forall m a. Monoid m => (a -> m) -> t a -> m)
-> (forall m a. Monoid m => (a -> m) -> t a -> m)
-> (forall a b. (a -> b -> b) -> b -> t a -> b)
-> (forall a b. (a -> b -> b) -> b -> t a -> b)
-> (forall b a. (b -> a -> b) -> b -> t a -> b)
-> (forall b a. (b -> a -> b) -> b -> t a -> b)
-> (forall a. (a -> a -> a) -> t a -> a)
-> (forall a. (a -> a -> a) -> t a -> a)
-> (forall a. t a -> [a])
-> (forall a. t a -> Bool)
-> (forall a. t a -> Int)
-> (forall a. Eq a => a -> t a -> Bool)
-> (forall a. Ord a => t a -> a)
-> (forall a. Ord a => t a -> a)
-> (forall a. Num a => t a -> a)
-> (forall a. Num a => t a -> a)
-> Foldable t
$cfold :: forall m. Monoid m => Closure m -> m
fold :: forall m. Monoid m => Closure m -> m
$cfoldMap :: forall m a. Monoid m => (a -> m) -> Closure a -> m
foldMap :: forall m a. Monoid m => (a -> m) -> Closure a -> m
$cfoldMap' :: forall m a. Monoid m => (a -> m) -> Closure a -> m
foldMap' :: forall m a. Monoid m => (a -> m) -> Closure a -> m
$cfoldr :: forall a b. (a -> b -> b) -> b -> Closure a -> b
foldr :: forall a b. (a -> b -> b) -> b -> Closure a -> b
$cfoldr' :: forall a b. (a -> b -> b) -> b -> Closure a -> b
foldr' :: forall a b. (a -> b -> b) -> b -> Closure a -> b
$cfoldl :: forall b a. (b -> a -> b) -> b -> Closure a -> b
foldl :: forall b a. (b -> a -> b) -> b -> Closure a -> b
$cfoldl' :: forall b a. (b -> a -> b) -> b -> Closure a -> b
foldl' :: forall b a. (b -> a -> b) -> b -> Closure a -> b
$cfoldr1 :: forall a. (a -> a -> a) -> Closure a -> a
foldr1 :: forall a. (a -> a -> a) -> Closure a -> a
$cfoldl1 :: forall a. (a -> a -> a) -> Closure a -> a
foldl1 :: forall a. (a -> a -> a) -> Closure a -> a
$ctoList :: forall a. Closure a -> [a]
toList :: forall a. Closure a -> [a]
$cnull :: forall a. Closure a -> Bool
null :: forall a. Closure a -> Bool
$clength :: forall a. Closure a -> Int
length :: forall a. Closure a -> Int
$celem :: forall a. Eq a => a -> Closure a -> Bool
elem :: forall a. Eq a => a -> Closure a -> Bool
$cmaximum :: forall a. Ord a => Closure a -> a
maximum :: forall a. Ord a => Closure a -> a
$cminimum :: forall a. Ord a => Closure a -> a
minimum :: forall a. Ord a => Closure a -> a
$csum :: forall a. Num a => Closure a -> a
sum :: forall a. Num a => Closure a -> a
$cproduct :: forall a. Num a => Closure a -> a
product :: forall a. Num a => Closure a -> a
Foldable, (forall x. Closure a -> Rep (Closure a) x)
-> (forall x. Rep (Closure a) x -> Closure a)
-> Generic (Closure a)
forall x. Rep (Closure a) x -> Closure a
forall x. Closure a -> Rep (Closure a) x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
forall a x. Rep (Closure a) x -> Closure a
forall a x. Closure a -> Rep (Closure a) x
$cfrom :: forall a x. Closure a -> Rep (Closure a) x
from :: forall x. Closure a -> Rep (Closure a) x
$cto :: forall a x. Rep (Closure a) x -> Closure a
to :: forall x. Rep (Closure a) x -> Closure a
Generic)
instance Show a => Show (Closure a) where
show :: Closure a -> String
show Closure a
cl = String
"Closure { clValue = " String -> ShowS
forall a. [a] -> [a] -> [a]
++ a -> String
forall a. Show a => a -> String
show (Closure a -> a
forall a. Closure a -> a
clValue Closure a
cl) String -> ShowS
forall a. [a] -> [a] -> [a]
++ String
" }"
instance HasRange a => HasRange (Closure a) where
getRange :: Closure a -> Range
getRange = a -> Range
forall a. HasRange a => a -> Range
getRange (a -> Range) -> (Closure a -> a) -> Closure a -> Range
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Closure a -> a
forall a. Closure a -> a
clValue
class LensClosure b a | b -> a where
lensClosure :: Lens' b (Closure a)
instance LensClosure (Closure a) a where
lensClosure :: Lens' (Closure a) (Closure a)
lensClosure = (Closure a -> f (Closure a)) -> Closure a -> f (Closure a)
forall a. a -> a
id
instance LensTCEnv (Closure a) where
lensTCEnv :: Lens' (Closure a) TCEnv
lensTCEnv TCEnv -> f TCEnv
f Closure a
cl = (TCEnv -> f TCEnv
f (TCEnv -> f TCEnv) -> TCEnv -> f TCEnv
forall a b. (a -> b) -> a -> b
$! Closure a -> TCEnv
forall a. Closure a -> TCEnv
clEnv Closure a
cl) f TCEnv -> (TCEnv -> Closure a) -> f (Closure a)
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ TCEnv
env -> Closure a
cl { clEnv = env }
{-# SPECIALIZE buildClosure :: a -> TCM (Closure a) #-}
buildClosure :: (MonadTCEnv m, ReadTCState m) => a -> m (Closure a)
buildClosure :: forall (m :: * -> *) a.
(MonadTCEnv m, ReadTCState m) =>
a -> m (Closure a)
buildClosure a
x = do
env <- m TCEnv
forall (m :: * -> *). MonadTCEnv m => m TCEnv
askTC
sig <- useR stSignature
scope <- useR stScope
cps <- useR stModuleCheckpoints
return $ Closure sig env scope cps x
type Constraints = [ProblemConstraint]
data ProblemConstraint = PConstr
{ ProblemConstraint -> Set ProblemId
constraintProblems :: Set ProblemId
, ProblemConstraint -> Blocker
constraintUnblocker :: Blocker
, ProblemConstraint -> Closure Constraint
theConstraint :: Closure Constraint
}
deriving (Int -> ProblemConstraint -> ShowS
Constraints -> ShowS
ProblemConstraint -> String
(Int -> ProblemConstraint -> ShowS)
-> (ProblemConstraint -> String)
-> (Constraints -> ShowS)
-> Show ProblemConstraint
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> ProblemConstraint -> ShowS
showsPrec :: Int -> ProblemConstraint -> ShowS
$cshow :: ProblemConstraint -> String
show :: ProblemConstraint -> String
$cshowList :: Constraints -> ShowS
showList :: Constraints -> ShowS
Show, (forall x. ProblemConstraint -> Rep ProblemConstraint x)
-> (forall x. Rep ProblemConstraint x -> ProblemConstraint)
-> Generic ProblemConstraint
forall x. Rep ProblemConstraint x -> ProblemConstraint
forall x. ProblemConstraint -> Rep ProblemConstraint x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. ProblemConstraint -> Rep ProblemConstraint x
from :: forall x. ProblemConstraint -> Rep ProblemConstraint x
$cto :: forall x. Rep ProblemConstraint x -> ProblemConstraint
to :: forall x. Rep ProblemConstraint x -> ProblemConstraint
Generic)
instance HasRange ProblemConstraint where
getRange :: ProblemConstraint -> Range
getRange = Closure Constraint -> Range
forall a. HasRange a => a -> Range
getRange (Closure Constraint -> Range)
-> (ProblemConstraint -> Closure Constraint)
-> ProblemConstraint
-> Range
forall b c a. (b -> c) -> (a -> b) -> a -> c
. ProblemConstraint -> Closure Constraint
theConstraint
data WhyCheckModality
= ConstructorType
| IndexedClause
| IndexedClauseArg Name Name
| GeneratedClause
deriving (Int -> WhyCheckModality -> ShowS
[WhyCheckModality] -> ShowS
WhyCheckModality -> String
(Int -> WhyCheckModality -> ShowS)
-> (WhyCheckModality -> String)
-> ([WhyCheckModality] -> ShowS)
-> Show WhyCheckModality
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> WhyCheckModality -> ShowS
showsPrec :: Int -> WhyCheckModality -> ShowS
$cshow :: WhyCheckModality -> String
show :: WhyCheckModality -> String
$cshowList :: [WhyCheckModality] -> ShowS
showList :: [WhyCheckModality] -> ShowS
Show, (forall x. WhyCheckModality -> Rep WhyCheckModality x)
-> (forall x. Rep WhyCheckModality x -> WhyCheckModality)
-> Generic WhyCheckModality
forall x. Rep WhyCheckModality x -> WhyCheckModality
forall x. WhyCheckModality -> Rep WhyCheckModality x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. WhyCheckModality -> Rep WhyCheckModality x
from :: forall x. WhyCheckModality -> Rep WhyCheckModality x
$cto :: forall x. Rep WhyCheckModality x -> WhyCheckModality
to :: forall x. Rep WhyCheckModality x -> WhyCheckModality
Generic)
data Constraint
= ValueCmp Comparison CompareAs Term Term
| ValueCmpOnFace Comparison Term Type Term Term
| ElimCmp [Polarity] [IsForced] Type Term [Elim] [Elim]
| SortCmp Comparison Sort Sort
| LevelCmp Comparison Level Level
| HasBiggerSort Sort
| HasPTSRule (Dom Type) (Abs Sort)
| CheckDataSort QName Sort
| CheckMetaInst MetaId
| CheckType Type
| UnBlock MetaId
| IsEmpty Range Type
| CheckSizeLtSat Term
| FindInstance MetaId (Maybe [Candidate])
| ResolveInstanceHead QName
| CheckFunDef A.DefInfo QName [A.Clause] TCErr
| UnquoteTactic Term Term Type
| CheckLockedVars Term Type (Arg Term) Type
| UsableAtModality WhyCheckModality (Maybe Sort) Modality Term
deriving (Int -> Constraint -> ShowS
[Constraint] -> ShowS
Constraint -> String
(Int -> Constraint -> ShowS)
-> (Constraint -> String)
-> ([Constraint] -> ShowS)
-> Show Constraint
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> Constraint -> ShowS
showsPrec :: Int -> Constraint -> ShowS
$cshow :: Constraint -> String
show :: Constraint -> String
$cshowList :: [Constraint] -> ShowS
showList :: [Constraint] -> ShowS
Show, (forall x. Constraint -> Rep Constraint x)
-> (forall x. Rep Constraint x -> Constraint) -> Generic Constraint
forall x. Rep Constraint x -> Constraint
forall x. Constraint -> Rep Constraint x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. Constraint -> Rep Constraint x
from :: forall x. Constraint -> Rep Constraint x
$cto :: forall x. Rep Constraint x -> Constraint
to :: forall x. Rep Constraint x -> Constraint
Generic)
instance HasRange Constraint where
getRange :: Constraint -> Range
getRange (IsEmpty Range
r Type
t) = Range
r
getRange Constraint
_ = Range
forall a. Range' a
noRange
instance Free Constraint where
freeVars' :: forall a c. IsVarSet a c => Constraint -> FreeM a c
freeVars' Constraint
c =
case Constraint
c of
ValueCmp Comparison
_ CompareAs
t Term
u Term
v -> (CompareAs, (Term, Term)) -> FreeM a c
forall t a c. (Free t, IsVarSet a c) => t -> FreeM a c
forall a c. IsVarSet a c => (CompareAs, (Term, Term)) -> FreeM a c
freeVars' (CompareAs
t, (Term
u, Term
v))
ValueCmpOnFace Comparison
_ Term
p Type
t Term
u Term
v -> (Term, (Type, (Term, Term))) -> FreeM a c
forall t a c. (Free t, IsVarSet a c) => t -> FreeM a c
forall a c.
IsVarSet a c =>
(Term, (Type, (Term, Term))) -> FreeM a c
freeVars' (Term
p, (Type
t, (Term
u, Term
v)))
ElimCmp [Polarity]
_ [IsForced]
_ Type
t Term
u [Elim]
es [Elim]
es' -> ((Type, Term), ([Elim], [Elim])) -> FreeM a c
forall t a c. (Free t, IsVarSet a c) => t -> FreeM a c
forall a c.
IsVarSet a c =>
((Type, Term), ([Elim], [Elim])) -> FreeM a c
freeVars' ((Type
t, Term
u), ([Elim]
es, [Elim]
es'))
SortCmp Comparison
_ Sort
s Sort
s' -> (Sort, Sort) -> FreeM a c
forall t a c. (Free t, IsVarSet a c) => t -> FreeM a c
forall a c. IsVarSet a c => (Sort, Sort) -> FreeM a c
freeVars' (Sort
s, Sort
s')
LevelCmp Comparison
_ Level
l Level
l' -> (Level, Level) -> FreeM a c
forall t a c. (Free t, IsVarSet a c) => t -> FreeM a c
forall a c. IsVarSet a c => (Level, Level) -> FreeM a c
freeVars' (Level
l, Level
l')
UnBlock MetaId
_ -> FreeM a c
forall a. Monoid a => a
mempty
IsEmpty Range
_ Type
t -> Type -> FreeM a c
forall t a c. (Free t, IsVarSet a c) => t -> FreeM a c
forall a c. IsVarSet a c => Type -> FreeM a c
freeVars' Type
t
CheckSizeLtSat Term
u -> Term -> FreeM a c
forall t a c. (Free t, IsVarSet a c) => t -> FreeM a c
forall a c. IsVarSet a c => Term -> FreeM a c
freeVars' Term
u
FindInstance MetaId
_ Maybe [Candidate]
cs -> Maybe [Candidate] -> FreeM a c
forall t a c. (Free t, IsVarSet a c) => t -> FreeM a c
forall a c. IsVarSet a c => Maybe [Candidate] -> FreeM a c
freeVars' Maybe [Candidate]
cs
ResolveInstanceHead QName
q -> FreeM a c
forall a. Monoid a => a
mempty
CheckFunDef{} -> FreeM a c
forall a. Monoid a => a
mempty
HasBiggerSort Sort
s -> Sort -> FreeM a c
forall t a c. (Free t, IsVarSet a c) => t -> FreeM a c
forall a c. IsVarSet a c => Sort -> FreeM a c
freeVars' Sort
s
HasPTSRule Dom Type
a Abs Sort
s -> (Dom Type, Abs Sort) -> FreeM a c
forall t a c. (Free t, IsVarSet a c) => t -> FreeM a c
forall a c. IsVarSet a c => (Dom Type, Abs Sort) -> FreeM a c
freeVars' (Dom Type
a , Abs Sort
s)
CheckLockedVars Term
a Type
b Arg Term
c Type
d -> ((Term, Type), (Arg Term, Type)) -> FreeM a c
forall t a c. (Free t, IsVarSet a c) => t -> FreeM a c
forall a c.
IsVarSet a c =>
((Term, Type), (Arg Term, Type)) -> FreeM a c
freeVars' ((Term
a,Type
b),(Arg Term
c,Type
d))
UnquoteTactic Term
t Term
h Type
g -> (Term, (Term, Type)) -> FreeM a c
forall t a c. (Free t, IsVarSet a c) => t -> FreeM a c
forall a c. IsVarSet a c => (Term, (Term, Type)) -> FreeM a c
freeVars' (Term
t, (Term
h, Type
g))
CheckDataSort QName
_ Sort
s -> Sort -> FreeM a c
forall t a c. (Free t, IsVarSet a c) => t -> FreeM a c
forall a c. IsVarSet a c => Sort -> FreeM a c
freeVars' Sort
s
CheckMetaInst MetaId
m -> FreeM a c
forall a. Monoid a => a
mempty
CheckType Type
t -> Type -> FreeM a c
forall t a c. (Free t, IsVarSet a c) => t -> FreeM a c
forall a c. IsVarSet a c => Type -> FreeM a c
freeVars' Type
t
UsableAtModality WhyCheckModality
_ Maybe Sort
ms Modality
mod Term
t -> (Maybe Sort, Term) -> FreeM a c
forall t a c. (Free t, IsVarSet a c) => t -> FreeM a c
forall a c. IsVarSet a c => (Maybe Sort, Term) -> FreeM a c
freeVars' (Maybe Sort
ms, Term
t)
instance TermLike Constraint where
foldTerm :: forall m. Monoid m => (Term -> m) -> Constraint -> m
foldTerm Term -> m
f = \case
ValueCmp Comparison
_ CompareAs
t Term
u Term
v -> (Term -> m) -> (CompareAs, Term, Term) -> m
forall m. Monoid m => (Term -> m) -> (CompareAs, Term, Term) -> m
forall a m. (TermLike a, Monoid m) => (Term -> m) -> a -> m
foldTerm Term -> m
f (CompareAs
t, Term
u, Term
v)
ValueCmpOnFace Comparison
_ Term
p Type
t Term
u Term
v -> (Term -> m) -> (Term, Type, Term, Term) -> m
forall m. Monoid m => (Term -> m) -> (Term, Type, Term, Term) -> m
forall a m. (TermLike a, Monoid m) => (Term -> m) -> a -> m
foldTerm Term -> m
f (Term
p, Type
t, Term
u, Term
v)
ElimCmp [Polarity]
_ [IsForced]
_ Type
t Term
u [Elim]
es [Elim]
es' -> (Term -> m) -> (Type, Term, [Elim], [Elim]) -> m
forall m.
Monoid m =>
(Term -> m) -> (Type, Term, [Elim], [Elim]) -> m
forall a m. (TermLike a, Monoid m) => (Term -> m) -> a -> m
foldTerm Term -> m
f (Type
t, Term
u, [Elim]
es, [Elim]
es')
LevelCmp Comparison
_ Level
l Level
l' -> (Term -> m) -> (Term, Term) -> m
forall m. Monoid m => (Term -> m) -> (Term, Term) -> m
forall a m. (TermLike a, Monoid m) => (Term -> m) -> a -> m
foldTerm Term -> m
f (Level -> Term
Level Level
l, Level -> Term
Level Level
l')
IsEmpty Range
_ Type
t -> (Term -> m) -> Type -> m
forall m. Monoid m => (Term -> m) -> Type -> m
forall a m. (TermLike a, Monoid m) => (Term -> m) -> a -> m
foldTerm Term -> m
f Type
t
CheckSizeLtSat Term
u -> (Term -> m) -> Term -> m
forall m. Monoid m => (Term -> m) -> Term -> m
forall a m. (TermLike a, Monoid m) => (Term -> m) -> a -> m
foldTerm Term -> m
f Term
u
UnquoteTactic Term
t Term
h Type
g -> (Term -> m) -> (Term, Term, Type) -> m
forall m. Monoid m => (Term -> m) -> (Term, Term, Type) -> m
forall a m. (TermLike a, Monoid m) => (Term -> m) -> a -> m
foldTerm Term -> m
f (Term
t, Term
h, Type
g)
SortCmp Comparison
_ Sort
s1 Sort
s2 -> (Term -> m) -> (Term, Term) -> m
forall m. Monoid m => (Term -> m) -> (Term, Term) -> m
forall a m. (TermLike a, Monoid m) => (Term -> m) -> a -> m
foldTerm Term -> m
f (Sort -> Term
Sort Sort
s1, Sort -> Term
Sort Sort
s2)
UnBlock MetaId
_ -> m
forall a. Monoid a => a
mempty
CheckLockedVars Term
a Type
b Arg Term
c Type
d -> (Term -> m) -> (Term, Type, Arg Term, Type) -> m
forall m.
Monoid m =>
(Term -> m) -> (Term, Type, Arg Term, Type) -> m
forall a m. (TermLike a, Monoid m) => (Term -> m) -> a -> m
foldTerm Term -> m
f (Term
a, Type
b, Arg Term
c, Type
d)
FindInstance MetaId
_ Maybe [Candidate]
_ -> m
forall a. Monoid a => a
mempty
ResolveInstanceHead QName
q -> m
forall a. Monoid a => a
mempty
CheckFunDef{} -> m
forall a. Monoid a => a
mempty
HasBiggerSort Sort
s -> (Term -> m) -> Sort -> m
forall m. Monoid m => (Term -> m) -> Sort -> m
forall a m. (TermLike a, Monoid m) => (Term -> m) -> a -> m
foldTerm Term -> m
f Sort
s
HasPTSRule Dom Type
a Abs Sort
s -> (Term -> m) -> (Dom Type, Abs Term) -> m
forall m. Monoid m => (Term -> m) -> (Dom Type, Abs Term) -> m
forall a m. (TermLike a, Monoid m) => (Term -> m) -> a -> m
foldTerm Term -> m
f (Dom Type
a, Sort -> Term
Sort (Sort -> Term) -> Abs Sort -> Abs Term
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> Abs Sort
s)
CheckDataSort QName
_ Sort
s -> (Term -> m) -> Sort -> m
forall m. Monoid m => (Term -> m) -> Sort -> m
forall a m. (TermLike a, Monoid m) => (Term -> m) -> a -> m
foldTerm Term -> m
f Sort
s
CheckMetaInst MetaId
m -> m
forall a. Monoid a => a
mempty
CheckType Type
t -> (Term -> m) -> Type -> m
forall m. Monoid m => (Term -> m) -> Type -> m
forall a m. (TermLike a, Monoid m) => (Term -> m) -> a -> m
foldTerm Term -> m
f Type
t
UsableAtModality WhyCheckModality
_ Maybe Sort
ms Modality
m Term
t -> (Term -> m) -> (Maybe Term, Term) -> m
forall m. Monoid m => (Term -> m) -> (Maybe Term, Term) -> m
forall a m. (TermLike a, Monoid m) => (Term -> m) -> a -> m
foldTerm Term -> m
f (Sort -> Term
Sort (Sort -> Term) -> Maybe Sort -> Maybe Term
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> Maybe Sort
ms, Term
t)
traverseTermM :: forall (m :: * -> *).
Monad m =>
(Term -> m Term) -> Constraint -> m Constraint
traverseTermM Term -> m Term
f Constraint
c = m Constraint
forall a. HasCallStack => a
__IMPOSSIBLE__
instance AllMetas Constraint
instance Pretty Comparison where
pretty :: Comparison -> Doc
pretty Comparison
CmpEq = Doc
"="
pretty Comparison
CmpLeq = Doc
"=<"
data CompareDirection = DirEq | DirLeq | DirGeq
deriving (CompareDirection -> CompareDirection -> Bool
(CompareDirection -> CompareDirection -> Bool)
-> (CompareDirection -> CompareDirection -> Bool)
-> Eq CompareDirection
forall a. (a -> a -> Bool) -> (a -> a -> Bool) -> Eq a
$c== :: CompareDirection -> CompareDirection -> Bool
== :: CompareDirection -> CompareDirection -> Bool
$c/= :: CompareDirection -> CompareDirection -> Bool
/= :: CompareDirection -> CompareDirection -> Bool
Eq, Int -> CompareDirection -> ShowS
[CompareDirection] -> ShowS
CompareDirection -> String
(Int -> CompareDirection -> ShowS)
-> (CompareDirection -> String)
-> ([CompareDirection] -> ShowS)
-> Show CompareDirection
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> CompareDirection -> ShowS
showsPrec :: Int -> CompareDirection -> ShowS
$cshow :: CompareDirection -> String
show :: CompareDirection -> String
$cshowList :: [CompareDirection] -> ShowS
showList :: [CompareDirection] -> ShowS
Show)
instance Pretty CompareDirection where
pretty :: CompareDirection -> Doc
pretty = String -> Doc
forall a. String -> Doc a
text (String -> Doc)
-> (CompareDirection -> String) -> CompareDirection -> Doc
forall b c a. (b -> c) -> (a -> b) -> a -> c
. \case
CompareDirection
DirEq -> String
"="
CompareDirection
DirLeq -> String
"=<"
CompareDirection
DirGeq -> String
">="
fromCmp :: Comparison -> CompareDirection
fromCmp :: Comparison -> CompareDirection
fromCmp Comparison
CmpEq = CompareDirection
DirEq
fromCmp Comparison
CmpLeq = CompareDirection
DirLeq
flipCmp :: CompareDirection -> CompareDirection
flipCmp :: CompareDirection -> CompareDirection
flipCmp CompareDirection
DirEq = CompareDirection
DirEq
flipCmp CompareDirection
DirLeq = CompareDirection
DirGeq
flipCmp CompareDirection
DirGeq = CompareDirection
DirLeq
dirToCmp :: (Comparison -> a -> a -> c) -> CompareDirection -> a -> a -> c
dirToCmp :: forall a c.
(Comparison -> a -> a -> c) -> CompareDirection -> a -> a -> c
dirToCmp Comparison -> a -> a -> c
cont CompareDirection
DirEq = Comparison -> a -> a -> c
cont Comparison
CmpEq
dirToCmp Comparison -> a -> a -> c
cont CompareDirection
DirLeq = Comparison -> a -> a -> c
cont Comparison
CmpLeq
dirToCmp Comparison -> a -> a -> c
cont CompareDirection
DirGeq = (a -> a -> c) -> a -> a -> c
forall a b c. (a -> b -> c) -> b -> a -> c
flip ((a -> a -> c) -> a -> a -> c) -> (a -> a -> c) -> a -> a -> c
forall a b. (a -> b) -> a -> b
$ Comparison -> a -> a -> c
cont Comparison
CmpLeq
data CompareAs
= AsTermsOf Type
| AsSizes
| AsTypes
deriving (Int -> CompareAs -> ShowS
[CompareAs] -> ShowS
CompareAs -> String
(Int -> CompareAs -> ShowS)
-> (CompareAs -> String)
-> ([CompareAs] -> ShowS)
-> Show CompareAs
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> CompareAs -> ShowS
showsPrec :: Int -> CompareAs -> ShowS
$cshow :: CompareAs -> String
show :: CompareAs -> String
$cshowList :: [CompareAs] -> ShowS
showList :: [CompareAs] -> ShowS
Show, (forall x. CompareAs -> Rep CompareAs x)
-> (forall x. Rep CompareAs x -> CompareAs) -> Generic CompareAs
forall x. Rep CompareAs x -> CompareAs
forall x. CompareAs -> Rep CompareAs x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. CompareAs -> Rep CompareAs x
from :: forall x. CompareAs -> Rep CompareAs x
$cto :: forall x. Rep CompareAs x -> CompareAs
to :: forall x. Rep CompareAs x -> CompareAs
Generic)
instance Free CompareAs where
freeVars' :: forall a c. IsVarSet a c => CompareAs -> FreeM a c
freeVars' (AsTermsOf Type
a) = Type -> FreeM a c
forall t a c. (Free t, IsVarSet a c) => t -> FreeM a c
forall a c. IsVarSet a c => Type -> FreeM a c
freeVars' Type
a
freeVars' CompareAs
AsSizes = FreeM a c
forall a. Monoid a => a
mempty
freeVars' CompareAs
AsTypes = FreeM a c
forall a. Monoid a => a
mempty
instance TermLike CompareAs where
foldTerm :: forall m. Monoid m => (Term -> m) -> CompareAs -> m
foldTerm Term -> m
f (AsTermsOf Type
a) = (Term -> m) -> Type -> m
forall m. Monoid m => (Term -> m) -> Type -> m
forall a m. (TermLike a, Monoid m) => (Term -> m) -> a -> m
foldTerm Term -> m
f Type
a
foldTerm Term -> m
f CompareAs
AsSizes = m
forall a. Monoid a => a
mempty
foldTerm Term -> m
f CompareAs
AsTypes = m
forall a. Monoid a => a
mempty
traverseTermM :: forall (m :: * -> *).
Monad m =>
(Term -> m Term) -> CompareAs -> m CompareAs
traverseTermM Term -> m Term
f = \case
AsTermsOf Type
a -> Type -> CompareAs
AsTermsOf (Type -> CompareAs) -> m Type -> m CompareAs
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> (Term -> m Term) -> Type -> m Type
forall a (m :: * -> *).
(TermLike a, Monad m) =>
(Term -> m Term) -> a -> m a
forall (m :: * -> *). Monad m => (Term -> m Term) -> Type -> m Type
traverseTermM Term -> m Term
f Type
a
CompareAs
AsSizes -> CompareAs -> m CompareAs
forall a. a -> m a
forall (m :: * -> *) a. Monad m => a -> m a
return CompareAs
AsSizes
CompareAs
AsTypes -> CompareAs -> m CompareAs
forall a. a -> m a
forall (m :: * -> *) a. Monad m => a -> m a
return CompareAs
AsTypes
instance AllMetas CompareAs
instance Pretty CompareAs where
pretty :: CompareAs -> Doc
pretty (AsTermsOf Type
a) = Doc
":" Doc -> Doc -> Doc
forall a. Doc a -> Doc a -> Doc a
<+> Type -> Doc
forall a. Pretty a => a -> Doc
pretty Type
a
pretty CompareAs
AsSizes = Doc
":" Doc -> Doc -> Doc
forall a. Doc a -> Doc a -> Doc a
<+> String -> Doc
forall a. String -> Doc a
text String
"Size"
pretty CompareAs
AsTypes = Doc
forall a. Null a => a
empty
data Open a = OpenThing { forall a. Open a -> CheckpointId
openThingCheckpoint :: CheckpointId
, forall a. Open a -> Map CheckpointId Substitution
openThingCheckpointMap :: Map CheckpointId Substitution
, forall a. Open a -> ModuleNameHash
openThingModule :: ModuleNameHash
, forall a. Open a -> a
openThing :: a }
deriving (Int -> Open a -> ShowS
[Open a] -> ShowS
Open a -> String
(Int -> Open a -> ShowS)
-> (Open a -> String) -> ([Open a] -> ShowS) -> Show (Open a)
forall a. Show a => Int -> Open a -> ShowS
forall a. Show a => [Open a] -> ShowS
forall a. Show a => Open a -> String
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: forall a. Show a => Int -> Open a -> ShowS
showsPrec :: Int -> Open a -> ShowS
$cshow :: forall a. Show a => Open a -> String
show :: Open a -> String
$cshowList :: forall a. Show a => [Open a] -> ShowS
showList :: [Open a] -> ShowS
Show, (forall a b. (a -> b) -> Open a -> Open b)
-> (forall a b. a -> Open b -> Open a) -> Functor Open
forall a b. a -> Open b -> Open a
forall a b. (a -> b) -> Open a -> Open b
forall (f :: * -> *).
(forall a b. (a -> b) -> f a -> f b)
-> (forall a b. a -> f b -> f a) -> Functor f
$cfmap :: forall a b. (a -> b) -> Open a -> Open b
fmap :: forall a b. (a -> b) -> Open a -> Open b
$c<$ :: forall a b. a -> Open b -> Open a
<$ :: forall a b. a -> Open b -> Open a
Functor, (forall m. Monoid m => Open m -> m)
-> (forall m a. Monoid m => (a -> m) -> Open a -> m)
-> (forall m a. Monoid m => (a -> m) -> Open a -> m)
-> (forall a b. (a -> b -> b) -> b -> Open a -> b)
-> (forall a b. (a -> b -> b) -> b -> Open a -> b)
-> (forall b a. (b -> a -> b) -> b -> Open a -> b)
-> (forall b a. (b -> a -> b) -> b -> Open a -> b)
-> (forall a. (a -> a -> a) -> Open a -> a)
-> (forall a. (a -> a -> a) -> Open a -> a)
-> (forall a. Open a -> [a])
-> (forall a. Open a -> Bool)
-> (forall a. Open a -> Int)
-> (forall a. Eq a => a -> Open a -> Bool)
-> (forall a. Ord a => Open a -> a)
-> (forall a. Ord a => Open a -> a)
-> (forall a. Num a => Open a -> a)
-> (forall a. Num a => Open a -> a)
-> Foldable Open
forall a. Eq a => a -> Open a -> Bool
forall a. Num a => Open a -> a
forall a. Ord a => Open a -> a
forall m. Monoid m => Open m -> m
forall a. Open a -> Bool
forall a. Open a -> Int
forall a. Open a -> [a]
forall a. (a -> a -> a) -> Open a -> a
forall m a. Monoid m => (a -> m) -> Open a -> m
forall b a. (b -> a -> b) -> b -> Open a -> b
forall a b. (a -> b -> b) -> b -> Open a -> b
forall (t :: * -> *).
(forall m. Monoid m => t m -> m)
-> (forall m a. Monoid m => (a -> m) -> t a -> m)
-> (forall m a. Monoid m => (a -> m) -> t a -> m)
-> (forall a b. (a -> b -> b) -> b -> t a -> b)
-> (forall a b. (a -> b -> b) -> b -> t a -> b)
-> (forall b a. (b -> a -> b) -> b -> t a -> b)
-> (forall b a. (b -> a -> b) -> b -> t a -> b)
-> (forall a. (a -> a -> a) -> t a -> a)
-> (forall a. (a -> a -> a) -> t a -> a)
-> (forall a. t a -> [a])
-> (forall a. t a -> Bool)
-> (forall a. t a -> Int)
-> (forall a. Eq a => a -> t a -> Bool)
-> (forall a. Ord a => t a -> a)
-> (forall a. Ord a => t a -> a)
-> (forall a. Num a => t a -> a)
-> (forall a. Num a => t a -> a)
-> Foldable t
$cfold :: forall m. Monoid m => Open m -> m
fold :: forall m. Monoid m => Open m -> m
$cfoldMap :: forall m a. Monoid m => (a -> m) -> Open a -> m
foldMap :: forall m a. Monoid m => (a -> m) -> Open a -> m
$cfoldMap' :: forall m a. Monoid m => (a -> m) -> Open a -> m
foldMap' :: forall m a. Monoid m => (a -> m) -> Open a -> m
$cfoldr :: forall a b. (a -> b -> b) -> b -> Open a -> b
foldr :: forall a b. (a -> b -> b) -> b -> Open a -> b
$cfoldr' :: forall a b. (a -> b -> b) -> b -> Open a -> b
foldr' :: forall a b. (a -> b -> b) -> b -> Open a -> b
$cfoldl :: forall b a. (b -> a -> b) -> b -> Open a -> b
foldl :: forall b a. (b -> a -> b) -> b -> Open a -> b
$cfoldl' :: forall b a. (b -> a -> b) -> b -> Open a -> b
foldl' :: forall b a. (b -> a -> b) -> b -> Open a -> b
$cfoldr1 :: forall a. (a -> a -> a) -> Open a -> a
foldr1 :: forall a. (a -> a -> a) -> Open a -> a
$cfoldl1 :: forall a. (a -> a -> a) -> Open a -> a
foldl1 :: forall a. (a -> a -> a) -> Open a -> a
$ctoList :: forall a. Open a -> [a]
toList :: forall a. Open a -> [a]
$cnull :: forall a. Open a -> Bool
null :: forall a. Open a -> Bool
$clength :: forall a. Open a -> Int
length :: forall a. Open a -> Int
$celem :: forall a. Eq a => a -> Open a -> Bool
elem :: forall a. Eq a => a -> Open a -> Bool
$cmaximum :: forall a. Ord a => Open a -> a
maximum :: forall a. Ord a => Open a -> a
$cminimum :: forall a. Ord a => Open a -> a
minimum :: forall a. Ord a => Open a -> a
$csum :: forall a. Num a => Open a -> a
sum :: forall a. Num a => Open a -> a
$cproduct :: forall a. Num a => Open a -> a
product :: forall a. Num a => Open a -> a
Foldable, Functor Open
Foldable Open
(Functor Open, Foldable Open) =>
(forall (f :: * -> *) a b.
Applicative f =>
(a -> f b) -> Open a -> f (Open b))
-> (forall (f :: * -> *) a.
Applicative f =>
Open (f a) -> f (Open a))
-> (forall (m :: * -> *) a b.
Monad m =>
(a -> m b) -> Open a -> m (Open b))
-> (forall (m :: * -> *) a. Monad m => Open (m a) -> m (Open a))
-> Traversable Open
forall (t :: * -> *).
(Functor t, Foldable t) =>
(forall (f :: * -> *) a b.
Applicative f =>
(a -> f b) -> t a -> f (t b))
-> (forall (f :: * -> *) a. Applicative f => t (f a) -> f (t a))
-> (forall (m :: * -> *) a b.
Monad m =>
(a -> m b) -> t a -> m (t b))
-> (forall (m :: * -> *) a. Monad m => t (m a) -> m (t a))
-> Traversable t
forall (m :: * -> *) a. Monad m => Open (m a) -> m (Open a)
forall (f :: * -> *) a. Applicative f => Open (f a) -> f (Open a)
forall (m :: * -> *) a b.
Monad m =>
(a -> m b) -> Open a -> m (Open b)
forall (f :: * -> *) a b.
Applicative f =>
(a -> f b) -> Open a -> f (Open b)
$ctraverse :: forall (f :: * -> *) a b.
Applicative f =>
(a -> f b) -> Open a -> f (Open b)
traverse :: forall (f :: * -> *) a b.
Applicative f =>
(a -> f b) -> Open a -> f (Open b)
$csequenceA :: forall (f :: * -> *) a. Applicative f => Open (f a) -> f (Open a)
sequenceA :: forall (f :: * -> *) a. Applicative f => Open (f a) -> f (Open a)
$cmapM :: forall (m :: * -> *) a b.
Monad m =>
(a -> m b) -> Open a -> m (Open b)
mapM :: forall (m :: * -> *) a b.
Monad m =>
(a -> m b) -> Open a -> m (Open b)
$csequence :: forall (m :: * -> *) a. Monad m => Open (m a) -> m (Open a)
sequence :: forall (m :: * -> *) a. Monad m => Open (m a) -> m (Open a)
Traversable, (forall x. Open a -> Rep (Open a) x)
-> (forall x. Rep (Open a) x -> Open a) -> Generic (Open a)
forall x. Rep (Open a) x -> Open a
forall x. Open a -> Rep (Open a) x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
forall a x. Rep (Open a) x -> Open a
forall a x. Open a -> Rep (Open a) x
$cfrom :: forall a x. Open a -> Rep (Open a) x
from :: forall x. Open a -> Rep (Open a) x
$cto :: forall a x. Rep (Open a) x -> Open a
to :: forall x. Rep (Open a) x -> Open a
Generic)
instance Decoration Open where
traverseF :: forall (m :: * -> *) a b.
Functor m =>
(a -> m b) -> Open a -> m (Open b)
traverseF a -> m b
f (OpenThing CheckpointId
cp Map CheckpointId Substitution
env ModuleNameHash
m a
x) = CheckpointId
-> Map CheckpointId Substitution -> ModuleNameHash -> b -> Open b
forall a.
CheckpointId
-> Map CheckpointId Substitution -> ModuleNameHash -> a -> Open a
OpenThing CheckpointId
cp Map CheckpointId Substitution
env ModuleNameHash
m (b -> Open b) -> m b -> m (Open b)
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> a -> m b
f a
x
instance Pretty a => Pretty (Open a) where
prettyPrec :: Int -> Open a -> Doc
prettyPrec Int
p (OpenThing CheckpointId
cp Map CheckpointId Substitution
env ModuleNameHash
_ a
x) = Bool -> Doc -> Doc
mparens (Int
p Int -> Int -> Bool
forall a. Ord a => a -> a -> Bool
> Int
9) (Doc -> Doc) -> Doc -> Doc
forall a b. (a -> b) -> a -> b
$
Doc
"OpenThing" Doc -> Doc -> Doc
forall a. Doc a -> Doc a -> Doc a
<+> CheckpointId -> Doc
forall a. Pretty a => a -> Doc
pretty CheckpointId
cp Doc -> Doc -> Doc
forall a. Doc a -> Doc a -> Doc a
<+> [(CheckpointId, Substitution)] -> Doc
forall a. Pretty a => a -> Doc
pretty (Map CheckpointId Substitution -> [(CheckpointId, Substitution)]
forall k a. Map k a -> [(k, a)]
Map.toList Map CheckpointId Substitution
env) Doc -> Doc -> Doc
<?> Int -> a -> Doc
forall a. Pretty a => Int -> a -> Doc
prettyPrec Int
10 a
x
data Judgement a
= HasType
{ forall a. Judgement a -> a
jMetaId :: a
, forall a. Judgement a -> Comparison
jComparison :: Comparison
, forall a. Judgement a -> Type
jMetaType :: Type
}
| IsSort
{ jMetaId :: a
, jMetaType :: Type
}
deriving (Int -> Judgement a -> ShowS
[Judgement a] -> ShowS
Judgement a -> String
(Int -> Judgement a -> ShowS)
-> (Judgement a -> String)
-> ([Judgement a] -> ShowS)
-> Show (Judgement a)
forall a. Show a => Int -> Judgement a -> ShowS
forall a. Show a => [Judgement a] -> ShowS
forall a. Show a => Judgement a -> String
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: forall a. Show a => Int -> Judgement a -> ShowS
showsPrec :: Int -> Judgement a -> ShowS
$cshow :: forall a. Show a => Judgement a -> String
show :: Judgement a -> String
$cshowList :: forall a. Show a => [Judgement a] -> ShowS
showList :: [Judgement a] -> ShowS
Show, (forall x. Judgement a -> Rep (Judgement a) x)
-> (forall x. Rep (Judgement a) x -> Judgement a)
-> Generic (Judgement a)
forall x. Rep (Judgement a) x -> Judgement a
forall x. Judgement a -> Rep (Judgement a) x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
forall a x. Rep (Judgement a) x -> Judgement a
forall a x. Judgement a -> Rep (Judgement a) x
$cfrom :: forall a x. Judgement a -> Rep (Judgement a) x
from :: forall x. Judgement a -> Rep (Judgement a) x
$cto :: forall a x. Rep (Judgement a) x -> Judgement a
to :: forall x. Rep (Judgement a) x -> Judgement a
Generic)
instance Pretty a => Pretty (Judgement a) where
pretty :: Judgement a -> Doc
pretty (HasType a
a Comparison
cmp Type
t) = [Doc] -> Doc
forall (t :: * -> *). Foldable t => t Doc -> Doc
hsep [ a -> Doc
forall a. Pretty a => a -> Doc
pretty a
a, Doc
":" , Type -> Doc
forall a. Pretty a => a -> Doc
pretty Type
t ]
pretty (IsSort a
a Type
t) = [Doc] -> Doc
forall (t :: * -> *). Foldable t => t Doc -> Doc
hsep [ a -> Doc
forall a. Pretty a => a -> Doc
pretty a
a, Doc
":sort", Type -> Doc
forall a. Pretty a => a -> Doc
pretty Type
t ]
data DoGeneralize
= YesGeneralizeVar
| YesGeneralizeMeta
| NoGeneralize
deriving (DoGeneralize -> DoGeneralize -> Bool
(DoGeneralize -> DoGeneralize -> Bool)
-> (DoGeneralize -> DoGeneralize -> Bool) -> Eq DoGeneralize
forall a. (a -> a -> Bool) -> (a -> a -> Bool) -> Eq a
$c== :: DoGeneralize -> DoGeneralize -> Bool
== :: DoGeneralize -> DoGeneralize -> Bool
$c/= :: DoGeneralize -> DoGeneralize -> Bool
/= :: DoGeneralize -> DoGeneralize -> Bool
Eq, Eq DoGeneralize
Eq DoGeneralize =>
(DoGeneralize -> DoGeneralize -> Ordering)
-> (DoGeneralize -> DoGeneralize -> Bool)
-> (DoGeneralize -> DoGeneralize -> Bool)
-> (DoGeneralize -> DoGeneralize -> Bool)
-> (DoGeneralize -> DoGeneralize -> Bool)
-> (DoGeneralize -> DoGeneralize -> DoGeneralize)
-> (DoGeneralize -> DoGeneralize -> DoGeneralize)
-> Ord DoGeneralize
DoGeneralize -> DoGeneralize -> Bool
DoGeneralize -> DoGeneralize -> Ordering
DoGeneralize -> DoGeneralize -> DoGeneralize
forall a.
Eq a =>
(a -> a -> Ordering)
-> (a -> a -> Bool)
-> (a -> a -> Bool)
-> (a -> a -> Bool)
-> (a -> a -> Bool)
-> (a -> a -> a)
-> (a -> a -> a)
-> Ord a
$ccompare :: DoGeneralize -> DoGeneralize -> Ordering
compare :: DoGeneralize -> DoGeneralize -> Ordering
$c< :: DoGeneralize -> DoGeneralize -> Bool
< :: DoGeneralize -> DoGeneralize -> Bool
$c<= :: DoGeneralize -> DoGeneralize -> Bool
<= :: DoGeneralize -> DoGeneralize -> Bool
$c> :: DoGeneralize -> DoGeneralize -> Bool
> :: DoGeneralize -> DoGeneralize -> Bool
$c>= :: DoGeneralize -> DoGeneralize -> Bool
>= :: DoGeneralize -> DoGeneralize -> Bool
$cmax :: DoGeneralize -> DoGeneralize -> DoGeneralize
max :: DoGeneralize -> DoGeneralize -> DoGeneralize
$cmin :: DoGeneralize -> DoGeneralize -> DoGeneralize
min :: DoGeneralize -> DoGeneralize -> DoGeneralize
Ord, Int -> DoGeneralize -> ShowS
[DoGeneralize] -> ShowS
DoGeneralize -> String
(Int -> DoGeneralize -> ShowS)
-> (DoGeneralize -> String)
-> ([DoGeneralize] -> ShowS)
-> Show DoGeneralize
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> DoGeneralize -> ShowS
showsPrec :: Int -> DoGeneralize -> ShowS
$cshow :: DoGeneralize -> String
show :: DoGeneralize -> String
$cshowList :: [DoGeneralize] -> ShowS
showList :: [DoGeneralize] -> ShowS
Show, (forall x. DoGeneralize -> Rep DoGeneralize x)
-> (forall x. Rep DoGeneralize x -> DoGeneralize)
-> Generic DoGeneralize
forall x. Rep DoGeneralize x -> DoGeneralize
forall x. DoGeneralize -> Rep DoGeneralize x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. DoGeneralize -> Rep DoGeneralize x
from :: forall x. DoGeneralize -> Rep DoGeneralize x
$cto :: forall x. Rep DoGeneralize x -> DoGeneralize
to :: forall x. Rep DoGeneralize x -> DoGeneralize
Generic)
data GeneralizedValue = GeneralizedValue
{ GeneralizedValue -> CheckpointId
genvalCheckpoint :: CheckpointId
, GeneralizedValue -> Term
genvalTerm :: Term
, GeneralizedValue -> Type
genvalType :: Type
} deriving (Int -> GeneralizedValue -> ShowS
[GeneralizedValue] -> ShowS
GeneralizedValue -> String
(Int -> GeneralizedValue -> ShowS)
-> (GeneralizedValue -> String)
-> ([GeneralizedValue] -> ShowS)
-> Show GeneralizedValue
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> GeneralizedValue -> ShowS
showsPrec :: Int -> GeneralizedValue -> ShowS
$cshow :: GeneralizedValue -> String
show :: GeneralizedValue -> String
$cshowList :: [GeneralizedValue] -> ShowS
showList :: [GeneralizedValue] -> ShowS
Show, (forall x. GeneralizedValue -> Rep GeneralizedValue x)
-> (forall x. Rep GeneralizedValue x -> GeneralizedValue)
-> Generic GeneralizedValue
forall x. Rep GeneralizedValue x -> GeneralizedValue
forall x. GeneralizedValue -> Rep GeneralizedValue x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. GeneralizedValue -> Rep GeneralizedValue x
from :: forall x. GeneralizedValue -> Rep GeneralizedValue x
$cto :: forall x. Rep GeneralizedValue x -> GeneralizedValue
to :: forall x. Rep GeneralizedValue x -> GeneralizedValue
Generic)
data MetaVariable =
MetaVar { MetaVariable -> MetaInfo
mvInfo :: MetaInfo
, MetaVariable -> MetaPriority
mvPriority :: MetaPriority
, MetaVariable -> Permutation
mvPermutation :: Permutation
, MetaVariable -> Judgement MetaId
mvJudgement :: Judgement MetaId
, MetaVariable -> MetaInstantiation
mvInstantiation :: MetaInstantiation
, MetaVariable -> Set Listener
mvListeners :: Set Listener
, MetaVariable -> Frozen
mvFrozen :: Frozen
, MetaVariable -> Maybe MetaId
mvTwin :: Maybe MetaId
}
deriving (forall x. MetaVariable -> Rep MetaVariable x)
-> (forall x. Rep MetaVariable x -> MetaVariable)
-> Generic MetaVariable
forall x. Rep MetaVariable x -> MetaVariable
forall x. MetaVariable -> Rep MetaVariable x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. MetaVariable -> Rep MetaVariable x
from :: forall x. MetaVariable -> Rep MetaVariable x
$cto :: forall x. Rep MetaVariable x -> MetaVariable
to :: forall x. Rep MetaVariable x -> MetaVariable
Generic
data Listener = EtaExpand MetaId
| CheckConstraint Nat ProblemConstraint
deriving (forall x. Listener -> Rep Listener x)
-> (forall x. Rep Listener x -> Listener) -> Generic Listener
forall x. Rep Listener x -> Listener
forall x. Listener -> Rep Listener x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. Listener -> Rep Listener x
from :: forall x. Listener -> Rep Listener x
$cto :: forall x. Rep Listener x -> Listener
to :: forall x. Rep Listener x -> Listener
Generic
instance Eq Listener where
EtaExpand MetaId
x == :: Listener -> Listener -> Bool
== EtaExpand MetaId
y = MetaId
x MetaId -> MetaId -> Bool
forall a. Eq a => a -> a -> Bool
== MetaId
y
CheckConstraint Int
x ProblemConstraint
_ == CheckConstraint Int
y ProblemConstraint
_ = Int
x Int -> Int -> Bool
forall a. Eq a => a -> a -> Bool
== Int
y
Listener
_ == Listener
_ = Bool
False
instance Ord Listener where
EtaExpand MetaId
x compare :: Listener -> Listener -> Ordering
`compare` EtaExpand MetaId
y = MetaId
x MetaId -> MetaId -> Ordering
forall a. Ord a => a -> a -> Ordering
`compare` MetaId
y
CheckConstraint Int
x ProblemConstraint
_ `compare` CheckConstraint Int
y ProblemConstraint
_ = Int
x Int -> Int -> Ordering
forall a. Ord a => a -> a -> Ordering
`compare` Int
y
EtaExpand{} `compare` CheckConstraint{} = Ordering
LT
CheckConstraint{} `compare` EtaExpand{} = Ordering
GT
data Frozen
= Frozen
| Instantiable
deriving (Frozen -> Frozen -> Bool
(Frozen -> Frozen -> Bool)
-> (Frozen -> Frozen -> Bool) -> Eq Frozen
forall a. (a -> a -> Bool) -> (a -> a -> Bool) -> Eq a
$c== :: Frozen -> Frozen -> Bool
== :: Frozen -> Frozen -> Bool
$c/= :: Frozen -> Frozen -> Bool
/= :: Frozen -> Frozen -> Bool
Eq, Int -> Frozen -> ShowS
[Frozen] -> ShowS
Frozen -> String
(Int -> Frozen -> ShowS)
-> (Frozen -> String) -> ([Frozen] -> ShowS) -> Show Frozen
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> Frozen -> ShowS
showsPrec :: Int -> Frozen -> ShowS
$cshow :: Frozen -> String
show :: Frozen -> String
$cshowList :: [Frozen] -> ShowS
showList :: [Frozen] -> ShowS
Show, (forall x. Frozen -> Rep Frozen x)
-> (forall x. Rep Frozen x -> Frozen) -> Generic Frozen
forall x. Rep Frozen x -> Frozen
forall x. Frozen -> Rep Frozen x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. Frozen -> Rep Frozen x
from :: forall x. Frozen -> Rep Frozen x
$cto :: forall x. Rep Frozen x -> Frozen
to :: forall x. Rep Frozen x -> Frozen
Generic)
data MetaInstantiation
= InstV Instantiation
| OpenMeta MetaKind
| BlockedConst Term
| PostponedTypeCheckingProblem (Closure TypeCheckingProblem)
deriving (forall x. MetaInstantiation -> Rep MetaInstantiation x)
-> (forall x. Rep MetaInstantiation x -> MetaInstantiation)
-> Generic MetaInstantiation
forall x. Rep MetaInstantiation x -> MetaInstantiation
forall x. MetaInstantiation -> Rep MetaInstantiation x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. MetaInstantiation -> Rep MetaInstantiation x
from :: forall x. MetaInstantiation -> Rep MetaInstantiation x
$cto :: forall x. Rep MetaInstantiation x -> MetaInstantiation
to :: forall x. Rep MetaInstantiation x -> MetaInstantiation
Generic
data Instantiation = Instantiation
{ Instantiation -> [Arg String]
instTel :: [Arg String]
, Instantiation -> Term
instBody :: Term
}
deriving (Int -> Instantiation -> ShowS
[Instantiation] -> ShowS
Instantiation -> String
(Int -> Instantiation -> ShowS)
-> (Instantiation -> String)
-> ([Instantiation] -> ShowS)
-> Show Instantiation
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> Instantiation -> ShowS
showsPrec :: Int -> Instantiation -> ShowS
$cshow :: Instantiation -> String
show :: Instantiation -> String
$cshowList :: [Instantiation] -> ShowS
showList :: [Instantiation] -> ShowS
Show, (forall x. Instantiation -> Rep Instantiation x)
-> (forall x. Rep Instantiation x -> Instantiation)
-> Generic Instantiation
forall x. Rep Instantiation x -> Instantiation
forall x. Instantiation -> Rep Instantiation x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. Instantiation -> Rep Instantiation x
from :: forall x. Instantiation -> Rep Instantiation x
$cto :: forall x. Rep Instantiation x -> Instantiation
to :: forall x. Rep Instantiation x -> Instantiation
Generic)
data RemoteMetaVariable = RemoteMetaVariable
{ RemoteMetaVariable -> Instantiation
rmvInstantiation :: Instantiation
, RemoteMetaVariable -> Modality
rmvModality :: Modality
, RemoteMetaVariable -> Judgement MetaId
rmvJudgement :: Judgement MetaId
}
deriving (Int -> RemoteMetaVariable -> ShowS
[RemoteMetaVariable] -> ShowS
RemoteMetaVariable -> String
(Int -> RemoteMetaVariable -> ShowS)
-> (RemoteMetaVariable -> String)
-> ([RemoteMetaVariable] -> ShowS)
-> Show RemoteMetaVariable
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> RemoteMetaVariable -> ShowS
showsPrec :: Int -> RemoteMetaVariable -> ShowS
$cshow :: RemoteMetaVariable -> String
show :: RemoteMetaVariable -> String
$cshowList :: [RemoteMetaVariable] -> ShowS
showList :: [RemoteMetaVariable] -> ShowS
Show, (forall x. RemoteMetaVariable -> Rep RemoteMetaVariable x)
-> (forall x. Rep RemoteMetaVariable x -> RemoteMetaVariable)
-> Generic RemoteMetaVariable
forall x. Rep RemoteMetaVariable x -> RemoteMetaVariable
forall x. RemoteMetaVariable -> Rep RemoteMetaVariable x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. RemoteMetaVariable -> Rep RemoteMetaVariable x
from :: forall x. RemoteMetaVariable -> Rep RemoteMetaVariable x
$cto :: forall x. Rep RemoteMetaVariable x -> RemoteMetaVariable
to :: forall x. Rep RemoteMetaVariable x -> RemoteMetaVariable
Generic)
data CheckedTarget = CheckedTarget (Maybe ProblemId)
| NotCheckedTarget
data PrincipalArgTypeMetas = PrincipalArgTypeMetas
{ PrincipalArgTypeMetas -> Args
patmMetas :: Args
, PrincipalArgTypeMetas -> Type
patmRemainder :: Type
}
deriving (forall x. PrincipalArgTypeMetas -> Rep PrincipalArgTypeMetas x)
-> (forall x. Rep PrincipalArgTypeMetas x -> PrincipalArgTypeMetas)
-> Generic PrincipalArgTypeMetas
forall x. Rep PrincipalArgTypeMetas x -> PrincipalArgTypeMetas
forall x. PrincipalArgTypeMetas -> Rep PrincipalArgTypeMetas x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. PrincipalArgTypeMetas -> Rep PrincipalArgTypeMetas x
from :: forall x. PrincipalArgTypeMetas -> Rep PrincipalArgTypeMetas x
$cto :: forall x. Rep PrincipalArgTypeMetas x -> PrincipalArgTypeMetas
to :: forall x. Rep PrincipalArgTypeMetas x -> PrincipalArgTypeMetas
Generic
data TypeCheckingProblem
= CheckExpr Comparison A.Expr Type
| CheckArgs Comparison ExpandHidden Range [NamedArg A.Expr] Type Type (ArgsCheckState CheckedTarget -> TCM Term)
| CheckProjAppToKnownPrincipalArg Comparison A.Expr ProjOrigin (List1 QName) A.Args Type Int Term Type PrincipalArgTypeMetas
| CheckLambda Comparison (Arg (List1 (WithHiding Name), Maybe Type)) A.Expr Type
| DoQuoteTerm Comparison Term Type
deriving (forall x. TypeCheckingProblem -> Rep TypeCheckingProblem x)
-> (forall x. Rep TypeCheckingProblem x -> TypeCheckingProblem)
-> Generic TypeCheckingProblem
forall x. Rep TypeCheckingProblem x -> TypeCheckingProblem
forall x. TypeCheckingProblem -> Rep TypeCheckingProblem x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. TypeCheckingProblem -> Rep TypeCheckingProblem x
from :: forall x. TypeCheckingProblem -> Rep TypeCheckingProblem x
$cto :: forall x. Rep TypeCheckingProblem x -> TypeCheckingProblem
to :: forall x. Rep TypeCheckingProblem x -> TypeCheckingProblem
Generic
instance Pretty MetaInstantiation where
pretty :: MetaInstantiation -> Doc
pretty = \case
OpenMeta MetaKind
UnificationMeta -> Doc
"Open"
OpenMeta MetaKind
InstanceMeta -> Doc
"OpenInstance"
PostponedTypeCheckingProblem{} -> Doc
"PostponedTypeCheckingProblem (...)"
BlockedConst Term
t -> [Doc] -> Doc
forall (t :: * -> *). Foldable t => t Doc -> Doc
hsep [ Doc
"BlockedConst", Doc -> Doc
parens (Term -> Doc
forall a. Pretty a => a -> Doc
pretty Term
t) ]
InstV Instantiation{ [Arg String]
instTel :: Instantiation -> [Arg String]
instTel :: [Arg String]
instTel, Term
instBody :: Instantiation -> Term
instBody :: Term
instBody } -> [Doc] -> Doc
forall (t :: * -> *). Foldable t => t Doc -> Doc
hsep [ Doc
"InstV", [Arg String] -> Doc
forall a. Pretty a => a -> Doc
pretty [Arg String]
instTel, Doc -> Doc
parens (Term -> Doc
forall a. Pretty a => a -> Doc
pretty Term
instBody) ]
newtype MetaPriority = MetaPriority Int
deriving (MetaPriority -> MetaPriority -> Bool
(MetaPriority -> MetaPriority -> Bool)
-> (MetaPriority -> MetaPriority -> Bool) -> Eq MetaPriority
forall a. (a -> a -> Bool) -> (a -> a -> Bool) -> Eq a
$c== :: MetaPriority -> MetaPriority -> Bool
== :: MetaPriority -> MetaPriority -> Bool
$c/= :: MetaPriority -> MetaPriority -> Bool
/= :: MetaPriority -> MetaPriority -> Bool
Eq, Eq MetaPriority
Eq MetaPriority =>
(MetaPriority -> MetaPriority -> Ordering)
-> (MetaPriority -> MetaPriority -> Bool)
-> (MetaPriority -> MetaPriority -> Bool)
-> (MetaPriority -> MetaPriority -> Bool)
-> (MetaPriority -> MetaPriority -> Bool)
-> (MetaPriority -> MetaPriority -> MetaPriority)
-> (MetaPriority -> MetaPriority -> MetaPriority)
-> Ord MetaPriority
MetaPriority -> MetaPriority -> Bool
MetaPriority -> MetaPriority -> Ordering
MetaPriority -> MetaPriority -> MetaPriority
forall a.
Eq a =>
(a -> a -> Ordering)
-> (a -> a -> Bool)
-> (a -> a -> Bool)
-> (a -> a -> Bool)
-> (a -> a -> Bool)
-> (a -> a -> a)
-> (a -> a -> a)
-> Ord a
$ccompare :: MetaPriority -> MetaPriority -> Ordering
compare :: MetaPriority -> MetaPriority -> Ordering
$c< :: MetaPriority -> MetaPriority -> Bool
< :: MetaPriority -> MetaPriority -> Bool
$c<= :: MetaPriority -> MetaPriority -> Bool
<= :: MetaPriority -> MetaPriority -> Bool
$c> :: MetaPriority -> MetaPriority -> Bool
> :: MetaPriority -> MetaPriority -> Bool
$c>= :: MetaPriority -> MetaPriority -> Bool
>= :: MetaPriority -> MetaPriority -> Bool
$cmax :: MetaPriority -> MetaPriority -> MetaPriority
max :: MetaPriority -> MetaPriority -> MetaPriority
$cmin :: MetaPriority -> MetaPriority -> MetaPriority
min :: MetaPriority -> MetaPriority -> MetaPriority
Ord, Int -> MetaPriority -> ShowS
[MetaPriority] -> ShowS
MetaPriority -> String
(Int -> MetaPriority -> ShowS)
-> (MetaPriority -> String)
-> ([MetaPriority] -> ShowS)
-> Show MetaPriority
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> MetaPriority -> ShowS
showsPrec :: Int -> MetaPriority -> ShowS
$cshow :: MetaPriority -> String
show :: MetaPriority -> String
$cshowList :: [MetaPriority] -> ShowS
showList :: [MetaPriority] -> ShowS
Show, MetaPriority -> ()
(MetaPriority -> ()) -> NFData MetaPriority
forall a. (a -> ()) -> NFData a
$crnf :: MetaPriority -> ()
rnf :: MetaPriority -> ()
NFData)
data RunMetaOccursCheck
= RunMetaOccursCheck
| DontRunMetaOccursCheck
deriving (RunMetaOccursCheck -> RunMetaOccursCheck -> Bool
(RunMetaOccursCheck -> RunMetaOccursCheck -> Bool)
-> (RunMetaOccursCheck -> RunMetaOccursCheck -> Bool)
-> Eq RunMetaOccursCheck
forall a. (a -> a -> Bool) -> (a -> a -> Bool) -> Eq a
$c== :: RunMetaOccursCheck -> RunMetaOccursCheck -> Bool
== :: RunMetaOccursCheck -> RunMetaOccursCheck -> Bool
$c/= :: RunMetaOccursCheck -> RunMetaOccursCheck -> Bool
/= :: RunMetaOccursCheck -> RunMetaOccursCheck -> Bool
Eq, Eq RunMetaOccursCheck
Eq RunMetaOccursCheck =>
(RunMetaOccursCheck -> RunMetaOccursCheck -> Ordering)
-> (RunMetaOccursCheck -> RunMetaOccursCheck -> Bool)
-> (RunMetaOccursCheck -> RunMetaOccursCheck -> Bool)
-> (RunMetaOccursCheck -> RunMetaOccursCheck -> Bool)
-> (RunMetaOccursCheck -> RunMetaOccursCheck -> Bool)
-> (RunMetaOccursCheck -> RunMetaOccursCheck -> RunMetaOccursCheck)
-> (RunMetaOccursCheck -> RunMetaOccursCheck -> RunMetaOccursCheck)
-> Ord RunMetaOccursCheck
RunMetaOccursCheck -> RunMetaOccursCheck -> Bool
RunMetaOccursCheck -> RunMetaOccursCheck -> Ordering
RunMetaOccursCheck -> RunMetaOccursCheck -> RunMetaOccursCheck
forall a.
Eq a =>
(a -> a -> Ordering)
-> (a -> a -> Bool)
-> (a -> a -> Bool)
-> (a -> a -> Bool)
-> (a -> a -> Bool)
-> (a -> a -> a)
-> (a -> a -> a)
-> Ord a
$ccompare :: RunMetaOccursCheck -> RunMetaOccursCheck -> Ordering
compare :: RunMetaOccursCheck -> RunMetaOccursCheck -> Ordering
$c< :: RunMetaOccursCheck -> RunMetaOccursCheck -> Bool
< :: RunMetaOccursCheck -> RunMetaOccursCheck -> Bool
$c<= :: RunMetaOccursCheck -> RunMetaOccursCheck -> Bool
<= :: RunMetaOccursCheck -> RunMetaOccursCheck -> Bool
$c> :: RunMetaOccursCheck -> RunMetaOccursCheck -> Bool
> :: RunMetaOccursCheck -> RunMetaOccursCheck -> Bool
$c>= :: RunMetaOccursCheck -> RunMetaOccursCheck -> Bool
>= :: RunMetaOccursCheck -> RunMetaOccursCheck -> Bool
$cmax :: RunMetaOccursCheck -> RunMetaOccursCheck -> RunMetaOccursCheck
max :: RunMetaOccursCheck -> RunMetaOccursCheck -> RunMetaOccursCheck
$cmin :: RunMetaOccursCheck -> RunMetaOccursCheck -> RunMetaOccursCheck
min :: RunMetaOccursCheck -> RunMetaOccursCheck -> RunMetaOccursCheck
Ord, Int -> RunMetaOccursCheck -> ShowS
[RunMetaOccursCheck] -> ShowS
RunMetaOccursCheck -> String
(Int -> RunMetaOccursCheck -> ShowS)
-> (RunMetaOccursCheck -> String)
-> ([RunMetaOccursCheck] -> ShowS)
-> Show RunMetaOccursCheck
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> RunMetaOccursCheck -> ShowS
showsPrec :: Int -> RunMetaOccursCheck -> ShowS
$cshow :: RunMetaOccursCheck -> String
show :: RunMetaOccursCheck -> String
$cshowList :: [RunMetaOccursCheck] -> ShowS
showList :: [RunMetaOccursCheck] -> ShowS
Show, (forall x. RunMetaOccursCheck -> Rep RunMetaOccursCheck x)
-> (forall x. Rep RunMetaOccursCheck x -> RunMetaOccursCheck)
-> Generic RunMetaOccursCheck
forall x. Rep RunMetaOccursCheck x -> RunMetaOccursCheck
forall x. RunMetaOccursCheck -> Rep RunMetaOccursCheck x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. RunMetaOccursCheck -> Rep RunMetaOccursCheck x
from :: forall x. RunMetaOccursCheck -> Rep RunMetaOccursCheck x
$cto :: forall x. Rep RunMetaOccursCheck x -> RunMetaOccursCheck
to :: forall x. Rep RunMetaOccursCheck x -> RunMetaOccursCheck
Generic)
data MetaInfo = MetaInfo
{ MetaInfo -> Closure Range
miClosRange :: Closure Range
, MetaInfo -> Modality
miModality :: Modality
, MetaInfo -> RunMetaOccursCheck
miMetaOccursCheck :: RunMetaOccursCheck
, MetaInfo -> String
miNameSuggestion :: MetaNameSuggestion
, MetaInfo -> Arg DoGeneralize
miGeneralizable :: Arg DoGeneralize
}
deriving (forall x. MetaInfo -> Rep MetaInfo x)
-> (forall x. Rep MetaInfo x -> MetaInfo) -> Generic MetaInfo
forall x. Rep MetaInfo x -> MetaInfo
forall x. MetaInfo -> Rep MetaInfo x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. MetaInfo -> Rep MetaInfo x
from :: forall x. MetaInfo -> Rep MetaInfo x
$cto :: forall x. Rep MetaInfo x -> MetaInfo
to :: forall x. Rep MetaInfo x -> MetaInfo
Generic
instance LensModality MetaInfo where
getModality :: MetaInfo -> Modality
getModality = MetaInfo -> Modality
miModality
setModality :: Modality -> MetaInfo -> MetaInfo
setModality Modality
mod MetaInfo
mi = MetaInfo
mi { miModality = mod }
mapModality :: (Modality -> Modality) -> MetaInfo -> MetaInfo
mapModality Modality -> Modality
f MetaInfo
mi = MetaInfo
mi { miModality = f $ miModality mi }
instance LensQuantity MetaInfo where
getQuantity :: MetaInfo -> Quantity
getQuantity = Modality -> Quantity
forall a. LensQuantity a => a -> Quantity
getQuantity (Modality -> Quantity)
-> (MetaInfo -> Modality) -> MetaInfo -> Quantity
forall b c a. (b -> c) -> (a -> b) -> a -> c
. MetaInfo -> Modality
forall a. LensModality a => a -> Modality
getModality
mapQuantity :: (Quantity -> Quantity) -> MetaInfo -> MetaInfo
mapQuantity Quantity -> Quantity
f = (Modality -> Modality) -> MetaInfo -> MetaInfo
forall a. LensModality a => (Modality -> Modality) -> a -> a
mapModality ((Quantity -> Quantity) -> Modality -> Modality
forall a. LensQuantity a => (Quantity -> Quantity) -> a -> a
mapQuantity Quantity -> Quantity
f)
instance LensRelevance MetaInfo where
mapRelevance :: (Relevance -> Relevance) -> MetaInfo -> MetaInfo
mapRelevance Relevance -> Relevance
f = (Modality -> Modality) -> MetaInfo -> MetaInfo
forall a. LensModality a => (Modality -> Modality) -> a -> a
mapModality ((Relevance -> Relevance) -> Modality -> Modality
forall a. LensRelevance a => (Relevance -> Relevance) -> a -> a
mapRelevance Relevance -> Relevance
f)
suffixNameSuggestion :: MetaNameSuggestion -> ArgName -> MetaNameSuggestion
suffixNameSuggestion :: String -> ShowS
suffixNameSuggestion String
"_" String
field = String
field
suffixNameSuggestion String
"" String
field = String
field
suffixNameSuggestion String
record String
field = String
record String -> ShowS
forall a. [a] -> [a] -> [a]
++ String
"." String -> ShowS
forall a. [a] -> [a] -> [a]
++ String
field
instance Pretty NamedMeta where
pretty :: NamedMeta -> Doc
pretty (NamedMeta String
"" MetaId
x) = MetaId -> Doc
forall a. Pretty a => a -> Doc
pretty MetaId
x
pretty (NamedMeta String
"_" MetaId
x) = MetaId -> Doc
forall a. Pretty a => a -> Doc
pretty MetaId
x
pretty (NamedMeta String
s MetaId
x) = String -> Doc
forall a. String -> Doc a
text (String -> Doc) -> String -> Doc
forall a b. (a -> b) -> a -> b
$ String
"_" String -> ShowS
forall a. [a] -> [a] -> [a]
++ String
s String -> ShowS
forall a. [a] -> [a] -> [a]
++ MetaId -> String
forall a. Pretty a => a -> String
prettyShow MetaId
x
type LocalMetaStore = Map MetaId MetaVariable
{-# SPECIALIZE Map.insert :: MetaId -> v -> Map MetaId v -> Map MetaId v #-}
{-# SPECIALIZE Map.lookup :: MetaId -> Map MetaId v -> Maybe v #-}
type RemoteMetaStore = HashMap MetaId RemoteMetaVariable
instance HasRange MetaInfo where
getRange :: MetaInfo -> Range
getRange = Closure Range -> Range
forall a. Closure a -> a
clValue (Closure Range -> Range)
-> (MetaInfo -> Closure Range) -> MetaInfo -> Range
forall b c a. (b -> c) -> (a -> b) -> a -> c
. MetaInfo -> Closure Range
miClosRange
instance HasRange MetaVariable where
getRange :: MetaVariable -> Range
getRange MetaVariable
m = Closure Range -> Range
forall a. HasRange a => a -> Range
getRange (Closure Range -> Range) -> Closure Range -> Range
forall a b. (a -> b) -> a -> b
$ MetaVariable -> Closure Range
getMetaInfo MetaVariable
m
instance SetRange MetaInfo where
setRange :: Range -> MetaInfo -> MetaInfo
setRange Range
r MetaInfo
m = MetaInfo
m { miClosRange = (miClosRange m) { clValue = r }}
instance SetRange MetaVariable where
setRange :: Range -> MetaVariable -> MetaVariable
setRange Range
r MetaVariable
m = MetaVariable
m { mvInfo = setRange r (mvInfo m) }
instance LensModality MetaVariable where
getModality :: MetaVariable -> Modality
getModality = MetaInfo -> Modality
forall a. LensModality a => a -> Modality
getModality (MetaInfo -> Modality)
-> (MetaVariable -> MetaInfo) -> MetaVariable -> Modality
forall b c a. (b -> c) -> (a -> b) -> a -> c
. MetaVariable -> MetaInfo
mvInfo
setModality :: Modality -> MetaVariable -> MetaVariable
setModality Modality
mod MetaVariable
mv = MetaVariable
mv { mvInfo = setModality mod $ mvInfo mv }
mapModality :: (Modality -> Modality) -> MetaVariable -> MetaVariable
mapModality Modality -> Modality
f MetaVariable
mv = MetaVariable
mv { mvInfo = mapModality f $ mvInfo mv }
instance LensRelevance MetaVariable where
setRelevance :: Relevance -> MetaVariable -> MetaVariable
setRelevance Relevance
mod MetaVariable
mv = MetaVariable
mv { mvInfo = setRelevance mod $ mvInfo mv }
instance LensQuantity MetaVariable where
getQuantity :: MetaVariable -> Quantity
getQuantity = Modality -> Quantity
forall a. LensQuantity a => a -> Quantity
getQuantity (Modality -> Quantity)
-> (MetaVariable -> Modality) -> MetaVariable -> Quantity
forall b c a. (b -> c) -> (a -> b) -> a -> c
. MetaVariable -> Modality
forall a. LensModality a => a -> Modality
getModality
mapQuantity :: (Quantity -> Quantity) -> MetaVariable -> MetaVariable
mapQuantity Quantity -> Quantity
f = (Modality -> Modality) -> MetaVariable -> MetaVariable
forall a. LensModality a => (Modality -> Modality) -> a -> a
mapModality ((Quantity -> Quantity) -> Modality -> Modality
forall a. LensQuantity a => (Quantity -> Quantity) -> a -> a
mapQuantity Quantity -> Quantity
f)
instance LensModality RemoteMetaVariable where
getModality :: RemoteMetaVariable -> Modality
getModality = RemoteMetaVariable -> Modality
rmvModality
mapModality :: (Modality -> Modality) -> RemoteMetaVariable -> RemoteMetaVariable
mapModality Modality -> Modality
f RemoteMetaVariable
mv = RemoteMetaVariable
mv { rmvModality = f $ rmvModality mv }
instance LensRelevance RemoteMetaVariable where
mapRelevance :: (Relevance -> Relevance)
-> RemoteMetaVariable -> RemoteMetaVariable
mapRelevance Relevance -> Relevance
f = (Modality -> Modality) -> RemoteMetaVariable -> RemoteMetaVariable
forall a. LensModality a => (Modality -> Modality) -> a -> a
mapModality ((Relevance -> Relevance) -> Modality -> Modality
forall a. LensRelevance a => (Relevance -> Relevance) -> a -> a
mapRelevance Relevance -> Relevance
f)
instance LensQuantity RemoteMetaVariable where
mapQuantity :: (Quantity -> Quantity) -> RemoteMetaVariable -> RemoteMetaVariable
mapQuantity Quantity -> Quantity
f = (Modality -> Modality) -> RemoteMetaVariable -> RemoteMetaVariable
forall a. LensModality a => (Modality -> Modality) -> a -> a
mapModality ((Quantity -> Quantity) -> Modality -> Modality
forall a. LensQuantity a => (Quantity -> Quantity) -> a -> a
mapQuantity Quantity -> Quantity
f)
normalMetaPriority :: MetaPriority
normalMetaPriority :: MetaPriority
normalMetaPriority = Int -> MetaPriority
MetaPriority Int
0
lowMetaPriority :: MetaPriority
lowMetaPriority :: MetaPriority
lowMetaPriority = Int -> MetaPriority
MetaPriority (-Int
10)
highMetaPriority :: MetaPriority
highMetaPriority :: MetaPriority
highMetaPriority = Int -> MetaPriority
MetaPriority Int
10
getMetaInfo :: MetaVariable -> Closure Range
getMetaInfo :: MetaVariable -> Closure Range
getMetaInfo = MetaInfo -> Closure Range
miClosRange (MetaInfo -> Closure Range)
-> (MetaVariable -> MetaInfo) -> MetaVariable -> Closure Range
forall b c a. (b -> c) -> (a -> b) -> a -> c
. MetaVariable -> MetaInfo
mvInfo
getMetaScope :: MetaVariable -> ScopeInfo
getMetaScope :: MetaVariable -> ScopeInfo
getMetaScope MetaVariable
m = Closure Range -> ScopeInfo
forall a. Closure a -> ScopeInfo
clScope (Closure Range -> ScopeInfo) -> Closure Range -> ScopeInfo
forall a b. (a -> b) -> a -> b
$ MetaVariable -> Closure Range
getMetaInfo MetaVariable
m
getMetaEnv :: MetaVariable -> TCEnv
getMetaEnv :: MetaVariable -> TCEnv
getMetaEnv MetaVariable
m = Closure Range -> TCEnv
forall a. Closure a -> TCEnv
clEnv (Closure Range -> TCEnv) -> Closure Range -> TCEnv
forall a b. (a -> b) -> a -> b
$ MetaVariable -> Closure Range
getMetaInfo MetaVariable
m
getMetaSig :: MetaVariable -> Signature
getMetaSig :: MetaVariable -> Signature
getMetaSig MetaVariable
m = Closure Range -> Signature
forall a. Closure a -> Signature
clSignature (Closure Range -> Signature) -> Closure Range -> Signature
forall a b. (a -> b) -> a -> b
$ MetaVariable -> Closure Range
getMetaInfo MetaVariable
m
metaFrozen :: Lens' MetaVariable Frozen
metaFrozen :: Lens' MetaVariable Frozen
metaFrozen Frozen -> f Frozen
f MetaVariable
mv = Frozen -> f Frozen
f (MetaVariable -> Frozen
mvFrozen MetaVariable
mv) f Frozen -> (Frozen -> MetaVariable) -> f MetaVariable
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ Frozen
x -> MetaVariable
mv { mvFrozen = x }
_mvInfo :: Lens' MetaVariable MetaInfo
_mvInfo :: Lens' MetaVariable MetaInfo
_mvInfo MetaInfo -> f MetaInfo
f MetaVariable
mv = (MetaInfo -> f MetaInfo
f (MetaInfo -> f MetaInfo) -> MetaInfo -> f MetaInfo
forall a b. (a -> b) -> a -> b
$! MetaVariable -> MetaInfo
mvInfo MetaVariable
mv) f MetaInfo -> (MetaInfo -> MetaVariable) -> f MetaVariable
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ MetaInfo
mi -> MetaVariable
mv { mvInfo = mi }
instance LensClosure MetaInfo Range where
lensClosure :: Lens' MetaInfo (Closure Range)
lensClosure Closure Range -> f (Closure Range)
f MetaInfo
mi = (Closure Range -> f (Closure Range)
f (Closure Range -> f (Closure Range))
-> Closure Range -> f (Closure Range)
forall a b. (a -> b) -> a -> b
$! MetaInfo -> Closure Range
miClosRange MetaInfo
mi) f (Closure Range) -> (Closure Range -> MetaInfo) -> f MetaInfo
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ Closure Range
cl -> MetaInfo
mi { miClosRange = cl }
instance LensClosure MetaVariable Range where
lensClosure :: Lens' MetaVariable (Closure Range)
lensClosure = (MetaInfo -> f MetaInfo) -> MetaVariable -> f MetaVariable
Lens' MetaVariable MetaInfo
_mvInfo ((MetaInfo -> f MetaInfo) -> MetaVariable -> f MetaVariable)
-> ((Closure Range -> f (Closure Range)) -> MetaInfo -> f MetaInfo)
-> (Closure Range -> f (Closure Range))
-> MetaVariable
-> f MetaVariable
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (Closure Range -> f (Closure Range)) -> MetaInfo -> f MetaInfo
forall b a. LensClosure b a => Lens' b (Closure a)
Lens' MetaInfo (Closure Range)
lensClosure
instance LensIsAbstract TCEnv where
lensIsAbstract :: Lens' TCEnv IsAbstract
lensIsAbstract IsAbstract -> f IsAbstract
f TCEnv
env =
(IsAbstract -> f IsAbstract
f (IsAbstract -> f IsAbstract) -> IsAbstract -> f IsAbstract
forall a b. (a -> b) -> a -> b
$! IsAbstract -> Maybe IsAbstract -> IsAbstract
forall a. a -> Maybe a -> a
fromMaybe IsAbstract
forall a. HasCallStack => a
__IMPOSSIBLE__ (AbstractMode -> Maybe IsAbstract
aModeToDef (AbstractMode -> Maybe IsAbstract)
-> AbstractMode -> Maybe IsAbstract
forall a b. (a -> b) -> a -> b
$ TCEnv -> AbstractMode
envAbstractMode TCEnv
env))
f IsAbstract -> (IsAbstract -> TCEnv) -> f TCEnv
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ IsAbstract
a -> TCEnv
env { envAbstractMode = aDefToMode a }
instance LensIsAbstract (Closure a) where
lensIsAbstract :: Lens' (Closure a) IsAbstract
lensIsAbstract = (TCEnv -> f TCEnv) -> Closure a -> f (Closure a)
forall a. LensTCEnv a => Lens' a TCEnv
Lens' (Closure a) TCEnv
lensTCEnv ((TCEnv -> f TCEnv) -> Closure a -> f (Closure a))
-> ((IsAbstract -> f IsAbstract) -> TCEnv -> f TCEnv)
-> (IsAbstract -> f IsAbstract)
-> Closure a
-> f (Closure a)
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (IsAbstract -> f IsAbstract) -> TCEnv -> f TCEnv
forall a. LensIsAbstract a => Lens' a IsAbstract
Lens' TCEnv IsAbstract
lensIsAbstract
instance LensIsAbstract MetaInfo where
lensIsAbstract :: Lens' MetaInfo IsAbstract
lensIsAbstract = (Closure Range -> f (Closure Range)) -> MetaInfo -> f MetaInfo
forall b a. LensClosure b a => Lens' b (Closure a)
Lens' MetaInfo (Closure Range)
lensClosure ((Closure Range -> f (Closure Range)) -> MetaInfo -> f MetaInfo)
-> ((IsAbstract -> f IsAbstract)
-> Closure Range -> f (Closure Range))
-> (IsAbstract -> f IsAbstract)
-> MetaInfo
-> f MetaInfo
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (IsAbstract -> f IsAbstract) -> Closure Range -> f (Closure Range)
forall a. LensIsAbstract a => Lens' a IsAbstract
Lens' (Closure Range) IsAbstract
lensIsAbstract
instance LensIsOpaque TCEnv where
lensIsOpaque :: Lens' TCEnv IsOpaque
lensIsOpaque IsOpaque -> f IsOpaque
f TCEnv
env =
(IsOpaque -> f IsOpaque
f (IsOpaque -> f IsOpaque) -> IsOpaque -> f IsOpaque
forall a b. (a -> b) -> a -> b
$! case TCEnv -> Maybe OpaqueId
envCurrentOpaqueId TCEnv
env of { Just OpaqueId
x -> OpaqueId -> IsOpaque
OpaqueDef OpaqueId
x ; Maybe OpaqueId
Nothing -> IsOpaque
TransparentDef })
f IsOpaque -> (IsOpaque -> TCEnv) -> f TCEnv
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \case { OpaqueDef OpaqueId
x -> TCEnv
env { envCurrentOpaqueId = Just x }
; IsOpaque
TransparentDef -> TCEnv
env { envCurrentOpaqueId = Nothing }
}
data InteractionPoint = InteractionPoint
{ InteractionPoint -> Range
ipRange :: Range
, InteractionPoint -> Maybe MetaId
ipMeta :: Maybe MetaId
, InteractionPoint -> Bool
ipSolved :: Bool
, InteractionPoint -> IPClause
ipClause :: IPClause
, InteractionPoint -> IPBoundary
ipBoundary :: IPBoundary
}
deriving (forall x. InteractionPoint -> Rep InteractionPoint x)
-> (forall x. Rep InteractionPoint x -> InteractionPoint)
-> Generic InteractionPoint
forall x. Rep InteractionPoint x -> InteractionPoint
forall x. InteractionPoint -> Rep InteractionPoint x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. InteractionPoint -> Rep InteractionPoint x
from :: forall x. InteractionPoint -> Rep InteractionPoint x
$cto :: forall x. Rep InteractionPoint x -> InteractionPoint
to :: forall x. Rep InteractionPoint x -> InteractionPoint
Generic
instance Eq InteractionPoint where == :: InteractionPoint -> InteractionPoint -> Bool
(==) = Maybe MetaId -> Maybe MetaId -> Bool
forall a. Eq a => a -> a -> Bool
(==) (Maybe MetaId -> Maybe MetaId -> Bool)
-> (InteractionPoint -> Maybe MetaId)
-> InteractionPoint
-> InteractionPoint
-> Bool
forall b c a. (b -> b -> c) -> (a -> b) -> a -> a -> c
`on` InteractionPoint -> Maybe MetaId
ipMeta
instance HasTag InteractionPoint where
type Tag InteractionPoint = MetaId
tag :: InteractionPoint -> Maybe (Tag InteractionPoint)
tag = InteractionPoint -> Maybe (Tag InteractionPoint)
InteractionPoint -> Maybe MetaId
ipMeta
type InteractionPoints = BiMap InteractionId InteractionPoint
data Overapplied = Overapplied | NotOverapplied
deriving (Overapplied -> Overapplied -> Bool
(Overapplied -> Overapplied -> Bool)
-> (Overapplied -> Overapplied -> Bool) -> Eq Overapplied
forall a. (a -> a -> Bool) -> (a -> a -> Bool) -> Eq a
$c== :: Overapplied -> Overapplied -> Bool
== :: Overapplied -> Overapplied -> Bool
$c/= :: Overapplied -> Overapplied -> Bool
/= :: Overapplied -> Overapplied -> Bool
Eq, Int -> Overapplied -> ShowS
[Overapplied] -> ShowS
Overapplied -> String
(Int -> Overapplied -> ShowS)
-> (Overapplied -> String)
-> ([Overapplied] -> ShowS)
-> Show Overapplied
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> Overapplied -> ShowS
showsPrec :: Int -> Overapplied -> ShowS
$cshow :: Overapplied -> String
show :: Overapplied -> String
$cshowList :: [Overapplied] -> ShowS
showList :: [Overapplied] -> ShowS
Show, (forall x. Overapplied -> Rep Overapplied x)
-> (forall x. Rep Overapplied x -> Overapplied)
-> Generic Overapplied
forall x. Rep Overapplied x -> Overapplied
forall x. Overapplied -> Rep Overapplied x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. Overapplied -> Rep Overapplied x
from :: forall x. Overapplied -> Rep Overapplied x
$cto :: forall x. Rep Overapplied x -> Overapplied
to :: forall x. Rep Overapplied x -> Overapplied
Generic)
newtype IPBoundary' t = IPBoundary
{ forall t. IPBoundary' t -> Map (IntMap Bool) t
getBoundary :: Map (IntMap Bool) t
}
deriving (Int -> IPBoundary' t -> ShowS
[IPBoundary' t] -> ShowS
IPBoundary' t -> String
(Int -> IPBoundary' t -> ShowS)
-> (IPBoundary' t -> String)
-> ([IPBoundary' t] -> ShowS)
-> Show (IPBoundary' t)
forall t. Show t => Int -> IPBoundary' t -> ShowS
forall t. Show t => [IPBoundary' t] -> ShowS
forall t. Show t => IPBoundary' t -> String
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: forall t. Show t => Int -> IPBoundary' t -> ShowS
showsPrec :: Int -> IPBoundary' t -> ShowS
$cshow :: forall t. Show t => IPBoundary' t -> String
show :: IPBoundary' t -> String
$cshowList :: forall t. Show t => [IPBoundary' t] -> ShowS
showList :: [IPBoundary' t] -> ShowS
Show, (forall a b. (a -> b) -> IPBoundary' a -> IPBoundary' b)
-> (forall a b. a -> IPBoundary' b -> IPBoundary' a)
-> Functor IPBoundary'
forall a b. a -> IPBoundary' b -> IPBoundary' a
forall a b. (a -> b) -> IPBoundary' a -> IPBoundary' b
forall (f :: * -> *).
(forall a b. (a -> b) -> f a -> f b)
-> (forall a b. a -> f b -> f a) -> Functor f
$cfmap :: forall a b. (a -> b) -> IPBoundary' a -> IPBoundary' b
fmap :: forall a b. (a -> b) -> IPBoundary' a -> IPBoundary' b
$c<$ :: forall a b. a -> IPBoundary' b -> IPBoundary' a
<$ :: forall a b. a -> IPBoundary' b -> IPBoundary' a
Functor, (forall m. Monoid m => IPBoundary' m -> m)
-> (forall m a. Monoid m => (a -> m) -> IPBoundary' a -> m)
-> (forall m a. Monoid m => (a -> m) -> IPBoundary' a -> m)
-> (forall a b. (a -> b -> b) -> b -> IPBoundary' a -> b)
-> (forall a b. (a -> b -> b) -> b -> IPBoundary' a -> b)
-> (forall b a. (b -> a -> b) -> b -> IPBoundary' a -> b)
-> (forall b a. (b -> a -> b) -> b -> IPBoundary' a -> b)
-> (forall a. (a -> a -> a) -> IPBoundary' a -> a)
-> (forall a. (a -> a -> a) -> IPBoundary' a -> a)
-> (forall a. IPBoundary' a -> [a])
-> (forall a. IPBoundary' a -> Bool)
-> (forall a. IPBoundary' a -> Int)
-> (forall a. Eq a => a -> IPBoundary' a -> Bool)
-> (forall a. Ord a => IPBoundary' a -> a)
-> (forall a. Ord a => IPBoundary' a -> a)
-> (forall a. Num a => IPBoundary' a -> a)
-> (forall a. Num a => IPBoundary' a -> a)
-> Foldable IPBoundary'
forall a. Eq a => a -> IPBoundary' a -> Bool
forall a. Num a => IPBoundary' a -> a
forall a. Ord a => IPBoundary' a -> a
forall m. Monoid m => IPBoundary' m -> m
forall a. IPBoundary' a -> Bool
forall a. IPBoundary' a -> Int
forall a. IPBoundary' a -> [a]
forall a. (a -> a -> a) -> IPBoundary' a -> a
forall m a. Monoid m => (a -> m) -> IPBoundary' a -> m
forall b a. (b -> a -> b) -> b -> IPBoundary' a -> b
forall a b. (a -> b -> b) -> b -> IPBoundary' a -> b
forall (t :: * -> *).
(forall m. Monoid m => t m -> m)
-> (forall m a. Monoid m => (a -> m) -> t a -> m)
-> (forall m a. Monoid m => (a -> m) -> t a -> m)
-> (forall a b. (a -> b -> b) -> b -> t a -> b)
-> (forall a b. (a -> b -> b) -> b -> t a -> b)
-> (forall b a. (b -> a -> b) -> b -> t a -> b)
-> (forall b a. (b -> a -> b) -> b -> t a -> b)
-> (forall a. (a -> a -> a) -> t a -> a)
-> (forall a. (a -> a -> a) -> t a -> a)
-> (forall a. t a -> [a])
-> (forall a. t a -> Bool)
-> (forall a. t a -> Int)
-> (forall a. Eq a => a -> t a -> Bool)
-> (forall a. Ord a => t a -> a)
-> (forall a. Ord a => t a -> a)
-> (forall a. Num a => t a -> a)
-> (forall a. Num a => t a -> a)
-> Foldable t
$cfold :: forall m. Monoid m => IPBoundary' m -> m
fold :: forall m. Monoid m => IPBoundary' m -> m
$cfoldMap :: forall m a. Monoid m => (a -> m) -> IPBoundary' a -> m
foldMap :: forall m a. Monoid m => (a -> m) -> IPBoundary' a -> m
$cfoldMap' :: forall m a. Monoid m => (a -> m) -> IPBoundary' a -> m
foldMap' :: forall m a. Monoid m => (a -> m) -> IPBoundary' a -> m
$cfoldr :: forall a b. (a -> b -> b) -> b -> IPBoundary' a -> b
foldr :: forall a b. (a -> b -> b) -> b -> IPBoundary' a -> b
$cfoldr' :: forall a b. (a -> b -> b) -> b -> IPBoundary' a -> b
foldr' :: forall a b. (a -> b -> b) -> b -> IPBoundary' a -> b
$cfoldl :: forall b a. (b -> a -> b) -> b -> IPBoundary' a -> b
foldl :: forall b a. (b -> a -> b) -> b -> IPBoundary' a -> b
$cfoldl' :: forall b a. (b -> a -> b) -> b -> IPBoundary' a -> b
foldl' :: forall b a. (b -> a -> b) -> b -> IPBoundary' a -> b
$cfoldr1 :: forall a. (a -> a -> a) -> IPBoundary' a -> a
foldr1 :: forall a. (a -> a -> a) -> IPBoundary' a -> a
$cfoldl1 :: forall a. (a -> a -> a) -> IPBoundary' a -> a
foldl1 :: forall a. (a -> a -> a) -> IPBoundary' a -> a
$ctoList :: forall a. IPBoundary' a -> [a]
toList :: forall a. IPBoundary' a -> [a]
$cnull :: forall a. IPBoundary' a -> Bool
null :: forall a. IPBoundary' a -> Bool
$clength :: forall a. IPBoundary' a -> Int
length :: forall a. IPBoundary' a -> Int
$celem :: forall a. Eq a => a -> IPBoundary' a -> Bool
elem :: forall a. Eq a => a -> IPBoundary' a -> Bool
$cmaximum :: forall a. Ord a => IPBoundary' a -> a
maximum :: forall a. Ord a => IPBoundary' a -> a
$cminimum :: forall a. Ord a => IPBoundary' a -> a
minimum :: forall a. Ord a => IPBoundary' a -> a
$csum :: forall a. Num a => IPBoundary' a -> a
sum :: forall a. Num a => IPBoundary' a -> a
$cproduct :: forall a. Num a => IPBoundary' a -> a
product :: forall a. Num a => IPBoundary' a -> a
Foldable, Functor IPBoundary'
Foldable IPBoundary'
(Functor IPBoundary', Foldable IPBoundary') =>
(forall (f :: * -> *) a b.
Applicative f =>
(a -> f b) -> IPBoundary' a -> f (IPBoundary' b))
-> (forall (f :: * -> *) a.
Applicative f =>
IPBoundary' (f a) -> f (IPBoundary' a))
-> (forall (m :: * -> *) a b.
Monad m =>
(a -> m b) -> IPBoundary' a -> m (IPBoundary' b))
-> (forall (m :: * -> *) a.
Monad m =>
IPBoundary' (m a) -> m (IPBoundary' a))
-> Traversable IPBoundary'
forall (t :: * -> *).
(Functor t, Foldable t) =>
(forall (f :: * -> *) a b.
Applicative f =>
(a -> f b) -> t a -> f (t b))
-> (forall (f :: * -> *) a. Applicative f => t (f a) -> f (t a))
-> (forall (m :: * -> *) a b.
Monad m =>
(a -> m b) -> t a -> m (t b))
-> (forall (m :: * -> *) a. Monad m => t (m a) -> m (t a))
-> Traversable t
forall (m :: * -> *) a.
Monad m =>
IPBoundary' (m a) -> m (IPBoundary' a)
forall (f :: * -> *) a.
Applicative f =>
IPBoundary' (f a) -> f (IPBoundary' a)
forall (m :: * -> *) a b.
Monad m =>
(a -> m b) -> IPBoundary' a -> m (IPBoundary' b)
forall (f :: * -> *) a b.
Applicative f =>
(a -> f b) -> IPBoundary' a -> f (IPBoundary' b)
$ctraverse :: forall (f :: * -> *) a b.
Applicative f =>
(a -> f b) -> IPBoundary' a -> f (IPBoundary' b)
traverse :: forall (f :: * -> *) a b.
Applicative f =>
(a -> f b) -> IPBoundary' a -> f (IPBoundary' b)
$csequenceA :: forall (f :: * -> *) a.
Applicative f =>
IPBoundary' (f a) -> f (IPBoundary' a)
sequenceA :: forall (f :: * -> *) a.
Applicative f =>
IPBoundary' (f a) -> f (IPBoundary' a)
$cmapM :: forall (m :: * -> *) a b.
Monad m =>
(a -> m b) -> IPBoundary' a -> m (IPBoundary' b)
mapM :: forall (m :: * -> *) a b.
Monad m =>
(a -> m b) -> IPBoundary' a -> m (IPBoundary' b)
$csequence :: forall (m :: * -> *) a.
Monad m =>
IPBoundary' (m a) -> m (IPBoundary' a)
sequence :: forall (m :: * -> *) a.
Monad m =>
IPBoundary' (m a) -> m (IPBoundary' a)
Traversable, (forall x. IPBoundary' t -> Rep (IPBoundary' t) x)
-> (forall x. Rep (IPBoundary' t) x -> IPBoundary' t)
-> Generic (IPBoundary' t)
forall x. Rep (IPBoundary' t) x -> IPBoundary' t
forall x. IPBoundary' t -> Rep (IPBoundary' t) x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
forall t x. Rep (IPBoundary' t) x -> IPBoundary' t
forall t x. IPBoundary' t -> Rep (IPBoundary' t) x
$cfrom :: forall t x. IPBoundary' t -> Rep (IPBoundary' t) x
from :: forall x. IPBoundary' t -> Rep (IPBoundary' t) x
$cto :: forall t x. Rep (IPBoundary' t) x -> IPBoundary' t
to :: forall x. Rep (IPBoundary' t) x -> IPBoundary' t
Generic)
type IPBoundary = IPBoundary' Term
data IPClause = IPClause
{ IPClause -> QName
ipcQName :: QName
, IPClause -> Int
ipcClauseNo :: Int
, IPClause -> Type
ipcType :: Type
, IPClause -> Maybe Substitution
ipcWithSub :: Maybe Substitution
, IPClause -> SpineClause
ipcClause :: A.SpineClause
, IPClause -> Closure ()
ipcClosure :: Closure ()
}
| IPNoClause
deriving ((forall x. IPClause -> Rep IPClause x)
-> (forall x. Rep IPClause x -> IPClause) -> Generic IPClause
forall x. Rep IPClause x -> IPClause
forall x. IPClause -> Rep IPClause x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. IPClause -> Rep IPClause x
from :: forall x. IPClause -> Rep IPClause x
$cto :: forall x. Rep IPClause x -> IPClause
to :: forall x. Rep IPClause x -> IPClause
Generic)
instance Eq IPClause where
IPClause
IPNoClause == :: IPClause -> IPClause -> Bool
== IPClause
IPNoClause = Bool
True
IPClause QName
x Int
i Type
_ Maybe Substitution
_ SpineClause
_ Closure ()
_ == IPClause QName
x' Int
i' Type
_ Maybe Substitution
_ SpineClause
_ Closure ()
_ = QName
x QName -> QName -> Bool
forall a. Eq a => a -> a -> Bool
== QName
x' Bool -> Bool -> Bool
&& Int
i Int -> Int -> Bool
forall a. Eq a => a -> a -> Bool
== Int
i'
IPClause
_ == IPClause
_ = Bool
False
data Signature = Sig
{ Signature -> Sections
_sigSections :: Sections
, Signature -> Definitions
_sigDefinitions :: Definitions
, Signature -> RewriteRuleMap
_sigRewriteRules :: RewriteRuleMap
, Signature -> InstanceTable
_sigInstances :: InstanceTable
}
deriving (Int -> Signature -> ShowS
[Signature] -> ShowS
Signature -> String
(Int -> Signature -> ShowS)
-> (Signature -> String)
-> ([Signature] -> ShowS)
-> Show Signature
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> Signature -> ShowS
showsPrec :: Int -> Signature -> ShowS
$cshow :: Signature -> String
show :: Signature -> String
$cshowList :: [Signature] -> ShowS
showList :: [Signature] -> ShowS
Show, (forall x. Signature -> Rep Signature x)
-> (forall x. Rep Signature x -> Signature) -> Generic Signature
forall x. Rep Signature x -> Signature
forall x. Signature -> Rep Signature x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. Signature -> Rep Signature x
from :: forall x. Signature -> Rep Signature x
$cto :: forall x. Rep Signature x -> Signature
to :: forall x. Rep Signature x -> Signature
Generic)
sigSections :: Lens' Signature Sections
sigSections :: Lens' Signature Sections
sigSections Sections -> f Sections
f Signature
s =
Sections -> f Sections
f (Signature -> Sections
_sigSections Signature
s) f Sections -> (Sections -> Signature) -> f Signature
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&>
\Sections
x -> Signature
s {_sigSections = x}
sigDefinitions :: Lens' Signature Definitions
sigDefinitions :: Lens' Signature Definitions
sigDefinitions Definitions -> f Definitions
f Signature
s =
Definitions -> f Definitions
f (Signature -> Definitions
_sigDefinitions Signature
s) f Definitions -> (Definitions -> Signature) -> f Signature
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&>
\Definitions
x -> Signature
s {_sigDefinitions = x}
sigInstances :: Lens' Signature InstanceTable
sigInstances :: Lens' Signature InstanceTable
sigInstances InstanceTable -> f InstanceTable
f Signature
s = InstanceTable -> f InstanceTable
f (Signature -> InstanceTable
_sigInstances Signature
s) f InstanceTable -> (InstanceTable -> Signature) -> f Signature
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \InstanceTable
x -> Signature
s {_sigInstances = x}
sigRewriteRules :: Lens' Signature RewriteRuleMap
sigRewriteRules :: Lens' Signature RewriteRuleMap
sigRewriteRules RewriteRuleMap -> f RewriteRuleMap
f Signature
s =
RewriteRuleMap -> f RewriteRuleMap
f (Signature -> RewriteRuleMap
_sigRewriteRules Signature
s) f RewriteRuleMap -> (RewriteRuleMap -> Signature) -> f Signature
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&>
\RewriteRuleMap
x -> Signature
s {_sigRewriteRules = x}
type Sections = Map ModuleName Section
type Definitions = HashMap QName Definition
type RewriteRuleMap = HashMap QName RewriteRules
type DisplayForms = HashMap QName [LocalDisplayForm]
#if __GLASGOW_HASKELL__ >= 900
{-# SPECIALIZE HMap.insert :: QName -> v -> HashMap QName v -> HashMap QName v #-}
#endif
{-# SPECIALIZE HMap.lookup :: QName -> HashMap QName v -> Maybe v #-}
newtype Section = Section { Section -> Telescope
_secTelescope :: Telescope }
deriving (Int -> Section -> ShowS
[Section] -> ShowS
Section -> String
(Int -> Section -> ShowS)
-> (Section -> String) -> ([Section] -> ShowS) -> Show Section
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> Section -> ShowS
showsPrec :: Int -> Section -> ShowS
$cshow :: Section -> String
show :: Section -> String
$cshowList :: [Section] -> ShowS
showList :: [Section] -> ShowS
Show, Section -> ()
(Section -> ()) -> NFData Section
forall a. (a -> ()) -> NFData a
$crnf :: Section -> ()
rnf :: Section -> ()
NFData)
instance Pretty Section where
pretty :: Section -> Doc
pretty = Telescope -> Doc
forall a. Pretty a => a -> Doc
pretty (Telescope -> Doc) -> (Section -> Telescope) -> Section -> Doc
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Section -> Telescope
_secTelescope
secTelescope :: Lens' Section Telescope
secTelescope :: Lens' Section Telescope
secTelescope Telescope -> f Telescope
f Section
s =
Telescope -> f Telescope
f (Section -> Telescope
_secTelescope Section
s) f Telescope -> (Telescope -> Section) -> f Section
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&>
\Telescope
x -> Section
s {_secTelescope = x}
emptySignature :: Signature
emptySignature :: Signature
emptySignature = Sections
-> Definitions -> RewriteRuleMap -> InstanceTable -> Signature
Sig Sections
forall k a. Map k a
Map.empty Definitions
forall k v. HashMap k v
HMap.empty RewriteRuleMap
forall k v. HashMap k v
HMap.empty InstanceTable
forall a. Monoid a => a
mempty
data DisplayForm = Display
{ DisplayForm -> Int
dfPatternVars :: Nat
, DisplayForm -> [Elim]
dfPats :: Elims
, DisplayForm -> DisplayTerm
dfRHS :: DisplayTerm
}
deriving (Int -> DisplayForm -> ShowS
[DisplayForm] -> ShowS
DisplayForm -> String
(Int -> DisplayForm -> ShowS)
-> (DisplayForm -> String)
-> ([DisplayForm] -> ShowS)
-> Show DisplayForm
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> DisplayForm -> ShowS
showsPrec :: Int -> DisplayForm -> ShowS
$cshow :: DisplayForm -> String
show :: DisplayForm -> String
$cshowList :: [DisplayForm] -> ShowS
showList :: [DisplayForm] -> ShowS
Show, (forall x. DisplayForm -> Rep DisplayForm x)
-> (forall x. Rep DisplayForm x -> DisplayForm)
-> Generic DisplayForm
forall x. Rep DisplayForm x -> DisplayForm
forall x. DisplayForm -> Rep DisplayForm x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. DisplayForm -> Rep DisplayForm x
from :: forall x. DisplayForm -> Rep DisplayForm x
$cto :: forall x. Rep DisplayForm x -> DisplayForm
to :: forall x. Rep DisplayForm x -> DisplayForm
Generic)
type LocalDisplayForm = Open DisplayForm
data DisplayTerm
= DWithApp DisplayTerm (List1 DisplayTerm) Elims
| DCon ConHead ConInfo [Arg DisplayTerm]
| DDef QName [Elim' DisplayTerm]
| DDot' Term Elims
| DTerm' Term Elims
deriving (Int -> DisplayTerm -> ShowS
[DisplayTerm] -> ShowS
DisplayTerm -> String
(Int -> DisplayTerm -> ShowS)
-> (DisplayTerm -> String)
-> ([DisplayTerm] -> ShowS)
-> Show DisplayTerm
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> DisplayTerm -> ShowS
showsPrec :: Int -> DisplayTerm -> ShowS
$cshow :: DisplayTerm -> String
show :: DisplayTerm -> String
$cshowList :: [DisplayTerm] -> ShowS
showList :: [DisplayTerm] -> ShowS
Show, (forall x. DisplayTerm -> Rep DisplayTerm x)
-> (forall x. Rep DisplayTerm x -> DisplayTerm)
-> Generic DisplayTerm
forall x. Rep DisplayTerm x -> DisplayTerm
forall x. DisplayTerm -> Rep DisplayTerm x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. DisplayTerm -> Rep DisplayTerm x
from :: forall x. DisplayTerm -> Rep DisplayTerm x
$cto :: forall x. Rep DisplayTerm x -> DisplayTerm
to :: forall x. Rep DisplayTerm x -> DisplayTerm
Generic)
pattern DDot :: Term -> DisplayTerm
pattern $mDDot :: forall {r}. DisplayTerm -> (Term -> r) -> ((# #) -> r) -> r
$bDDot :: Term -> DisplayTerm
DDot v = DDot' v []
pattern DTerm :: Term -> DisplayTerm
pattern $mDTerm :: forall {r}. DisplayTerm -> (Term -> r) -> ((# #) -> r) -> r
$bDTerm :: Term -> DisplayTerm
DTerm v = DTerm' v []
instance Free DisplayForm where
freeVars' :: forall a c. IsVarSet a c => DisplayForm -> FreeM a c
freeVars' (Display Int
n [Elim]
ps DisplayTerm
t) = FreeM a c -> FreeM a c
forall a b c (m :: * -> *) z.
MonadReader (FreeEnv' a b c) m =>
m z -> m z
underBinder ([Elim] -> FreeM a c
forall t a c. (Free t, IsVarSet a c) => t -> FreeM a c
forall a c. IsVarSet a c => [Elim] -> FreeM a c
freeVars' [Elim]
ps) FreeM a c -> FreeM a c -> FreeM a c
forall a. Monoid a => a -> a -> a
`mappend` Int -> FreeM a c -> FreeM a c
forall a b c (m :: * -> *) z.
MonadReader (FreeEnv' a b c) m =>
Int -> m z -> m z
underBinder' Int
n (DisplayTerm -> FreeM a c
forall t a c. (Free t, IsVarSet a c) => t -> FreeM a c
forall a c. IsVarSet a c => DisplayTerm -> FreeM a c
freeVars' DisplayTerm
t)
instance Free DisplayTerm where
freeVars' :: forall a c. IsVarSet a c => DisplayTerm -> FreeM a c
freeVars' (DWithApp DisplayTerm
t List1 DisplayTerm
ws [Elim]
es) = (DisplayTerm, (List1 DisplayTerm, [Elim])) -> FreeM a c
forall t a c. (Free t, IsVarSet a c) => t -> FreeM a c
forall a c.
IsVarSet a c =>
(DisplayTerm, (List1 DisplayTerm, [Elim])) -> FreeM a c
freeVars' (DisplayTerm
t, (List1 DisplayTerm
ws, [Elim]
es))
freeVars' (DCon ConHead
_ ConInfo
_ [Arg DisplayTerm]
vs) = [Arg DisplayTerm] -> FreeM a c
forall t a c. (Free t, IsVarSet a c) => t -> FreeM a c
forall a c. IsVarSet a c => [Arg DisplayTerm] -> FreeM a c
freeVars' [Arg DisplayTerm]
vs
freeVars' (DDef QName
_ [Elim' DisplayTerm]
es) = [Elim' DisplayTerm] -> FreeM a c
forall t a c. (Free t, IsVarSet a c) => t -> FreeM a c
forall a c. IsVarSet a c => [Elim' DisplayTerm] -> FreeM a c
freeVars' [Elim' DisplayTerm]
es
freeVars' (DDot' Term
v [Elim]
es) = (Term, [Elim]) -> FreeM a c
forall t a c. (Free t, IsVarSet a c) => t -> FreeM a c
forall a c. IsVarSet a c => (Term, [Elim]) -> FreeM a c
freeVars' (Term
v, [Elim]
es)
freeVars' (DTerm' Term
v [Elim]
es) = (Term, [Elim]) -> FreeM a c
forall t a c. (Free t, IsVarSet a c) => t -> FreeM a c
forall a c. IsVarSet a c => (Term, [Elim]) -> FreeM a c
freeVars' (Term
v, [Elim]
es)
instance Pretty DisplayTerm where
prettyPrec :: Int -> DisplayTerm -> Doc
prettyPrec Int
p DisplayTerm
v =
case DisplayTerm
v of
DTerm Term
v -> Int -> Term -> Doc
forall a. Pretty a => Int -> a -> Doc
prettyPrec Int
p Term
v
DTerm' Term
v [Elim]
es -> Int -> Term -> Doc
forall a. Pretty a => Int -> a -> Doc
prettyPrec Int
9 Term
v Doc -> [Elim] -> Doc
forall el. Pretty el => Doc -> [el] -> Doc
`pApp` [Elim]
es
DDot Term
v -> Doc
"." Doc -> Doc -> Doc
forall a. Semigroup a => a -> a -> a
<> Int -> Term -> Doc
forall a. Pretty a => Int -> a -> Doc
prettyPrec Int
10 Term
v
DDot' Term
v [Elim]
es -> Doc
"." Doc -> Doc -> Doc
forall a. Semigroup a => a -> a -> a
<> Doc -> Doc
parens (Int -> Term -> Doc
forall a. Pretty a => Int -> a -> Doc
prettyPrec Int
9 Term
v Doc -> [Elim] -> Doc
forall el. Pretty el => Doc -> [el] -> Doc
`pAp` [Elim]
es)
DDef QName
f [Elim' DisplayTerm]
es -> QName -> Doc
forall a. Pretty a => a -> Doc
pretty QName
f Doc -> [Elim' DisplayTerm] -> Doc
forall el. Pretty el => Doc -> [el] -> Doc
`pApp` [Elim' DisplayTerm]
es
DCon ConHead
c ConInfo
_ [Arg DisplayTerm]
vs -> QName -> Doc
forall a. Pretty a => a -> Doc
pretty (ConHead -> QName
conName ConHead
c) Doc -> [Elim' DisplayTerm] -> Doc
forall el. Pretty el => Doc -> [el] -> Doc
`pApp` (Arg DisplayTerm -> Elim' DisplayTerm)
-> [Arg DisplayTerm] -> [Elim' DisplayTerm]
forall a b. (a -> b) -> [a] -> [b]
map Arg DisplayTerm -> Elim' DisplayTerm
forall a. Arg a -> Elim' a
Apply [Arg DisplayTerm]
vs
DWithApp DisplayTerm
h List1 DisplayTerm
ws [Elim]
es ->
Bool -> Doc -> Doc
mparens (Int
p Int -> Int -> Bool
forall a. Ord a => a -> a -> Bool
> Int
0)
([Doc] -> Doc
forall (t :: * -> *). Foldable t => t Doc -> Doc
sep [ DisplayTerm -> Doc
forall a. Pretty a => a -> Doc
pretty DisplayTerm
h
, Int -> Doc -> Doc
forall a. Int -> Doc a -> Doc a
nest Int
2 (Doc -> Doc) -> Doc -> Doc
forall a b. (a -> b) -> a -> b
$ NonEmpty Doc -> Doc
forall (t :: * -> *). Foldable t => t Doc -> Doc
fsep (NonEmpty Doc -> Doc) -> NonEmpty Doc -> Doc
forall a b. (a -> b) -> a -> b
$ (DisplayTerm -> Doc) -> List1 DisplayTerm -> NonEmpty Doc
forall a b. (a -> b) -> NonEmpty a -> NonEmpty b
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
fmap (\ DisplayTerm
w -> Doc
"|" Doc -> Doc -> Doc
forall a. Doc a -> Doc a -> Doc a
<+> DisplayTerm -> Doc
forall a. Pretty a => a -> Doc
pretty DisplayTerm
w) List1 DisplayTerm
ws
])
Doc -> [Elim] -> Doc
forall el. Pretty el => Doc -> [el] -> Doc
`pApp` [Elim]
es
where
pApp :: Pretty el => Doc -> [el] -> Doc
pApp :: forall el. Pretty el => Doc -> [el] -> Doc
pApp Doc
d [el]
els = Bool -> Doc -> Doc
mparens (Bool -> Bool
not ([el] -> Bool
forall a. Null a => a -> Bool
null [el]
els) Bool -> Bool -> Bool
&& Int
p Int -> Int -> Bool
forall a. Ord a => a -> a -> Bool
> Int
9) (Doc -> Doc) -> Doc -> Doc
forall a b. (a -> b) -> a -> b
$ Doc -> [el] -> Doc
forall el. Pretty el => Doc -> [el] -> Doc
pAp Doc
d [el]
els
pAp :: Pretty el => Doc -> [el] -> Doc
pAp :: forall el. Pretty el => Doc -> [el] -> Doc
pAp Doc
d [el]
els = [Doc] -> Doc
forall (t :: * -> *). Foldable t => t Doc -> Doc
sep [Doc
d, Int -> Doc -> Doc
forall a. Int -> Doc a -> Doc a
nest Int
2 (Doc -> Doc) -> Doc -> Doc
forall a b. (a -> b) -> a -> b
$ [Doc] -> Doc
forall (t :: * -> *). Foldable t => t Doc -> Doc
fsep ((el -> Doc) -> [el] -> [Doc]
forall a b. (a -> b) -> [a] -> [b]
map (Int -> el -> Doc
forall a. Pretty a => Int -> a -> Doc
prettyPrec Int
10) [el]
els)]
instance Pretty DisplayForm where
prettyPrec :: Int -> DisplayForm -> Doc
prettyPrec Int
p (Display Int
fv [Elim]
lhs DisplayTerm
rhs) = Bool -> Doc -> Doc
mparens (Int
p Int -> Int -> Bool
forall a. Ord a => a -> a -> Bool
> Int
9) (Doc -> Doc) -> Doc -> Doc
forall a b. (a -> b) -> a -> b
$
Doc
"Display" Doc -> Doc -> Doc
<?> [Doc] -> Doc
forall (t :: * -> *). Foldable t => t Doc -> Doc
fsep [ Int -> Doc
forall a. Show a => a -> Doc
pshow Int
fv, Int -> [Elim] -> Doc
forall a. Pretty a => Int -> a -> Doc
prettyPrec Int
10 [Elim]
lhs, Int -> DisplayTerm -> Doc
forall a. Pretty a => Int -> a -> Doc
prettyPrec Int
10 DisplayTerm
rhs ]
defaultDisplayForm :: QName -> [LocalDisplayForm]
defaultDisplayForm :: QName -> [LocalDisplayForm]
defaultDisplayForm QName
c = []
data NLPat
= PVar !Int [Arg Int]
| PDef QName PElims
| PLam ArgInfo (Abs NLPat)
| PPi (Dom NLPType) (Abs NLPType)
| PSort NLPSort
| PBoundVar {-# UNPACK #-} !Int PElims
| PTerm Term
deriving (Int -> NLPat -> ShowS
[NLPat] -> ShowS
NLPat -> String
(Int -> NLPat -> ShowS)
-> (NLPat -> String) -> ([NLPat] -> ShowS) -> Show NLPat
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> NLPat -> ShowS
showsPrec :: Int -> NLPat -> ShowS
$cshow :: NLPat -> String
show :: NLPat -> String
$cshowList :: [NLPat] -> ShowS
showList :: [NLPat] -> ShowS
Show, (forall x. NLPat -> Rep NLPat x)
-> (forall x. Rep NLPat x -> NLPat) -> Generic NLPat
forall x. Rep NLPat x -> NLPat
forall x. NLPat -> Rep NLPat x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. NLPat -> Rep NLPat x
from :: forall x. NLPat -> Rep NLPat x
$cto :: forall x. Rep NLPat x -> NLPat
to :: forall x. Rep NLPat x -> NLPat
Generic)
type PElims = [Elim' NLPat]
type instance TypeOf NLPat = Type
type instance TypeOf [Elim' NLPat] = (Type, Elims -> Term)
instance TermLike NLPat where
traverseTermM :: forall (m :: * -> *).
Monad m =>
(Term -> m Term) -> NLPat -> m NLPat
traverseTermM Term -> m Term
f = \case
p :: NLPat
p@PVar{} -> NLPat -> m NLPat
forall a. a -> m a
forall (m :: * -> *) a. Monad m => a -> m a
return NLPat
p
PDef QName
d PElims
ps -> QName -> PElims -> NLPat
PDef QName
d (PElims -> NLPat) -> m PElims -> m NLPat
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> (Term -> m Term) -> PElims -> m PElims
forall a (m :: * -> *).
(TermLike a, Monad m) =>
(Term -> m Term) -> a -> m a
forall (m :: * -> *).
Monad m =>
(Term -> m Term) -> PElims -> m PElims
traverseTermM Term -> m Term
f PElims
ps
PLam ArgInfo
i Abs NLPat
p -> ArgInfo -> Abs NLPat -> NLPat
PLam ArgInfo
i (Abs NLPat -> NLPat) -> m (Abs NLPat) -> m NLPat
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> (Term -> m Term) -> Abs NLPat -> m (Abs NLPat)
forall a (m :: * -> *).
(TermLike a, Monad m) =>
(Term -> m Term) -> a -> m a
forall (m :: * -> *).
Monad m =>
(Term -> m Term) -> Abs NLPat -> m (Abs NLPat)
traverseTermM Term -> m Term
f Abs NLPat
p
PPi Dom NLPType
a Abs NLPType
b -> Dom NLPType -> Abs NLPType -> NLPat
PPi (Dom NLPType -> Abs NLPType -> NLPat)
-> m (Dom NLPType) -> m (Abs NLPType -> NLPat)
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> (Term -> m Term) -> Dom NLPType -> m (Dom NLPType)
forall a (m :: * -> *).
(TermLike a, Monad m) =>
(Term -> m Term) -> a -> m a
forall (m :: * -> *).
Monad m =>
(Term -> m Term) -> Dom NLPType -> m (Dom NLPType)
traverseTermM Term -> m Term
f Dom NLPType
a m (Abs NLPType -> NLPat) -> m (Abs NLPType) -> m NLPat
forall a b. m (a -> b) -> m a -> m b
forall (f :: * -> *) a b. Applicative f => f (a -> b) -> f a -> f b
<*> (Term -> m Term) -> Abs NLPType -> m (Abs NLPType)
forall a (m :: * -> *).
(TermLike a, Monad m) =>
(Term -> m Term) -> a -> m a
forall (m :: * -> *).
Monad m =>
(Term -> m Term) -> Abs NLPType -> m (Abs NLPType)
traverseTermM Term -> m Term
f Abs NLPType
b
PSort NLPSort
s -> NLPSort -> NLPat
PSort (NLPSort -> NLPat) -> m NLPSort -> m NLPat
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> (Term -> m Term) -> NLPSort -> m NLPSort
forall a (m :: * -> *).
(TermLike a, Monad m) =>
(Term -> m Term) -> a -> m a
forall (m :: * -> *).
Monad m =>
(Term -> m Term) -> NLPSort -> m NLPSort
traverseTermM Term -> m Term
f NLPSort
s
PBoundVar Int
i PElims
ps -> Int -> PElims -> NLPat
PBoundVar Int
i (PElims -> NLPat) -> m PElims -> m NLPat
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> (Term -> m Term) -> PElims -> m PElims
forall a (m :: * -> *).
(TermLike a, Monad m) =>
(Term -> m Term) -> a -> m a
forall (m :: * -> *).
Monad m =>
(Term -> m Term) -> PElims -> m PElims
traverseTermM Term -> m Term
f PElims
ps
PTerm Term
t -> Term -> NLPat
PTerm (Term -> NLPat) -> m Term -> m NLPat
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> Term -> m Term
f Term
t
foldTerm :: forall m. Monoid m => (Term -> m) -> NLPat -> m
foldTerm Term -> m
f NLPat
t = case NLPat
t of
PVar{} -> m
forall a. Monoid a => a
mempty
PDef QName
d PElims
ps -> (Term -> m) -> PElims -> m
forall m. Monoid m => (Term -> m) -> PElims -> m
forall a m. (TermLike a, Monoid m) => (Term -> m) -> a -> m
foldTerm Term -> m
f PElims
ps
PLam ArgInfo
i Abs NLPat
p -> (Term -> m) -> Abs NLPat -> m
forall m. Monoid m => (Term -> m) -> Abs NLPat -> m
forall a m. (TermLike a, Monoid m) => (Term -> m) -> a -> m
foldTerm Term -> m
f Abs NLPat
p
PPi Dom NLPType
a Abs NLPType
b -> (Term -> m) -> (Dom NLPType, Abs NLPType) -> m
forall m.
Monoid m =>
(Term -> m) -> (Dom NLPType, Abs NLPType) -> m
forall a m. (TermLike a, Monoid m) => (Term -> m) -> a -> m
foldTerm Term -> m
f (Dom NLPType
a, Abs NLPType
b)
PSort NLPSort
s -> (Term -> m) -> NLPSort -> m
forall m. Monoid m => (Term -> m) -> NLPSort -> m
forall a m. (TermLike a, Monoid m) => (Term -> m) -> a -> m
foldTerm Term -> m
f NLPSort
s
PBoundVar Int
i PElims
ps -> (Term -> m) -> PElims -> m
forall m. Monoid m => (Term -> m) -> PElims -> m
forall a m. (TermLike a, Monoid m) => (Term -> m) -> a -> m
foldTerm Term -> m
f PElims
ps
PTerm Term
t -> (Term -> m) -> Term -> m
forall m. Monoid m => (Term -> m) -> Term -> m
forall a m. (TermLike a, Monoid m) => (Term -> m) -> a -> m
foldTerm Term -> m
f Term
t
instance AllMetas NLPat
data NLPType = NLPType
{ NLPType -> NLPSort
nlpTypeSort :: NLPSort
, NLPType -> NLPat
nlpTypeUnEl :: NLPat
} deriving (Int -> NLPType -> ShowS
[NLPType] -> ShowS
NLPType -> String
(Int -> NLPType -> ShowS)
-> (NLPType -> String) -> ([NLPType] -> ShowS) -> Show NLPType
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> NLPType -> ShowS
showsPrec :: Int -> NLPType -> ShowS
$cshow :: NLPType -> String
show :: NLPType -> String
$cshowList :: [NLPType] -> ShowS
showList :: [NLPType] -> ShowS
Show, (forall x. NLPType -> Rep NLPType x)
-> (forall x. Rep NLPType x -> NLPType) -> Generic NLPType
forall x. Rep NLPType x -> NLPType
forall x. NLPType -> Rep NLPType x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. NLPType -> Rep NLPType x
from :: forall x. NLPType -> Rep NLPType x
$cto :: forall x. Rep NLPType x -> NLPType
to :: forall x. Rep NLPType x -> NLPType
Generic)
instance TermLike NLPType where
traverseTermM :: forall (m :: * -> *).
Monad m =>
(Term -> m Term) -> NLPType -> m NLPType
traverseTermM Term -> m Term
f (NLPType NLPSort
s NLPat
t) = NLPSort -> NLPat -> NLPType
NLPType (NLPSort -> NLPat -> NLPType) -> m NLPSort -> m (NLPat -> NLPType)
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> (Term -> m Term) -> NLPSort -> m NLPSort
forall a (m :: * -> *).
(TermLike a, Monad m) =>
(Term -> m Term) -> a -> m a
forall (m :: * -> *).
Monad m =>
(Term -> m Term) -> NLPSort -> m NLPSort
traverseTermM Term -> m Term
f NLPSort
s m (NLPat -> NLPType) -> m NLPat -> m NLPType
forall a b. m (a -> b) -> m a -> m b
forall (f :: * -> *) a b. Applicative f => f (a -> b) -> f a -> f b
<*> (Term -> m Term) -> NLPat -> m NLPat
forall a (m :: * -> *).
(TermLike a, Monad m) =>
(Term -> m Term) -> a -> m a
forall (m :: * -> *).
Monad m =>
(Term -> m Term) -> NLPat -> m NLPat
traverseTermM Term -> m Term
f NLPat
t
foldTerm :: forall m. Monoid m => (Term -> m) -> NLPType -> m
foldTerm Term -> m
f (NLPType NLPSort
s NLPat
t) = (Term -> m) -> (NLPSort, NLPat) -> m
forall m. Monoid m => (Term -> m) -> (NLPSort, NLPat) -> m
forall a m. (TermLike a, Monoid m) => (Term -> m) -> a -> m
foldTerm Term -> m
f (NLPSort
s, NLPat
t)
instance AllMetas NLPType
data NLPSort
= PUniv Univ NLPat
| PInf Univ Integer
| PSizeUniv
| PLockUniv
| PLevelUniv
| PIntervalUniv
deriving (Int -> NLPSort -> ShowS
[NLPSort] -> ShowS
NLPSort -> String
(Int -> NLPSort -> ShowS)
-> (NLPSort -> String) -> ([NLPSort] -> ShowS) -> Show NLPSort
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> NLPSort -> ShowS
showsPrec :: Int -> NLPSort -> ShowS
$cshow :: NLPSort -> String
show :: NLPSort -> String
$cshowList :: [NLPSort] -> ShowS
showList :: [NLPSort] -> ShowS
Show, (forall x. NLPSort -> Rep NLPSort x)
-> (forall x. Rep NLPSort x -> NLPSort) -> Generic NLPSort
forall x. Rep NLPSort x -> NLPSort
forall x. NLPSort -> Rep NLPSort x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. NLPSort -> Rep NLPSort x
from :: forall x. NLPSort -> Rep NLPSort x
$cto :: forall x. Rep NLPSort x -> NLPSort
to :: forall x. Rep NLPSort x -> NLPSort
Generic)
pattern PType, PProp, PSSet :: NLPat -> NLPSort
pattern $mPType :: forall {r}. NLPSort -> (NLPat -> r) -> ((# #) -> r) -> r
$bPType :: NLPat -> NLPSort
PType p = PUniv UType p
pattern $mPProp :: forall {r}. NLPSort -> (NLPat -> r) -> ((# #) -> r) -> r
$bPProp :: NLPat -> NLPSort
PProp p = PUniv UProp p
pattern $mPSSet :: forall {r}. NLPSort -> (NLPat -> r) -> ((# #) -> r) -> r
$bPSSet :: NLPat -> NLPSort
PSSet p = PUniv USSet p
{-# COMPLETE
PType, PSSet, PProp, PInf,
PSizeUniv, PLockUniv, PLevelUniv, PIntervalUniv #-}
instance TermLike NLPSort where
traverseTermM :: forall (m :: * -> *).
Monad m =>
(Term -> m Term) -> NLPSort -> m NLPSort
traverseTermM Term -> m Term
f = \case
PUniv Univ
u NLPat
p -> Univ -> NLPat -> NLPSort
PUniv Univ
u (NLPat -> NLPSort) -> m NLPat -> m NLPSort
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> (Term -> m Term) -> NLPat -> m NLPat
forall a (m :: * -> *).
(TermLike a, Monad m) =>
(Term -> m Term) -> a -> m a
forall (m :: * -> *).
Monad m =>
(Term -> m Term) -> NLPat -> m NLPat
traverseTermM Term -> m Term
f NLPat
p
s :: NLPSort
s@PInf{} -> NLPSort -> m NLPSort
forall a. a -> m a
forall (m :: * -> *) a. Monad m => a -> m a
return NLPSort
s
s :: NLPSort
s@PSizeUniv{} -> NLPSort -> m NLPSort
forall a. a -> m a
forall (m :: * -> *) a. Monad m => a -> m a
return NLPSort
s
s :: NLPSort
s@PLockUniv{} -> NLPSort -> m NLPSort
forall a. a -> m a
forall (m :: * -> *) a. Monad m => a -> m a
return NLPSort
s
s :: NLPSort
s@PLevelUniv{} -> NLPSort -> m NLPSort
forall a. a -> m a
forall (m :: * -> *) a. Monad m => a -> m a
return NLPSort
s
s :: NLPSort
s@PIntervalUniv{} -> NLPSort -> m NLPSort
forall a. a -> m a
forall (m :: * -> *) a. Monad m => a -> m a
return NLPSort
s
foldTerm :: forall m. Monoid m => (Term -> m) -> NLPSort -> m
foldTerm Term -> m
f NLPSort
t = case NLPSort
t of
PUniv Univ
_ NLPat
p -> (Term -> m) -> NLPat -> m
forall m. Monoid m => (Term -> m) -> NLPat -> m
forall a m. (TermLike a, Monoid m) => (Term -> m) -> a -> m
foldTerm Term -> m
f NLPat
p
s :: NLPSort
s@PInf{} -> m
forall a. Monoid a => a
mempty
s :: NLPSort
s@PSizeUniv{} -> m
forall a. Monoid a => a
mempty
s :: NLPSort
s@PLockUniv{} -> m
forall a. Monoid a => a
mempty
s :: NLPSort
s@PLevelUniv{} -> m
forall a. Monoid a => a
mempty
s :: NLPSort
s@PIntervalUniv{} -> m
forall a. Monoid a => a
mempty
instance AllMetas NLPSort
type RewriteRules = [RewriteRule]
data RewriteRule = RewriteRule
{ RewriteRule -> QName
rewName :: QName
, RewriteRule -> Telescope
rewContext :: Telescope
, RewriteRule -> QName
rewHead :: QName
, RewriteRule -> PElims
rewPats :: PElims
, RewriteRule -> Term
rewRHS :: Term
, RewriteRule -> Type
rewType :: Type
, RewriteRule -> Bool
rewFromClause :: Bool
}
deriving (Int -> RewriteRule -> ShowS
[RewriteRule] -> ShowS
RewriteRule -> String
(Int -> RewriteRule -> ShowS)
-> (RewriteRule -> String)
-> ([RewriteRule] -> ShowS)
-> Show RewriteRule
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> RewriteRule -> ShowS
showsPrec :: Int -> RewriteRule -> ShowS
$cshow :: RewriteRule -> String
show :: RewriteRule -> String
$cshowList :: [RewriteRule] -> ShowS
showList :: [RewriteRule] -> ShowS
Show, (forall x. RewriteRule -> Rep RewriteRule x)
-> (forall x. Rep RewriteRule x -> RewriteRule)
-> Generic RewriteRule
forall x. Rep RewriteRule x -> RewriteRule
forall x. RewriteRule -> Rep RewriteRule x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. RewriteRule -> Rep RewriteRule x
from :: forall x. RewriteRule -> Rep RewriteRule x
$cto :: forall x. Rep RewriteRule x -> RewriteRule
to :: forall x. Rep RewriteRule x -> RewriteRule
Generic)
data InstanceInfo = InstanceInfo
{ InstanceInfo -> QName
instanceClass :: QName
, InstanceInfo -> OverlapMode
instanceOverlap :: OverlapMode
}
deriving (Int -> InstanceInfo -> ShowS
[InstanceInfo] -> ShowS
InstanceInfo -> String
(Int -> InstanceInfo -> ShowS)
-> (InstanceInfo -> String)
-> ([InstanceInfo] -> ShowS)
-> Show InstanceInfo
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> InstanceInfo -> ShowS
showsPrec :: Int -> InstanceInfo -> ShowS
$cshow :: InstanceInfo -> String
show :: InstanceInfo -> String
$cshowList :: [InstanceInfo] -> ShowS
showList :: [InstanceInfo] -> ShowS
Show, (forall x. InstanceInfo -> Rep InstanceInfo x)
-> (forall x. Rep InstanceInfo x -> InstanceInfo)
-> Generic InstanceInfo
forall x. Rep InstanceInfo x -> InstanceInfo
forall x. InstanceInfo -> Rep InstanceInfo x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. InstanceInfo -> Rep InstanceInfo x
from :: forall x. InstanceInfo -> Rep InstanceInfo x
$cto :: forall x. Rep InstanceInfo x -> InstanceInfo
to :: forall x. Rep InstanceInfo x -> InstanceInfo
Generic)
data Definition = Defn
{ Definition -> ArgInfo
defArgInfo :: ArgInfo
, Definition -> QName
defName :: QName
, Definition -> Type
defType :: Type
, Definition -> [Polarity]
defPolarity :: [Polarity]
, Definition -> [Occurrence]
defArgOccurrences :: [Occurrence]
, Definition -> [Maybe Name]
defGeneralizedParams :: [Maybe Name]
, Definition -> [LocalDisplayForm]
defDisplay :: [LocalDisplayForm]
, Definition -> MutualId
defMutual :: MutualId
, Definition -> CompiledRepresentation
defCompiledRep :: CompiledRepresentation
, Definition -> Maybe InstanceInfo
defInstance :: Maybe InstanceInfo
, Definition -> Bool
defCopy :: Bool
, Definition -> Set QName
defMatchable :: Set QName
, Definition -> Bool
defNoCompilation :: Bool
, Definition -> Bool
defInjective :: Bool
, Definition -> Bool
defCopatternLHS :: Bool
, Definition -> Blocked_
defBlocked :: Blocked_
, Definition -> Language
defLanguage :: !Language
, Definition -> Defn
theDef :: Defn
}
deriving (Int -> Definition -> ShowS
[Definition] -> ShowS
Definition -> String
(Int -> Definition -> ShowS)
-> (Definition -> String)
-> ([Definition] -> ShowS)
-> Show Definition
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> Definition -> ShowS
showsPrec :: Int -> Definition -> ShowS
$cshow :: Definition -> String
show :: Definition -> String
$cshowList :: [Definition] -> ShowS
showList :: [Definition] -> ShowS
Show, (forall x. Definition -> Rep Definition x)
-> (forall x. Rep Definition x -> Definition) -> Generic Definition
forall x. Rep Definition x -> Definition
forall x. Definition -> Rep Definition x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. Definition -> Rep Definition x
from :: forall x. Definition -> Rep Definition x
$cto :: forall x. Rep Definition x -> Definition
to :: forall x. Rep Definition x -> Definition
Generic)
instance LensArgInfo Definition where
getArgInfo :: Definition -> ArgInfo
getArgInfo = Definition -> ArgInfo
defArgInfo
mapArgInfo :: (ArgInfo -> ArgInfo) -> Definition -> Definition
mapArgInfo ArgInfo -> ArgInfo
f Definition
def = Definition
def { defArgInfo = f $ defArgInfo def }
instance LensModality Definition where
instance LensQuantity Definition where
instance LensRelevance Definition where
data NumGeneralizableArgs
= NoGeneralizableArgs
| SomeGeneralizableArgs !Int
deriving Int -> NumGeneralizableArgs -> ShowS
[NumGeneralizableArgs] -> ShowS
NumGeneralizableArgs -> String
(Int -> NumGeneralizableArgs -> ShowS)
-> (NumGeneralizableArgs -> String)
-> ([NumGeneralizableArgs] -> ShowS)
-> Show NumGeneralizableArgs
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> NumGeneralizableArgs -> ShowS
showsPrec :: Int -> NumGeneralizableArgs -> ShowS
$cshow :: NumGeneralizableArgs -> String
show :: NumGeneralizableArgs -> String
$cshowList :: [NumGeneralizableArgs] -> ShowS
showList :: [NumGeneralizableArgs] -> ShowS
Show
lensTheDef :: Lens' Definition Defn
lensTheDef :: Lens' Definition Defn
lensTheDef Defn -> f Defn
f Definition
d = Defn -> f Defn
f (Definition -> Defn
theDef Definition
d) f Defn -> (Defn -> Definition) -> f Definition
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ Defn
df -> Definition
d { theDef = df }
defaultDefn ::
ArgInfo -> QName -> Type -> Language -> Defn -> Definition
defaultDefn :: ArgInfo -> QName -> Type -> Language -> Defn -> Definition
defaultDefn ArgInfo
info QName
x Type
t Language
lang Defn
def = Defn
{ defArgInfo :: ArgInfo
defArgInfo = ArgInfo
info
, defName :: QName
defName = QName
x
, defType :: Type
defType = Type
t
, defPolarity :: [Polarity]
defPolarity = []
, defArgOccurrences :: [Occurrence]
defArgOccurrences = []
, defGeneralizedParams :: [Maybe Name]
defGeneralizedParams = []
, defDisplay :: [LocalDisplayForm]
defDisplay = QName -> [LocalDisplayForm]
defaultDisplayForm QName
x
, defMutual :: MutualId
defMutual = MutualId
0
, defCompiledRep :: CompiledRepresentation
defCompiledRep = CompiledRepresentation
noCompiledRep
, defInstance :: Maybe InstanceInfo
defInstance = Maybe InstanceInfo
forall a. Maybe a
Nothing
, defCopy :: Bool
defCopy = Bool
False
, defMatchable :: Set QName
defMatchable = Set QName
forall a. Set a
Set.empty
, defNoCompilation :: Bool
defNoCompilation = Bool
False
, defInjective :: Bool
defInjective = Bool
False
, defCopatternLHS :: Bool
defCopatternLHS = Bool
False
, defBlocked :: Blocked_
defBlocked = NotBlocked' Term -> () -> Blocked_
forall t a. NotBlocked' t -> a -> Blocked' t a
NotBlocked NotBlocked' Term
forall t. NotBlocked' t
ReallyNotBlocked ()
, defLanguage :: Language
defLanguage = Language
lang
, theDef :: Defn
theDef = Defn
def
}
instance Pretty Polarity where
pretty :: Polarity -> Doc
pretty = String -> Doc
forall a. String -> Doc a
text (String -> Doc) -> (Polarity -> String) -> Polarity -> Doc
forall b c a. (b -> c) -> (a -> b) -> a -> c
. \case
Polarity
Covariant -> String
"+"
Polarity
Contravariant -> String
"-"
Polarity
Invariant -> String
"*"
Polarity
Nonvariant -> String
"_"
data IsForced
= Forced
| NotForced
deriving (Int -> IsForced -> ShowS
[IsForced] -> ShowS
IsForced -> String
(Int -> IsForced -> ShowS)
-> (IsForced -> String) -> ([IsForced] -> ShowS) -> Show IsForced
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> IsForced -> ShowS
showsPrec :: Int -> IsForced -> ShowS
$cshow :: IsForced -> String
show :: IsForced -> String
$cshowList :: [IsForced] -> ShowS
showList :: [IsForced] -> ShowS
Show, IsForced -> IsForced -> Bool
(IsForced -> IsForced -> Bool)
-> (IsForced -> IsForced -> Bool) -> Eq IsForced
forall a. (a -> a -> Bool) -> (a -> a -> Bool) -> Eq a
$c== :: IsForced -> IsForced -> Bool
== :: IsForced -> IsForced -> Bool
$c/= :: IsForced -> IsForced -> Bool
/= :: IsForced -> IsForced -> Bool
Eq, (forall x. IsForced -> Rep IsForced x)
-> (forall x. Rep IsForced x -> IsForced) -> Generic IsForced
forall x. Rep IsForced x -> IsForced
forall x. IsForced -> Rep IsForced x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. IsForced -> Rep IsForced x
from :: forall x. IsForced -> Rep IsForced x
$cto :: forall x. Rep IsForced x -> IsForced
to :: forall x. Rep IsForced x -> IsForced
Generic)
data CompilerPragma = CompilerPragma Range String
deriving (Int -> CompilerPragma -> ShowS
[CompilerPragma] -> ShowS
CompilerPragma -> String
(Int -> CompilerPragma -> ShowS)
-> (CompilerPragma -> String)
-> ([CompilerPragma] -> ShowS)
-> Show CompilerPragma
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> CompilerPragma -> ShowS
showsPrec :: Int -> CompilerPragma -> ShowS
$cshow :: CompilerPragma -> String
show :: CompilerPragma -> String
$cshowList :: [CompilerPragma] -> ShowS
showList :: [CompilerPragma] -> ShowS
Show, CompilerPragma -> CompilerPragma -> Bool
(CompilerPragma -> CompilerPragma -> Bool)
-> (CompilerPragma -> CompilerPragma -> Bool) -> Eq CompilerPragma
forall a. (a -> a -> Bool) -> (a -> a -> Bool) -> Eq a
$c== :: CompilerPragma -> CompilerPragma -> Bool
== :: CompilerPragma -> CompilerPragma -> Bool
$c/= :: CompilerPragma -> CompilerPragma -> Bool
/= :: CompilerPragma -> CompilerPragma -> Bool
Eq, (forall x. CompilerPragma -> Rep CompilerPragma x)
-> (forall x. Rep CompilerPragma x -> CompilerPragma)
-> Generic CompilerPragma
forall x. Rep CompilerPragma x -> CompilerPragma
forall x. CompilerPragma -> Rep CompilerPragma x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. CompilerPragma -> Rep CompilerPragma x
from :: forall x. CompilerPragma -> Rep CompilerPragma x
$cto :: forall x. Rep CompilerPragma x -> CompilerPragma
to :: forall x. Rep CompilerPragma x -> CompilerPragma
Generic)
instance HasRange CompilerPragma where
getRange :: CompilerPragma -> Range
getRange (CompilerPragma Range
r String
_) = Range
r
jsBackendName, ghcBackendName :: BackendName
jsBackendName :: BackendName
jsBackendName = BackendName
"JS"
ghcBackendName :: BackendName
ghcBackendName = BackendName
"GHC"
type CompiledRepresentation = Map BackendName [CompilerPragma]
noCompiledRep :: CompiledRepresentation
noCompiledRep :: CompiledRepresentation
noCompiledRep = CompiledRepresentation
forall k a. Map k a
Map.empty
type Face = [(Term,Bool)]
data System = System
{ System -> Telescope
systemTel :: Telescope
, System -> [(Face, Term)]
systemClauses :: [(Face,Term)]
} deriving (Int -> System -> ShowS
[System] -> ShowS
System -> String
(Int -> System -> ShowS)
-> (System -> String) -> ([System] -> ShowS) -> Show System
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> System -> ShowS
showsPrec :: Int -> System -> ShowS
$cshow :: System -> String
show :: System -> String
$cshowList :: [System] -> ShowS
showList :: [System] -> ShowS
Show, (forall x. System -> Rep System x)
-> (forall x. Rep System x -> System) -> Generic System
forall x. Rep System x -> System
forall x. System -> Rep System x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. System -> Rep System x
from :: forall x. System -> Rep System x
$cto :: forall x. Rep System x -> System
to :: forall x. Rep System x -> System
Generic)
data ExtLamInfo = ExtLamInfo
{ ExtLamInfo -> ModuleName
extLamModule :: ModuleName
, ExtLamInfo -> Bool
extLamAbsurd :: Bool
, ExtLamInfo -> Maybe System
extLamSys :: !(Strict.Maybe System)
} deriving (Int -> ExtLamInfo -> ShowS
[ExtLamInfo] -> ShowS
ExtLamInfo -> String
(Int -> ExtLamInfo -> ShowS)
-> (ExtLamInfo -> String)
-> ([ExtLamInfo] -> ShowS)
-> Show ExtLamInfo
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> ExtLamInfo -> ShowS
showsPrec :: Int -> ExtLamInfo -> ShowS
$cshow :: ExtLamInfo -> String
show :: ExtLamInfo -> String
$cshowList :: [ExtLamInfo] -> ShowS
showList :: [ExtLamInfo] -> ShowS
Show, (forall x. ExtLamInfo -> Rep ExtLamInfo x)
-> (forall x. Rep ExtLamInfo x -> ExtLamInfo) -> Generic ExtLamInfo
forall x. Rep ExtLamInfo x -> ExtLamInfo
forall x. ExtLamInfo -> Rep ExtLamInfo x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. ExtLamInfo -> Rep ExtLamInfo x
from :: forall x. ExtLamInfo -> Rep ExtLamInfo x
$cto :: forall x. Rep ExtLamInfo x -> ExtLamInfo
to :: forall x. Rep ExtLamInfo x -> ExtLamInfo
Generic)
modifySystem :: (System -> System) -> ExtLamInfo -> ExtLamInfo
modifySystem :: (System -> System) -> ExtLamInfo -> ExtLamInfo
modifySystem System -> System
f ExtLamInfo
e = let !e' :: ExtLamInfo
e' = ExtLamInfo
e { extLamSys = f <$> extLamSys e } in ExtLamInfo
e'
data Projection = Projection
{ Projection -> Maybe QName
projProper :: Maybe QName
, Projection -> QName
projOrig :: QName
, Projection -> Arg QName
projFromType :: Arg QName
, Projection -> Int
projIndex :: Int
, Projection -> ProjLams
projLams :: ProjLams
} deriving (Int -> Projection -> ShowS
[Projection] -> ShowS
Projection -> String
(Int -> Projection -> ShowS)
-> (Projection -> String)
-> ([Projection] -> ShowS)
-> Show Projection
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> Projection -> ShowS
showsPrec :: Int -> Projection -> ShowS
$cshow :: Projection -> String
show :: Projection -> String
$cshowList :: [Projection] -> ShowS
showList :: [Projection] -> ShowS
Show, (forall x. Projection -> Rep Projection x)
-> (forall x. Rep Projection x -> Projection) -> Generic Projection
forall x. Rep Projection x -> Projection
forall x. Projection -> Rep Projection x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. Projection -> Rep Projection x
from :: forall x. Projection -> Rep Projection x
$cto :: forall x. Rep Projection x -> Projection
to :: forall x. Rep Projection x -> Projection
Generic)
newtype ProjLams = ProjLams { ProjLams -> [Arg String]
getProjLams :: [Arg ArgName] }
deriving (Int -> ProjLams -> ShowS
[ProjLams] -> ShowS
ProjLams -> String
(Int -> ProjLams -> ShowS)
-> (ProjLams -> String) -> ([ProjLams] -> ShowS) -> Show ProjLams
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> ProjLams -> ShowS
showsPrec :: Int -> ProjLams -> ShowS
$cshow :: ProjLams -> String
show :: ProjLams -> String
$cshowList :: [ProjLams] -> ShowS
showList :: [ProjLams] -> ShowS
Show, ProjLams
ProjLams -> Bool
ProjLams -> (ProjLams -> Bool) -> Null ProjLams
forall a. a -> (a -> Bool) -> Null a
$cempty :: ProjLams
empty :: ProjLams
$cnull :: ProjLams -> Bool
null :: ProjLams -> Bool
Null, (forall x. ProjLams -> Rep ProjLams x)
-> (forall x. Rep ProjLams x -> ProjLams) -> Generic ProjLams
forall x. Rep ProjLams x -> ProjLams
forall x. ProjLams -> Rep ProjLams x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. ProjLams -> Rep ProjLams x
from :: forall x. ProjLams -> Rep ProjLams x
$cto :: forall x. Rep ProjLams x -> ProjLams
to :: forall x. Rep ProjLams x -> ProjLams
Generic)
projDropPars :: Projection -> ProjOrigin -> Term
projDropPars :: Projection -> ProjOrigin -> Term
projDropPars (Projection Just{} QName
d Arg QName
_ Int
_ ProjLams
lams) ProjOrigin
o =
case [Arg String] -> Maybe ([Arg String], Arg String)
forall a. [a] -> Maybe ([a], a)
initLast ([Arg String] -> Maybe ([Arg String], Arg String))
-> [Arg String] -> Maybe ([Arg String], Arg String)
forall a b. (a -> b) -> a -> b
$ ProjLams -> [Arg String]
getProjLams ProjLams
lams of
Maybe ([Arg String], Arg String)
Nothing -> QName -> [Elim] -> Term
Def QName
d []
Just ([Arg String]
pars, Arg ArgInfo
i String
y) ->
let core :: Term
core = ArgInfo -> Abs Term -> Term
Lam ArgInfo
i (Abs Term -> Term) -> Abs Term -> Term
forall a b. (a -> b) -> a -> b
$ String -> Term -> Abs Term
forall a. String -> a -> Abs a
Abs String
y (Term -> Abs Term) -> Term -> Abs Term
forall a b. (a -> b) -> a -> b
$ Int -> [Elim] -> Term
Var Int
0 [ProjOrigin -> QName -> Elim
forall a. ProjOrigin -> QName -> Elim' a
Proj ProjOrigin
o QName
d] in
(Arg String -> Term -> Term) -> Term -> [Arg String] -> Term
forall a b. (a -> b -> b) -> b -> [a] -> b
forall (t :: * -> *) a b.
Foldable t =>
(a -> b -> b) -> b -> t a -> b
List.foldr (\ (Arg ArgInfo
ai String
x) -> ArgInfo -> Abs Term -> Term
Lam ArgInfo
ai (Abs Term -> Term) -> (Term -> Abs Term) -> Term -> Term
forall b c a. (b -> c) -> (a -> b) -> a -> c
. String -> Term -> Abs Term
forall a. String -> a -> Abs a
NoAbs String
x) Term
core [Arg String]
pars
projDropPars (Projection Maybe QName
Nothing QName
d Arg QName
_ Int
_ ProjLams
lams) ProjOrigin
o =
(Arg String -> Term -> Term) -> Term -> [Arg String] -> Term
forall a b. (a -> b -> b) -> b -> [a] -> b
forall (t :: * -> *) a b.
Foldable t =>
(a -> b -> b) -> b -> t a -> b
List.foldr (\ (Arg ArgInfo
ai String
x) -> ArgInfo -> Abs Term -> Term
Lam ArgInfo
ai (Abs Term -> Term) -> (Term -> Abs Term) -> Term -> Term
forall b c a. (b -> c) -> (a -> b) -> a -> c
. String -> Term -> Abs Term
forall a. String -> a -> Abs a
NoAbs String
x) (QName -> [Elim] -> Term
Def QName
d []) ([Arg String] -> Term) -> [Arg String] -> Term
forall a b. (a -> b) -> a -> b
$
[Arg String] -> [Arg String] -> [Arg String]
forall a. [a] -> [a] -> [a]
initWithDefault [Arg String]
forall a. HasCallStack => a
__IMPOSSIBLE__ ([Arg String] -> [Arg String]) -> [Arg String] -> [Arg String]
forall a b. (a -> b) -> a -> b
$ ProjLams -> [Arg String]
getProjLams ProjLams
lams
projArgInfo :: Projection -> ArgInfo
projArgInfo :: Projection -> ArgInfo
projArgInfo (Projection Maybe QName
_ QName
_ Arg QName
_ Int
_ ProjLams
lams) =
ArgInfo -> (Arg String -> ArgInfo) -> Maybe (Arg String) -> ArgInfo
forall b a. b -> (a -> b) -> Maybe a -> b
maybe ArgInfo
forall a. HasCallStack => a
__IMPOSSIBLE__ Arg String -> ArgInfo
forall a. LensArgInfo a => a -> ArgInfo
getArgInfo (Maybe (Arg String) -> ArgInfo) -> Maybe (Arg String) -> ArgInfo
forall a b. (a -> b) -> a -> b
$ [Arg String] -> Maybe (Arg String)
forall a. [a] -> Maybe a
lastMaybe ([Arg String] -> Maybe (Arg String))
-> [Arg String] -> Maybe (Arg String)
forall a b. (a -> b) -> a -> b
$ ProjLams -> [Arg String]
getProjLams ProjLams
lams
data EtaEquality
= Specified { EtaEquality -> HasEta
theEtaEquality :: !HasEta }
| Inferred { theEtaEquality :: !HasEta }
deriving (Int -> EtaEquality -> ShowS
[EtaEquality] -> ShowS
EtaEquality -> String
(Int -> EtaEquality -> ShowS)
-> (EtaEquality -> String)
-> ([EtaEquality] -> ShowS)
-> Show EtaEquality
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> EtaEquality -> ShowS
showsPrec :: Int -> EtaEquality -> ShowS
$cshow :: EtaEquality -> String
show :: EtaEquality -> String
$cshowList :: [EtaEquality] -> ShowS
showList :: [EtaEquality] -> ShowS
Show, EtaEquality -> EtaEquality -> Bool
(EtaEquality -> EtaEquality -> Bool)
-> (EtaEquality -> EtaEquality -> Bool) -> Eq EtaEquality
forall a. (a -> a -> Bool) -> (a -> a -> Bool) -> Eq a
$c== :: EtaEquality -> EtaEquality -> Bool
== :: EtaEquality -> EtaEquality -> Bool
$c/= :: EtaEquality -> EtaEquality -> Bool
/= :: EtaEquality -> EtaEquality -> Bool
Eq, (forall x. EtaEquality -> Rep EtaEquality x)
-> (forall x. Rep EtaEquality x -> EtaEquality)
-> Generic EtaEquality
forall x. Rep EtaEquality x -> EtaEquality
forall x. EtaEquality -> Rep EtaEquality x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. EtaEquality -> Rep EtaEquality x
from :: forall x. EtaEquality -> Rep EtaEquality x
$cto :: forall x. Rep EtaEquality x -> EtaEquality
to :: forall x. Rep EtaEquality x -> EtaEquality
Generic)
instance PatternMatchingAllowed EtaEquality where
patternMatchingAllowed :: EtaEquality -> Bool
patternMatchingAllowed = HasEta -> Bool
forall a. PatternMatchingAllowed a => a -> Bool
patternMatchingAllowed (HasEta -> Bool) -> (EtaEquality -> HasEta) -> EtaEquality -> Bool
forall b c a. (b -> c) -> (a -> b) -> a -> c
. EtaEquality -> HasEta
theEtaEquality
instance CopatternMatchingAllowed EtaEquality where
copatternMatchingAllowed :: EtaEquality -> Bool
copatternMatchingAllowed = HasEta -> Bool
forall a. CopatternMatchingAllowed a => a -> Bool
copatternMatchingAllowed (HasEta -> Bool) -> (EtaEquality -> HasEta) -> EtaEquality -> Bool
forall b c a. (b -> c) -> (a -> b) -> a -> c
. EtaEquality -> HasEta
theEtaEquality
setEtaEquality :: EtaEquality -> HasEta -> EtaEquality
setEtaEquality :: EtaEquality -> HasEta -> EtaEquality
setEtaEquality e :: EtaEquality
e@Specified{} HasEta
_ = EtaEquality
e
setEtaEquality EtaEquality
_ HasEta
b = HasEta -> EtaEquality
Inferred HasEta
b
data FunctionFlag
= FunStatic
| FunInline
| FunMacro
| FunFirstOrder
| FunErasure
| FunAbstract
| FunProj
deriving (FunctionFlag -> FunctionFlag -> Bool
(FunctionFlag -> FunctionFlag -> Bool)
-> (FunctionFlag -> FunctionFlag -> Bool) -> Eq FunctionFlag
forall a. (a -> a -> Bool) -> (a -> a -> Bool) -> Eq a
$c== :: FunctionFlag -> FunctionFlag -> Bool
== :: FunctionFlag -> FunctionFlag -> Bool
$c/= :: FunctionFlag -> FunctionFlag -> Bool
/= :: FunctionFlag -> FunctionFlag -> Bool
Eq, Eq FunctionFlag
Eq FunctionFlag =>
(FunctionFlag -> FunctionFlag -> Ordering)
-> (FunctionFlag -> FunctionFlag -> Bool)
-> (FunctionFlag -> FunctionFlag -> Bool)
-> (FunctionFlag -> FunctionFlag -> Bool)
-> (FunctionFlag -> FunctionFlag -> Bool)
-> (FunctionFlag -> FunctionFlag -> FunctionFlag)
-> (FunctionFlag -> FunctionFlag -> FunctionFlag)
-> Ord FunctionFlag
FunctionFlag -> FunctionFlag -> Bool
FunctionFlag -> FunctionFlag -> Ordering
FunctionFlag -> FunctionFlag -> FunctionFlag
forall a.
Eq a =>
(a -> a -> Ordering)
-> (a -> a -> Bool)
-> (a -> a -> Bool)
-> (a -> a -> Bool)
-> (a -> a -> Bool)
-> (a -> a -> a)
-> (a -> a -> a)
-> Ord a
$ccompare :: FunctionFlag -> FunctionFlag -> Ordering
compare :: FunctionFlag -> FunctionFlag -> Ordering
$c< :: FunctionFlag -> FunctionFlag -> Bool
< :: FunctionFlag -> FunctionFlag -> Bool
$c<= :: FunctionFlag -> FunctionFlag -> Bool
<= :: FunctionFlag -> FunctionFlag -> Bool
$c> :: FunctionFlag -> FunctionFlag -> Bool
> :: FunctionFlag -> FunctionFlag -> Bool
$c>= :: FunctionFlag -> FunctionFlag -> Bool
>= :: FunctionFlag -> FunctionFlag -> Bool
$cmax :: FunctionFlag -> FunctionFlag -> FunctionFlag
max :: FunctionFlag -> FunctionFlag -> FunctionFlag
$cmin :: FunctionFlag -> FunctionFlag -> FunctionFlag
min :: FunctionFlag -> FunctionFlag -> FunctionFlag
Ord, Int -> FunctionFlag
FunctionFlag -> Int
FunctionFlag -> [FunctionFlag]
FunctionFlag -> FunctionFlag
FunctionFlag -> FunctionFlag -> [FunctionFlag]
FunctionFlag -> FunctionFlag -> FunctionFlag -> [FunctionFlag]
(FunctionFlag -> FunctionFlag)
-> (FunctionFlag -> FunctionFlag)
-> (Int -> FunctionFlag)
-> (FunctionFlag -> Int)
-> (FunctionFlag -> [FunctionFlag])
-> (FunctionFlag -> FunctionFlag -> [FunctionFlag])
-> (FunctionFlag -> FunctionFlag -> [FunctionFlag])
-> (FunctionFlag -> FunctionFlag -> FunctionFlag -> [FunctionFlag])
-> Enum FunctionFlag
forall a.
(a -> a)
-> (a -> a)
-> (Int -> a)
-> (a -> Int)
-> (a -> [a])
-> (a -> a -> [a])
-> (a -> a -> [a])
-> (a -> a -> a -> [a])
-> Enum a
$csucc :: FunctionFlag -> FunctionFlag
succ :: FunctionFlag -> FunctionFlag
$cpred :: FunctionFlag -> FunctionFlag
pred :: FunctionFlag -> FunctionFlag
$ctoEnum :: Int -> FunctionFlag
toEnum :: Int -> FunctionFlag
$cfromEnum :: FunctionFlag -> Int
fromEnum :: FunctionFlag -> Int
$cenumFrom :: FunctionFlag -> [FunctionFlag]
enumFrom :: FunctionFlag -> [FunctionFlag]
$cenumFromThen :: FunctionFlag -> FunctionFlag -> [FunctionFlag]
enumFromThen :: FunctionFlag -> FunctionFlag -> [FunctionFlag]
$cenumFromTo :: FunctionFlag -> FunctionFlag -> [FunctionFlag]
enumFromTo :: FunctionFlag -> FunctionFlag -> [FunctionFlag]
$cenumFromThenTo :: FunctionFlag -> FunctionFlag -> FunctionFlag -> [FunctionFlag]
enumFromThenTo :: FunctionFlag -> FunctionFlag -> FunctionFlag -> [FunctionFlag]
Enum, Int -> FunctionFlag -> ShowS
[FunctionFlag] -> ShowS
FunctionFlag -> String
(Int -> FunctionFlag -> ShowS)
-> (FunctionFlag -> String)
-> ([FunctionFlag] -> ShowS)
-> Show FunctionFlag
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> FunctionFlag -> ShowS
showsPrec :: Int -> FunctionFlag -> ShowS
$cshow :: FunctionFlag -> String
show :: FunctionFlag -> String
$cshowList :: [FunctionFlag] -> ShowS
showList :: [FunctionFlag] -> ShowS
Show, (forall x. FunctionFlag -> Rep FunctionFlag x)
-> (forall x. Rep FunctionFlag x -> FunctionFlag)
-> Generic FunctionFlag
forall x. Rep FunctionFlag x -> FunctionFlag
forall x. FunctionFlag -> Rep FunctionFlag x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. FunctionFlag -> Rep FunctionFlag x
from :: forall x. FunctionFlag -> Rep FunctionFlag x
$cto :: forall x. Rep FunctionFlag x -> FunctionFlag
to :: forall x. Rep FunctionFlag x -> FunctionFlag
Generic, Ord FunctionFlag
Ord FunctionFlag =>
((FunctionFlag, FunctionFlag) -> [FunctionFlag])
-> ((FunctionFlag, FunctionFlag) -> FunctionFlag -> Int)
-> ((FunctionFlag, FunctionFlag) -> FunctionFlag -> Int)
-> ((FunctionFlag, FunctionFlag) -> FunctionFlag -> Bool)
-> ((FunctionFlag, FunctionFlag) -> Int)
-> ((FunctionFlag, FunctionFlag) -> Int)
-> Ix FunctionFlag
(FunctionFlag, FunctionFlag) -> Int
(FunctionFlag, FunctionFlag) -> [FunctionFlag]
(FunctionFlag, FunctionFlag) -> FunctionFlag -> Bool
(FunctionFlag, FunctionFlag) -> FunctionFlag -> Int
forall a.
Ord a =>
((a, a) -> [a])
-> ((a, a) -> a -> Int)
-> ((a, a) -> a -> Int)
-> ((a, a) -> a -> Bool)
-> ((a, a) -> Int)
-> ((a, a) -> Int)
-> Ix a
$crange :: (FunctionFlag, FunctionFlag) -> [FunctionFlag]
range :: (FunctionFlag, FunctionFlag) -> [FunctionFlag]
$cindex :: (FunctionFlag, FunctionFlag) -> FunctionFlag -> Int
index :: (FunctionFlag, FunctionFlag) -> FunctionFlag -> Int
$cunsafeIndex :: (FunctionFlag, FunctionFlag) -> FunctionFlag -> Int
unsafeIndex :: (FunctionFlag, FunctionFlag) -> FunctionFlag -> Int
$cinRange :: (FunctionFlag, FunctionFlag) -> FunctionFlag -> Bool
inRange :: (FunctionFlag, FunctionFlag) -> FunctionFlag -> Bool
$crangeSize :: (FunctionFlag, FunctionFlag) -> Int
rangeSize :: (FunctionFlag, FunctionFlag) -> Int
$cunsafeRangeSize :: (FunctionFlag, FunctionFlag) -> Int
unsafeRangeSize :: (FunctionFlag, FunctionFlag) -> Int
Ix, FunctionFlag
FunctionFlag -> FunctionFlag -> Bounded FunctionFlag
forall a. a -> a -> Bounded a
$cminBound :: FunctionFlag
minBound :: FunctionFlag
$cmaxBound :: FunctionFlag
maxBound :: FunctionFlag
Bounded)
instance SmallSetElement FunctionFlag
instance KillRange (SmallSet FunctionFlag) where killRange :: KillRangeT (SmallSet FunctionFlag)
killRange = KillRangeT (SmallSet FunctionFlag)
forall a. a -> a
id
data CompKit = CompKit
{ CompKit -> Maybe QName
nameOfHComp :: Maybe QName
, CompKit -> Maybe QName
nameOfTransp :: Maybe QName
}
deriving (CompKit -> CompKit -> Bool
(CompKit -> CompKit -> Bool)
-> (CompKit -> CompKit -> Bool) -> Eq CompKit
forall a. (a -> a -> Bool) -> (a -> a -> Bool) -> Eq a
$c== :: CompKit -> CompKit -> Bool
== :: CompKit -> CompKit -> Bool
$c/= :: CompKit -> CompKit -> Bool
/= :: CompKit -> CompKit -> Bool
Eq, Eq CompKit
Eq CompKit =>
(CompKit -> CompKit -> Ordering)
-> (CompKit -> CompKit -> Bool)
-> (CompKit -> CompKit -> Bool)
-> (CompKit -> CompKit -> Bool)
-> (CompKit -> CompKit -> Bool)
-> (CompKit -> CompKit -> CompKit)
-> (CompKit -> CompKit -> CompKit)
-> Ord CompKit
CompKit -> CompKit -> Bool
CompKit -> CompKit -> Ordering
CompKit -> CompKit -> CompKit
forall a.
Eq a =>
(a -> a -> Ordering)
-> (a -> a -> Bool)
-> (a -> a -> Bool)
-> (a -> a -> Bool)
-> (a -> a -> Bool)
-> (a -> a -> a)
-> (a -> a -> a)
-> Ord a
$ccompare :: CompKit -> CompKit -> Ordering
compare :: CompKit -> CompKit -> Ordering
$c< :: CompKit -> CompKit -> Bool
< :: CompKit -> CompKit -> Bool
$c<= :: CompKit -> CompKit -> Bool
<= :: CompKit -> CompKit -> Bool
$c> :: CompKit -> CompKit -> Bool
> :: CompKit -> CompKit -> Bool
$c>= :: CompKit -> CompKit -> Bool
>= :: CompKit -> CompKit -> Bool
$cmax :: CompKit -> CompKit -> CompKit
max :: CompKit -> CompKit -> CompKit
$cmin :: CompKit -> CompKit -> CompKit
min :: CompKit -> CompKit -> CompKit
Ord, Int -> CompKit -> ShowS
[CompKit] -> ShowS
CompKit -> String
(Int -> CompKit -> ShowS)
-> (CompKit -> String) -> ([CompKit] -> ShowS) -> Show CompKit
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> CompKit -> ShowS
showsPrec :: Int -> CompKit -> ShowS
$cshow :: CompKit -> String
show :: CompKit -> String
$cshowList :: [CompKit] -> ShowS
showList :: [CompKit] -> ShowS
Show, (forall x. CompKit -> Rep CompKit x)
-> (forall x. Rep CompKit x -> CompKit) -> Generic CompKit
forall x. Rep CompKit x -> CompKit
forall x. CompKit -> Rep CompKit x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. CompKit -> Rep CompKit x
from :: forall x. CompKit -> Rep CompKit x
$cto :: forall x. Rep CompKit x -> CompKit
to :: forall x. Rep CompKit x -> CompKit
Generic)
emptyCompKit :: CompKit
emptyCompKit :: CompKit
emptyCompKit = Maybe QName -> Maybe QName -> CompKit
CompKit Maybe QName
forall a. Maybe a
Nothing Maybe QName
forall a. Maybe a
Nothing
defaultAxiom :: Defn
defaultAxiom :: Defn
defaultAxiom = Bool -> Defn
Axiom Bool
False
constTranspAxiom :: Defn
constTranspAxiom :: Defn
constTranspAxiom = Bool -> Defn
Axiom Bool
True
data Defn
= AxiomDefn AxiomData
| DataOrRecSigDefn DataOrRecSigData
| GeneralizableVar
NumGeneralizableArgs
| AbstractDefn Defn
| FunctionDefn FunctionData
| DatatypeDefn DatatypeData
| RecordDefn RecordData
| ConstructorDefn ConstructorData
| PrimitiveDefn PrimitiveData
| PrimitiveSortDefn PrimitiveSortData
deriving (Int -> Defn -> ShowS
[Defn] -> ShowS
Defn -> String
(Int -> Defn -> ShowS)
-> (Defn -> String) -> ([Defn] -> ShowS) -> Show Defn
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> Defn -> ShowS
showsPrec :: Int -> Defn -> ShowS
$cshow :: Defn -> String
show :: Defn -> String
$cshowList :: [Defn] -> ShowS
showList :: [Defn] -> ShowS
Show, (forall x. Defn -> Rep Defn x)
-> (forall x. Rep Defn x -> Defn) -> Generic Defn
forall x. Rep Defn x -> Defn
forall x. Defn -> Rep Defn x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. Defn -> Rep Defn x
from :: forall x. Defn -> Rep Defn x
$cto :: forall x. Rep Defn x -> Defn
to :: forall x. Rep Defn x -> Defn
Generic)
{-# COMPLETE
Axiom, DataOrRecSig, GeneralizableVar, AbstractDefn,
Function, Datatype, Record, Constructor, Primitive, PrimitiveSort #-}
data AxiomData = AxiomData
{ AxiomData -> Bool
_axiomConstTransp :: Bool
} deriving (Int -> AxiomData -> ShowS
[AxiomData] -> ShowS
AxiomData -> String
(Int -> AxiomData -> ShowS)
-> (AxiomData -> String)
-> ([AxiomData] -> ShowS)
-> Show AxiomData
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> AxiomData -> ShowS
showsPrec :: Int -> AxiomData -> ShowS
$cshow :: AxiomData -> String
show :: AxiomData -> String
$cshowList :: [AxiomData] -> ShowS
showList :: [AxiomData] -> ShowS
Show, (forall x. AxiomData -> Rep AxiomData x)
-> (forall x. Rep AxiomData x -> AxiomData) -> Generic AxiomData
forall x. Rep AxiomData x -> AxiomData
forall x. AxiomData -> Rep AxiomData x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. AxiomData -> Rep AxiomData x
from :: forall x. AxiomData -> Rep AxiomData x
$cto :: forall x. Rep AxiomData x -> AxiomData
to :: forall x. Rep AxiomData x -> AxiomData
Generic)
pattern Axiom :: Bool -> Defn
pattern $mAxiom :: forall {r}. Defn -> (Bool -> r) -> ((# #) -> r) -> r
$bAxiom :: Bool -> Defn
Axiom{ Defn -> Bool
axiomConstTransp } = AxiomDefn (AxiomData axiomConstTransp)
data DataOrRecSigData = DataOrRecSigData
{ DataOrRecSigData -> Int
_datarecPars :: Int
} deriving (Int -> DataOrRecSigData -> ShowS
[DataOrRecSigData] -> ShowS
DataOrRecSigData -> String
(Int -> DataOrRecSigData -> ShowS)
-> (DataOrRecSigData -> String)
-> ([DataOrRecSigData] -> ShowS)
-> Show DataOrRecSigData
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> DataOrRecSigData -> ShowS
showsPrec :: Int -> DataOrRecSigData -> ShowS
$cshow :: DataOrRecSigData -> String
show :: DataOrRecSigData -> String
$cshowList :: [DataOrRecSigData] -> ShowS
showList :: [DataOrRecSigData] -> ShowS
Show, (forall x. DataOrRecSigData -> Rep DataOrRecSigData x)
-> (forall x. Rep DataOrRecSigData x -> DataOrRecSigData)
-> Generic DataOrRecSigData
forall x. Rep DataOrRecSigData x -> DataOrRecSigData
forall x. DataOrRecSigData -> Rep DataOrRecSigData x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. DataOrRecSigData -> Rep DataOrRecSigData x
from :: forall x. DataOrRecSigData -> Rep DataOrRecSigData x
$cto :: forall x. Rep DataOrRecSigData x -> DataOrRecSigData
to :: forall x. Rep DataOrRecSigData x -> DataOrRecSigData
Generic)
pattern DataOrRecSig :: Int -> Defn
pattern $mDataOrRecSig :: forall {r}. Defn -> (Int -> r) -> ((# #) -> r) -> r
$bDataOrRecSig :: Int -> Defn
DataOrRecSig{ Defn -> Int
datarecPars } = DataOrRecSigDefn (DataOrRecSigData datarecPars)
data ProjectionLikenessMissing
= MaybeProjection
| NeverProjection
deriving (Int -> ProjectionLikenessMissing -> ShowS
[ProjectionLikenessMissing] -> ShowS
ProjectionLikenessMissing -> String
(Int -> ProjectionLikenessMissing -> ShowS)
-> (ProjectionLikenessMissing -> String)
-> ([ProjectionLikenessMissing] -> ShowS)
-> Show ProjectionLikenessMissing
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> ProjectionLikenessMissing -> ShowS
showsPrec :: Int -> ProjectionLikenessMissing -> ShowS
$cshow :: ProjectionLikenessMissing -> String
show :: ProjectionLikenessMissing -> String
$cshowList :: [ProjectionLikenessMissing] -> ShowS
showList :: [ProjectionLikenessMissing] -> ShowS
Show, (forall x.
ProjectionLikenessMissing -> Rep ProjectionLikenessMissing x)
-> (forall x.
Rep ProjectionLikenessMissing x -> ProjectionLikenessMissing)
-> Generic ProjectionLikenessMissing
forall x.
Rep ProjectionLikenessMissing x -> ProjectionLikenessMissing
forall x.
ProjectionLikenessMissing -> Rep ProjectionLikenessMissing x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x.
ProjectionLikenessMissing -> Rep ProjectionLikenessMissing x
from :: forall x.
ProjectionLikenessMissing -> Rep ProjectionLikenessMissing x
$cto :: forall x.
Rep ProjectionLikenessMissing x -> ProjectionLikenessMissing
to :: forall x.
Rep ProjectionLikenessMissing x -> ProjectionLikenessMissing
Generic, Int -> ProjectionLikenessMissing
ProjectionLikenessMissing -> Int
ProjectionLikenessMissing -> [ProjectionLikenessMissing]
ProjectionLikenessMissing -> ProjectionLikenessMissing
ProjectionLikenessMissing
-> ProjectionLikenessMissing -> [ProjectionLikenessMissing]
ProjectionLikenessMissing
-> ProjectionLikenessMissing
-> ProjectionLikenessMissing
-> [ProjectionLikenessMissing]
(ProjectionLikenessMissing -> ProjectionLikenessMissing)
-> (ProjectionLikenessMissing -> ProjectionLikenessMissing)
-> (Int -> ProjectionLikenessMissing)
-> (ProjectionLikenessMissing -> Int)
-> (ProjectionLikenessMissing -> [ProjectionLikenessMissing])
-> (ProjectionLikenessMissing
-> ProjectionLikenessMissing -> [ProjectionLikenessMissing])
-> (ProjectionLikenessMissing
-> ProjectionLikenessMissing -> [ProjectionLikenessMissing])
-> (ProjectionLikenessMissing
-> ProjectionLikenessMissing
-> ProjectionLikenessMissing
-> [ProjectionLikenessMissing])
-> Enum ProjectionLikenessMissing
forall a.
(a -> a)
-> (a -> a)
-> (Int -> a)
-> (a -> Int)
-> (a -> [a])
-> (a -> a -> [a])
-> (a -> a -> [a])
-> (a -> a -> a -> [a])
-> Enum a
$csucc :: ProjectionLikenessMissing -> ProjectionLikenessMissing
succ :: ProjectionLikenessMissing -> ProjectionLikenessMissing
$cpred :: ProjectionLikenessMissing -> ProjectionLikenessMissing
pred :: ProjectionLikenessMissing -> ProjectionLikenessMissing
$ctoEnum :: Int -> ProjectionLikenessMissing
toEnum :: Int -> ProjectionLikenessMissing
$cfromEnum :: ProjectionLikenessMissing -> Int
fromEnum :: ProjectionLikenessMissing -> Int
$cenumFrom :: ProjectionLikenessMissing -> [ProjectionLikenessMissing]
enumFrom :: ProjectionLikenessMissing -> [ProjectionLikenessMissing]
$cenumFromThen :: ProjectionLikenessMissing
-> ProjectionLikenessMissing -> [ProjectionLikenessMissing]
enumFromThen :: ProjectionLikenessMissing
-> ProjectionLikenessMissing -> [ProjectionLikenessMissing]
$cenumFromTo :: ProjectionLikenessMissing
-> ProjectionLikenessMissing -> [ProjectionLikenessMissing]
enumFromTo :: ProjectionLikenessMissing
-> ProjectionLikenessMissing -> [ProjectionLikenessMissing]
$cenumFromThenTo :: ProjectionLikenessMissing
-> ProjectionLikenessMissing
-> ProjectionLikenessMissing
-> [ProjectionLikenessMissing]
enumFromThenTo :: ProjectionLikenessMissing
-> ProjectionLikenessMissing
-> ProjectionLikenessMissing
-> [ProjectionLikenessMissing]
Enum, ProjectionLikenessMissing
ProjectionLikenessMissing
-> ProjectionLikenessMissing -> Bounded ProjectionLikenessMissing
forall a. a -> a -> Bounded a
$cminBound :: ProjectionLikenessMissing
minBound :: ProjectionLikenessMissing
$cmaxBound :: ProjectionLikenessMissing
maxBound :: ProjectionLikenessMissing
Bounded)
data FunctionData = FunctionData
{ FunctionData -> [Clause]
_funClauses :: [Clause]
, FunctionData -> Maybe CompiledClauses
_funCompiled :: Maybe CompiledClauses
, FunctionData -> Maybe SplitTree
_funSplitTree :: Maybe SplitTree
, FunctionData -> Maybe Compiled
_funTreeless :: Maybe Compiled
, FunctionData -> [Clause]
_funCovering :: [Clause]
, FunctionData -> FunctionInverse
_funInv :: FunctionInverse
, FunctionData -> Maybe [QName]
_funMutual :: Maybe [QName]
, FunctionData -> Either ProjectionLikenessMissing Projection
_funProjection :: Either ProjectionLikenessMissing Projection
, FunctionData -> SmallSet FunctionFlag
_funFlags :: SmallSet FunctionFlag
, FunctionData -> Maybe Bool
_funTerminates :: Maybe Bool
, FunctionData -> Maybe ExtLamInfo
_funExtLam :: Maybe ExtLamInfo
, FunctionData -> Maybe QName
_funWith :: Maybe QName
, FunctionData -> Maybe QName
_funIsKanOp :: Maybe QName
, FunctionData -> IsOpaque
_funOpaque :: IsOpaque
} deriving (Int -> FunctionData -> ShowS
[FunctionData] -> ShowS
FunctionData -> String
(Int -> FunctionData -> ShowS)
-> (FunctionData -> String)
-> ([FunctionData] -> ShowS)
-> Show FunctionData
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> FunctionData -> ShowS
showsPrec :: Int -> FunctionData -> ShowS
$cshow :: FunctionData -> String
show :: FunctionData -> String
$cshowList :: [FunctionData] -> ShowS
showList :: [FunctionData] -> ShowS
Show, (forall x. FunctionData -> Rep FunctionData x)
-> (forall x. Rep FunctionData x -> FunctionData)
-> Generic FunctionData
forall x. Rep FunctionData x -> FunctionData
forall x. FunctionData -> Rep FunctionData x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. FunctionData -> Rep FunctionData x
from :: forall x. FunctionData -> Rep FunctionData x
$cto :: forall x. Rep FunctionData x -> FunctionData
to :: forall x. Rep FunctionData x -> FunctionData
Generic)
pattern Function
:: [Clause]
-> Maybe CompiledClauses
-> Maybe SplitTree
-> Maybe Compiled
-> [Clause]
-> FunctionInverse
-> Maybe [QName]
-> Either ProjectionLikenessMissing Projection
-> SmallSet FunctionFlag
-> Maybe Bool
-> Maybe ExtLamInfo
-> Maybe QName
-> Maybe QName
-> IsOpaque
-> Defn
pattern $mFunction :: forall {r}.
Defn
-> ([Clause]
-> Maybe CompiledClauses
-> Maybe SplitTree
-> Maybe Compiled
-> [Clause]
-> FunctionInverse
-> Maybe [QName]
-> Either ProjectionLikenessMissing Projection
-> SmallSet FunctionFlag
-> Maybe Bool
-> Maybe ExtLamInfo
-> Maybe QName
-> Maybe QName
-> IsOpaque
-> r)
-> ((# #) -> r)
-> r
$bFunction :: [Clause]
-> Maybe CompiledClauses
-> Maybe SplitTree
-> Maybe Compiled
-> [Clause]
-> FunctionInverse
-> Maybe [QName]
-> Either ProjectionLikenessMissing Projection
-> SmallSet FunctionFlag
-> Maybe Bool
-> Maybe ExtLamInfo
-> Maybe QName
-> Maybe QName
-> IsOpaque
-> Defn
Function
{ Defn -> [Clause]
funClauses
, Defn -> Maybe CompiledClauses
funCompiled
, Defn -> Maybe SplitTree
funSplitTree
, Defn -> Maybe Compiled
funTreeless
, Defn -> [Clause]
funCovering
, Defn -> FunctionInverse
funInv
, Defn -> Maybe [QName]
funMutual
, Defn -> Either ProjectionLikenessMissing Projection
funProjection
, Defn -> SmallSet FunctionFlag
funFlags
, Defn -> Maybe Bool
funTerminates
, Defn -> Maybe ExtLamInfo
funExtLam
, Defn -> Maybe QName
funWith
, Defn -> Maybe QName
funIsKanOp
, Defn -> IsOpaque
funOpaque
} = FunctionDefn (FunctionData
funClauses
funCompiled
funSplitTree
funTreeless
funCovering
funInv
funMutual
funProjection
funFlags
funTerminates
funExtLam
funWith
funIsKanOp
funOpaque
)
data DatatypeData = DatatypeData
{ DatatypeData -> Int
_dataPars :: Nat
, DatatypeData -> Int
_dataIxs :: Nat
, DatatypeData -> Maybe Clause
_dataClause :: Maybe Clause
, DatatypeData -> [QName]
_dataCons :: [QName]
, DatatypeData -> Sort
_dataSort :: Sort
, DatatypeData -> Maybe [QName]
_dataMutual :: Maybe [QName]
, DatatypeData -> IsAbstract
_dataAbstr :: IsAbstract
, DatatypeData -> [QName]
_dataPathCons :: [QName]
, DatatypeData -> Maybe QName
_dataTranspIx :: Maybe QName
, DatatypeData -> Maybe QName
_dataTransp :: Maybe QName
} deriving (Int -> DatatypeData -> ShowS
[DatatypeData] -> ShowS
DatatypeData -> String
(Int -> DatatypeData -> ShowS)
-> (DatatypeData -> String)
-> ([DatatypeData] -> ShowS)
-> Show DatatypeData
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> DatatypeData -> ShowS
showsPrec :: Int -> DatatypeData -> ShowS
$cshow :: DatatypeData -> String
show :: DatatypeData -> String
$cshowList :: [DatatypeData] -> ShowS
showList :: [DatatypeData] -> ShowS
Show, (forall x. DatatypeData -> Rep DatatypeData x)
-> (forall x. Rep DatatypeData x -> DatatypeData)
-> Generic DatatypeData
forall x. Rep DatatypeData x -> DatatypeData
forall x. DatatypeData -> Rep DatatypeData x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. DatatypeData -> Rep DatatypeData x
from :: forall x. DatatypeData -> Rep DatatypeData x
$cto :: forall x. Rep DatatypeData x -> DatatypeData
to :: forall x. Rep DatatypeData x -> DatatypeData
Generic)
pattern Datatype
:: Nat
-> Nat
-> (Maybe Clause)
-> [QName]
-> Sort
-> Maybe [QName]
-> IsAbstract
-> [QName]
-> Maybe QName
-> Maybe QName
-> Defn
pattern $mDatatype :: forall {r}.
Defn
-> (Int
-> Int
-> Maybe Clause
-> [QName]
-> Sort
-> Maybe [QName]
-> IsAbstract
-> [QName]
-> Maybe QName
-> Maybe QName
-> r)
-> ((# #) -> r)
-> r
$bDatatype :: Int
-> Int
-> Maybe Clause
-> [QName]
-> Sort
-> Maybe [QName]
-> IsAbstract
-> [QName]
-> Maybe QName
-> Maybe QName
-> Defn
Datatype
{ Defn -> Int
dataPars
, Defn -> Int
dataIxs
, Defn -> Maybe Clause
dataClause
, Defn -> [QName]
dataCons
, Defn -> Sort
dataSort
, Defn -> Maybe [QName]
dataMutual
, Defn -> IsAbstract
dataAbstr
, Defn -> [QName]
dataPathCons
, Defn -> Maybe QName
dataTranspIx
, Defn -> Maybe QName
dataTransp
} = DatatypeDefn (DatatypeData
dataPars
dataIxs
dataClause
dataCons
dataSort
dataMutual
dataAbstr
dataPathCons
dataTranspIx
dataTransp
)
data RecordData = RecordData
{ RecordData -> Int
_recPars :: Nat
, RecordData -> Maybe Clause
_recClause :: Maybe Clause
, RecordData -> ConHead
_recConHead :: ConHead
, RecordData -> Bool
_recNamedCon :: Bool
, RecordData -> [Dom QName]
_recFields :: [Dom QName]
, RecordData -> Telescope
_recTel :: Telescope
, RecordData -> Maybe [QName]
_recMutual :: Maybe [QName]
, RecordData -> EtaEquality
_recEtaEquality' :: EtaEquality
, RecordData -> PatternOrCopattern
_recPatternMatching :: PatternOrCopattern
, RecordData -> Maybe Induction
_recInduction :: Maybe Induction
, RecordData -> Maybe Bool
_recTerminates :: Maybe Bool
, RecordData -> IsAbstract
_recAbstr :: IsAbstract
, RecordData -> CompKit
_recComp :: CompKit
} deriving (Int -> RecordData -> ShowS
[RecordData] -> ShowS
RecordData -> String
(Int -> RecordData -> ShowS)
-> (RecordData -> String)
-> ([RecordData] -> ShowS)
-> Show RecordData
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> RecordData -> ShowS
showsPrec :: Int -> RecordData -> ShowS
$cshow :: RecordData -> String
show :: RecordData -> String
$cshowList :: [RecordData] -> ShowS
showList :: [RecordData] -> ShowS
Show, (forall x. RecordData -> Rep RecordData x)
-> (forall x. Rep RecordData x -> RecordData) -> Generic RecordData
forall x. Rep RecordData x -> RecordData
forall x. RecordData -> Rep RecordData x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. RecordData -> Rep RecordData x
from :: forall x. RecordData -> Rep RecordData x
$cto :: forall x. Rep RecordData x -> RecordData
to :: forall x. Rep RecordData x -> RecordData
Generic)
pattern Record
:: Nat
-> Maybe Clause
-> ConHead
-> Bool
-> [Dom QName]
-> Telescope
-> Maybe [QName]
-> EtaEquality
-> PatternOrCopattern
-> Maybe Induction
-> Maybe Bool
-> IsAbstract
-> CompKit
-> Defn
pattern $mRecord :: forall {r}.
Defn
-> (Int
-> Maybe Clause
-> ConHead
-> Bool
-> [Dom QName]
-> Telescope
-> Maybe [QName]
-> EtaEquality
-> PatternOrCopattern
-> Maybe Induction
-> Maybe Bool
-> IsAbstract
-> CompKit
-> r)
-> ((# #) -> r)
-> r
$bRecord :: Int
-> Maybe Clause
-> ConHead
-> Bool
-> [Dom QName]
-> Telescope
-> Maybe [QName]
-> EtaEquality
-> PatternOrCopattern
-> Maybe Induction
-> Maybe Bool
-> IsAbstract
-> CompKit
-> Defn
Record
{ Defn -> Int
recPars
, Defn -> Maybe Clause
recClause
, Defn -> ConHead
recConHead
, Defn -> Bool
recNamedCon
, Defn -> [Dom QName]
recFields
, Defn -> Telescope
recTel
, Defn -> Maybe [QName]
recMutual
, Defn -> EtaEquality
recEtaEquality'
, Defn -> PatternOrCopattern
recPatternMatching
, Defn -> Maybe Induction
recInduction
, Defn -> Maybe Bool
recTerminates
, Defn -> IsAbstract
recAbstr
, Defn -> CompKit
recComp
} = RecordDefn (RecordData
recPars
recClause
recConHead
recNamedCon
recFields
recTel
recMutual
recEtaEquality'
recPatternMatching
recInduction
recTerminates
recAbstr
recComp
)
data ConstructorData = ConstructorData
{ ConstructorData -> Int
_conPars :: Int
, ConstructorData -> Int
_conArity :: Int
, ConstructorData -> ConHead
_conSrcCon :: ConHead
, ConstructorData -> QName
_conData :: QName
, ConstructorData -> IsAbstract
_conAbstr :: IsAbstract
, ConstructorData -> CompKit
_conComp :: CompKit
, ConstructorData -> Maybe [QName]
_conProj :: Maybe [QName]
, ConstructorData -> [IsForced]
_conForced :: [IsForced]
, ConstructorData -> Maybe [Bool]
_conErased :: Maybe [Bool]
, ConstructorData -> Bool
_conErasure :: !Bool
, ConstructorData -> Bool
_conInline :: !Bool
} deriving (Int -> ConstructorData -> ShowS
[ConstructorData] -> ShowS
ConstructorData -> String
(Int -> ConstructorData -> ShowS)
-> (ConstructorData -> String)
-> ([ConstructorData] -> ShowS)
-> Show ConstructorData
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> ConstructorData -> ShowS
showsPrec :: Int -> ConstructorData -> ShowS
$cshow :: ConstructorData -> String
show :: ConstructorData -> String
$cshowList :: [ConstructorData] -> ShowS
showList :: [ConstructorData] -> ShowS
Show, (forall x. ConstructorData -> Rep ConstructorData x)
-> (forall x. Rep ConstructorData x -> ConstructorData)
-> Generic ConstructorData
forall x. Rep ConstructorData x -> ConstructorData
forall x. ConstructorData -> Rep ConstructorData x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. ConstructorData -> Rep ConstructorData x
from :: forall x. ConstructorData -> Rep ConstructorData x
$cto :: forall x. Rep ConstructorData x -> ConstructorData
to :: forall x. Rep ConstructorData x -> ConstructorData
Generic)
pattern Constructor
:: Int
-> Int
-> ConHead
-> QName
-> IsAbstract
-> CompKit
-> Maybe [QName]
-> [IsForced]
-> Maybe [Bool]
-> Bool
-> Bool
-> Defn
pattern $mConstructor :: forall {r}.
Defn
-> (Int
-> Int
-> ConHead
-> QName
-> IsAbstract
-> CompKit
-> Maybe [QName]
-> [IsForced]
-> Maybe [Bool]
-> Bool
-> Bool
-> r)
-> ((# #) -> r)
-> r
$bConstructor :: Int
-> Int
-> ConHead
-> QName
-> IsAbstract
-> CompKit
-> Maybe [QName]
-> [IsForced]
-> Maybe [Bool]
-> Bool
-> Bool
-> Defn
Constructor
{ Defn -> Int
conPars
, Defn -> Int
conArity
, Defn -> ConHead
conSrcCon
, Defn -> QName
conData
, Defn -> IsAbstract
conAbstr
, Defn -> CompKit
conComp
, Defn -> Maybe [QName]
conProj
, Defn -> [IsForced]
conForced
, Defn -> Maybe [Bool]
conErased
, Defn -> Bool
conErasure
, Defn -> Bool
conInline
} = ConstructorDefn (ConstructorData
conPars
conArity
conSrcCon
conData
conAbstr
conComp
conProj
conForced
conErased
conErasure
conInline
)
data PrimitiveData = PrimitiveData
{ PrimitiveData -> IsAbstract
_primAbstr :: IsAbstract
, PrimitiveData -> PrimitiveId
_primName :: PrimitiveId
, PrimitiveData -> [Clause]
_primClauses :: [Clause]
, PrimitiveData -> FunctionInverse
_primInv :: FunctionInverse
, PrimitiveData -> Maybe CompiledClauses
_primCompiled :: Maybe CompiledClauses
, PrimitiveData -> IsOpaque
_primOpaque :: IsOpaque
} deriving (Int -> PrimitiveData -> ShowS
[PrimitiveData] -> ShowS
PrimitiveData -> String
(Int -> PrimitiveData -> ShowS)
-> (PrimitiveData -> String)
-> ([PrimitiveData] -> ShowS)
-> Show PrimitiveData
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> PrimitiveData -> ShowS
showsPrec :: Int -> PrimitiveData -> ShowS
$cshow :: PrimitiveData -> String
show :: PrimitiveData -> String
$cshowList :: [PrimitiveData] -> ShowS
showList :: [PrimitiveData] -> ShowS
Show, (forall x. PrimitiveData -> Rep PrimitiveData x)
-> (forall x. Rep PrimitiveData x -> PrimitiveData)
-> Generic PrimitiveData
forall x. Rep PrimitiveData x -> PrimitiveData
forall x. PrimitiveData -> Rep PrimitiveData x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. PrimitiveData -> Rep PrimitiveData x
from :: forall x. PrimitiveData -> Rep PrimitiveData x
$cto :: forall x. Rep PrimitiveData x -> PrimitiveData
to :: forall x. Rep PrimitiveData x -> PrimitiveData
Generic)
pattern Primitive
:: IsAbstract
-> PrimitiveId
-> [Clause]
-> FunctionInverse
-> Maybe CompiledClauses
-> IsOpaque
-> Defn
pattern $mPrimitive :: forall {r}.
Defn
-> (IsAbstract
-> PrimitiveId
-> [Clause]
-> FunctionInverse
-> Maybe CompiledClauses
-> IsOpaque
-> r)
-> ((# #) -> r)
-> r
$bPrimitive :: IsAbstract
-> PrimitiveId
-> [Clause]
-> FunctionInverse
-> Maybe CompiledClauses
-> IsOpaque
-> Defn
Primitive
{ Defn -> IsAbstract
primAbstr
, Defn -> PrimitiveId
primName
, Defn -> [Clause]
primClauses
, Defn -> FunctionInverse
primInv
, Defn -> Maybe CompiledClauses
primCompiled
, Defn -> IsOpaque
primOpaque
} = PrimitiveDefn (PrimitiveData
primAbstr
primName
primClauses
primInv
primCompiled
primOpaque
)
data PrimitiveSortData = PrimitiveSortData
{ PrimitiveSortData -> BuiltinSort
_primSortName :: BuiltinSort
, PrimitiveSortData -> Sort
_primSortSort :: Sort
} deriving (Int -> PrimitiveSortData -> ShowS
[PrimitiveSortData] -> ShowS
PrimitiveSortData -> String
(Int -> PrimitiveSortData -> ShowS)
-> (PrimitiveSortData -> String)
-> ([PrimitiveSortData] -> ShowS)
-> Show PrimitiveSortData
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> PrimitiveSortData -> ShowS
showsPrec :: Int -> PrimitiveSortData -> ShowS
$cshow :: PrimitiveSortData -> String
show :: PrimitiveSortData -> String
$cshowList :: [PrimitiveSortData] -> ShowS
showList :: [PrimitiveSortData] -> ShowS
Show, (forall x. PrimitiveSortData -> Rep PrimitiveSortData x)
-> (forall x. Rep PrimitiveSortData x -> PrimitiveSortData)
-> Generic PrimitiveSortData
forall x. Rep PrimitiveSortData x -> PrimitiveSortData
forall x. PrimitiveSortData -> Rep PrimitiveSortData x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. PrimitiveSortData -> Rep PrimitiveSortData x
from :: forall x. PrimitiveSortData -> Rep PrimitiveSortData x
$cto :: forall x. Rep PrimitiveSortData x -> PrimitiveSortData
to :: forall x. Rep PrimitiveSortData x -> PrimitiveSortData
Generic)
pattern PrimitiveSort
:: BuiltinSort
-> Sort
-> Defn
pattern $mPrimitiveSort :: forall {r}. Defn -> (BuiltinSort -> Sort -> r) -> ((# #) -> r) -> r
$bPrimitiveSort :: BuiltinSort -> Sort -> Defn
PrimitiveSort
{ Defn -> BuiltinSort
primSortName
, Defn -> Sort
primSortSort
} = PrimitiveSortDefn (PrimitiveSortData
primSortName
primSortSort
)
lensFunction :: Lens' Defn FunctionData
lensFunction :: Lens' Defn FunctionData
lensFunction FunctionData -> f FunctionData
f = \case
FunctionDefn FunctionData
d -> FunctionData -> Defn
FunctionDefn (FunctionData -> Defn) -> f FunctionData -> f Defn
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> FunctionData -> f FunctionData
f FunctionData
d
Defn
_ -> f Defn
forall a. HasCallStack => a
__IMPOSSIBLE__
lensConstructor :: Lens' Defn ConstructorData
lensConstructor :: Lens' Defn ConstructorData
lensConstructor ConstructorData -> f ConstructorData
f = \case
ConstructorDefn ConstructorData
d -> ConstructorData -> Defn
ConstructorDefn (ConstructorData -> Defn) -> f ConstructorData -> f Defn
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> ConstructorData -> f ConstructorData
f ConstructorData
d
Defn
_ -> f Defn
forall a. HasCallStack => a
__IMPOSSIBLE__
lensRecord :: Lens' Defn RecordData
lensRecord :: Lens' Defn RecordData
lensRecord RecordData -> f RecordData
f = \case
RecordDefn RecordData
d -> RecordData -> Defn
RecordDefn (RecordData -> Defn) -> f RecordData -> f Defn
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> RecordData -> f RecordData
f RecordData
d
Defn
_ -> f Defn
forall a. HasCallStack => a
__IMPOSSIBLE__
lensRecTel :: Lens' RecordData Telescope
lensRecTel :: Lens' RecordData Telescope
lensRecTel Telescope -> f Telescope
f RecordData
r =
Telescope -> f Telescope
f (RecordData -> Telescope
_recTel RecordData
r) f Telescope -> (Telescope -> RecordData) -> f RecordData
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ Telescope
tel -> RecordData
r { _recTel = tel }
lensRecEta :: Lens' RecordData EtaEquality
lensRecEta :: Lens' RecordData EtaEquality
lensRecEta EtaEquality -> f EtaEquality
f RecordData
r =
EtaEquality -> f EtaEquality
f (RecordData -> EtaEquality
_recEtaEquality' RecordData
r) f EtaEquality -> (EtaEquality -> RecordData) -> f RecordData
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ EtaEquality
eta -> RecordData
r { _recEtaEquality' = eta }
instance Pretty Definition where
pretty :: Definition -> Doc
pretty Defn{Bool
[Maybe Name]
[Occurrence]
[Polarity]
[LocalDisplayForm]
Maybe InstanceInfo
CompiledRepresentation
Set QName
ArgInfo
Language
QName
Blocked_
Type
MutualId
Defn
defArgInfo :: Definition -> ArgInfo
defName :: Definition -> QName
defType :: Definition -> Type
defPolarity :: Definition -> [Polarity]
defArgOccurrences :: Definition -> [Occurrence]
defGeneralizedParams :: Definition -> [Maybe Name]
defDisplay :: Definition -> [LocalDisplayForm]
defMutual :: Definition -> MutualId
defCompiledRep :: Definition -> CompiledRepresentation
defInstance :: Definition -> Maybe InstanceInfo
defCopy :: Definition -> Bool
defMatchable :: Definition -> Set QName
defNoCompilation :: Definition -> Bool
defInjective :: Definition -> Bool
defCopatternLHS :: Definition -> Bool
defBlocked :: Definition -> Blocked_
defLanguage :: Definition -> Language
theDef :: Definition -> Defn
defArgInfo :: ArgInfo
defName :: QName
defType :: Type
defPolarity :: [Polarity]
defArgOccurrences :: [Occurrence]
defGeneralizedParams :: [Maybe Name]
defDisplay :: [LocalDisplayForm]
defMutual :: MutualId
defCompiledRep :: CompiledRepresentation
defInstance :: Maybe InstanceInfo
defCopy :: Bool
defMatchable :: Set QName
defNoCompilation :: Bool
defInjective :: Bool
defCopatternLHS :: Bool
defBlocked :: Blocked_
defLanguage :: Language
theDef :: Defn
..} =
Doc
"Defn {" Doc -> Doc -> Doc
<?> [Doc] -> Doc
forall (t :: * -> *). Foldable t => t Doc -> Doc
vcat
[ Doc
"defArgInfo =" Doc -> Doc -> Doc
<?> ArgInfo -> Doc
forall a. Show a => a -> Doc
pshow ArgInfo
defArgInfo
, Doc
"defName =" Doc -> Doc -> Doc
<?> QName -> Doc
forall a. Pretty a => a -> Doc
pretty QName
defName
, Doc
"defType =" Doc -> Doc -> Doc
<?> Type -> Doc
forall a. Pretty a => a -> Doc
pretty Type
defType
, Doc
"defPolarity =" Doc -> Doc -> Doc
<?> [Polarity] -> Doc
forall a. Show a => a -> Doc
pshow [Polarity]
defPolarity
, Doc
"defArgOccurrences =" Doc -> Doc -> Doc
<?> [Occurrence] -> Doc
forall a. Show a => a -> Doc
pshow [Occurrence]
defArgOccurrences
, Doc
"defGeneralizedParams =" Doc -> Doc -> Doc
<?> [Maybe Name] -> Doc
forall a. Show a => a -> Doc
pshow [Maybe Name]
defGeneralizedParams
, Doc
"defDisplay =" Doc -> Doc -> Doc
<?> [LocalDisplayForm] -> Doc
forall a. Pretty a => a -> Doc
pretty [LocalDisplayForm]
defDisplay
, Doc
"defMutual =" Doc -> Doc -> Doc
<?> MutualId -> Doc
forall a. Show a => a -> Doc
pshow MutualId
defMutual
, Doc
"defCompiledRep =" Doc -> Doc -> Doc
<?> CompiledRepresentation -> Doc
forall a. Show a => a -> Doc
pshow CompiledRepresentation
defCompiledRep
, Doc
"defInstance =" Doc -> Doc -> Doc
<?> Maybe InstanceInfo -> Doc
forall a. Show a => a -> Doc
pshow Maybe InstanceInfo
defInstance
, Doc
"defCopy =" Doc -> Doc -> Doc
<?> Bool -> Doc
forall a. Show a => a -> Doc
pshow Bool
defCopy
, Doc
"defMatchable =" Doc -> Doc -> Doc
<?> [QName] -> Doc
forall a. Show a => a -> Doc
pshow (Set QName -> [QName]
forall a. Set a -> [a]
Set.toList Set QName
defMatchable)
, Doc
"defInjective =" Doc -> Doc -> Doc
<?> Bool -> Doc
forall a. Show a => a -> Doc
pshow Bool
defInjective
, Doc
"defCopatternLHS =" Doc -> Doc -> Doc
<?> Bool -> Doc
forall a. Show a => a -> Doc
pshow Bool
defCopatternLHS
, Doc
"theDef =" Doc -> Doc -> Doc
<?> Defn -> Doc
forall a. Pretty a => a -> Doc
pretty Defn
theDef ] Doc -> Doc -> Doc
forall a. Doc a -> Doc a -> Doc a
<+> Doc
"}"
instance Pretty Defn where
pretty :: Defn -> Doc
pretty = \case
AxiomDefn AxiomData
_ -> Doc
"Axiom"
DataOrRecSigDefn DataOrRecSigData
d -> DataOrRecSigData -> Doc
forall a. Pretty a => a -> Doc
pretty DataOrRecSigData
d
GeneralizableVar NumGeneralizableArgs
_ -> Doc
"GeneralizableVar"
AbstractDefn Defn
def -> Doc
"AbstractDefn" Doc -> Doc -> Doc
<?> Doc -> Doc
parens (Defn -> Doc
forall a. Pretty a => a -> Doc
pretty Defn
def)
FunctionDefn FunctionData
d -> FunctionData -> Doc
forall a. Pretty a => a -> Doc
pretty FunctionData
d
DatatypeDefn DatatypeData
d -> DatatypeData -> Doc
forall a. Pretty a => a -> Doc
pretty DatatypeData
d
RecordDefn RecordData
d -> RecordData -> Doc
forall a. Pretty a => a -> Doc
pretty RecordData
d
ConstructorDefn ConstructorData
d -> ConstructorData -> Doc
forall a. Pretty a => a -> Doc
pretty ConstructorData
d
PrimitiveDefn PrimitiveData
d -> PrimitiveData -> Doc
forall a. Pretty a => a -> Doc
pretty PrimitiveData
d
PrimitiveSortDefn PrimitiveSortData
d -> PrimitiveSortData -> Doc
forall a. Pretty a => a -> Doc
pretty PrimitiveSortData
d
instance Pretty DataOrRecSigData where
pretty :: DataOrRecSigData -> Doc
pretty (DataOrRecSigData Int
n) = Doc
"DataOrRecSig" Doc -> Doc -> Doc
forall a. Doc a -> Doc a -> Doc a
<+> Int -> Doc
forall a. Pretty a => a -> Doc
pretty Int
n
instance Pretty ProjectionLikenessMissing where
pretty :: ProjectionLikenessMissing -> Doc
pretty ProjectionLikenessMissing
MaybeProjection = Doc
"MaybeProjection"
pretty ProjectionLikenessMissing
NeverProjection = Doc
"NeverProjection"
instance Pretty FunctionData where
pretty :: FunctionData -> Doc
pretty (FunctionData
[Clause]
funClauses
Maybe CompiledClauses
funCompiled
Maybe SplitTree
funSplitTree
Maybe Compiled
funTreeless
[Clause]
_funCovering
FunctionInverse
funInv
Maybe [QName]
funMutual
Either ProjectionLikenessMissing Projection
funProjection
SmallSet FunctionFlag
funFlags
Maybe Bool
funTerminates
Maybe ExtLamInfo
_funExtLam
Maybe QName
funWith
Maybe QName
funIsKanOp
IsOpaque
funOpaque
) =
Doc
"Function {" Doc -> Doc -> Doc
<?> [Doc] -> Doc
forall (t :: * -> *). Foldable t => t Doc -> Doc
vcat
[ Doc
"funClauses =" Doc -> Doc -> Doc
<?> [Doc] -> Doc
forall (t :: * -> *). Foldable t => t Doc -> Doc
vcat ((Clause -> Doc) -> [Clause] -> [Doc]
forall a b. (a -> b) -> [a] -> [b]
map Clause -> Doc
forall a. Pretty a => a -> Doc
pretty [Clause]
funClauses)
, Doc
"funCompiled =" Doc -> Doc -> Doc
<?> Maybe CompiledClauses -> Doc
forall a. Pretty a => a -> Doc
pretty Maybe CompiledClauses
funCompiled
, Doc
"funSplitTree =" Doc -> Doc -> Doc
<?> Maybe SplitTree -> Doc
forall a. Pretty a => a -> Doc
pretty Maybe SplitTree
funSplitTree
, Doc
"funTreeless =" Doc -> Doc -> Doc
<?> Maybe Compiled -> Doc
forall a. Pretty a => a -> Doc
pretty Maybe Compiled
funTreeless
, Doc
"funInv =" Doc -> Doc -> Doc
<?> FunctionInverse -> Doc
forall a. Pretty a => a -> Doc
pretty FunctionInverse
funInv
, Doc
"funMutual =" Doc -> Doc -> Doc
<?> Maybe [QName] -> Doc
forall a. Show a => a -> Doc
pshow Maybe [QName]
funMutual
, Doc
"funProjection =" Doc -> Doc -> Doc
<?> Either ProjectionLikenessMissing Projection -> Doc
forall a. Pretty a => a -> Doc
pretty Either ProjectionLikenessMissing Projection
funProjection
, Doc
"funFlags =" Doc -> Doc -> Doc
<?> SmallSet FunctionFlag -> Doc
forall a. Show a => a -> Doc
pshow SmallSet FunctionFlag
funFlags
, Doc
"funTerminates =" Doc -> Doc -> Doc
<?> Maybe Bool -> Doc
forall a. Show a => a -> Doc
pshow Maybe Bool
funTerminates
, Doc
"funWith =" Doc -> Doc -> Doc
<?> Maybe QName -> Doc
forall a. Pretty a => a -> Doc
pretty Maybe QName
funWith
, Doc
"funIsKanOp =" Doc -> Doc -> Doc
<?> Maybe QName -> Doc
forall a. Pretty a => a -> Doc
pretty Maybe QName
funIsKanOp
, Doc
"funOpaque =" Doc -> Doc -> Doc
<?> IsOpaque -> Doc
forall a. Show a => a -> Doc
pshow IsOpaque
funOpaque
] Doc -> Doc -> Doc
<?> Doc
"}"
instance Pretty DatatypeData where
pretty :: DatatypeData -> Doc
pretty (DatatypeData
Int
dataPars
Int
dataIxs
Maybe Clause
dataClause
[QName]
dataCons
Sort
dataSort
Maybe [QName]
dataMutual
IsAbstract
_dataAbstr
[QName]
_dataPathCons
Maybe QName
_dataTranspIx
Maybe QName
_dataTransp
) =
Doc
"Datatype {" Doc -> Doc -> Doc
<?> [Doc] -> Doc
forall (t :: * -> *). Foldable t => t Doc -> Doc
vcat
[ Doc
"dataPars =" Doc -> Doc -> Doc
<?> Int -> Doc
forall a. Show a => a -> Doc
pshow Int
dataPars
, Doc
"dataIxs =" Doc -> Doc -> Doc
<?> Int -> Doc
forall a. Show a => a -> Doc
pshow Int
dataIxs
, Doc
"dataClause =" Doc -> Doc -> Doc
<?> Maybe Clause -> Doc
forall a. Pretty a => a -> Doc
pretty Maybe Clause
dataClause
, Doc
"dataCons =" Doc -> Doc -> Doc
<?> [QName] -> Doc
forall a. Show a => a -> Doc
pshow [QName]
dataCons
, Doc
"dataSort =" Doc -> Doc -> Doc
<?> Sort -> Doc
forall a. Pretty a => a -> Doc
pretty Sort
dataSort
, Doc
"dataMutual =" Doc -> Doc -> Doc
<?> Maybe [QName] -> Doc
forall a. Show a => a -> Doc
pshow Maybe [QName]
dataMutual
, Doc
"dataAbstr =" Doc -> Doc -> Doc
<?> (Defn -> IsAbstract) -> Doc
forall a. Show a => a -> Doc
pshow Defn -> IsAbstract
dataAbstr
] Doc -> Doc -> Doc
<?> Doc
"}"
instance Pretty RecordData where
pretty :: RecordData -> Doc
pretty (RecordData
Int
recPars
Maybe Clause
recClause
ConHead
recConHead
Bool
recNamedCon
[Dom QName]
recFields
Telescope
recTel
Maybe [QName]
recMutual
EtaEquality
recEtaEquality'
PatternOrCopattern
_recPatternMatching
Maybe Induction
recInduction
Maybe Bool
_recTerminates
IsAbstract
recAbstr
CompKit
_recComp
) =
Doc
"Record {" Doc -> Doc -> Doc
<?> [Doc] -> Doc
forall (t :: * -> *). Foldable t => t Doc -> Doc
vcat
[ Doc
"recPars =" Doc -> Doc -> Doc
<?> Int -> Doc
forall a. Show a => a -> Doc
pshow Int
recPars
, Doc
"recClause =" Doc -> Doc -> Doc
<?> Maybe Clause -> Doc
forall a. Pretty a => a -> Doc
pretty Maybe Clause
recClause
, Doc
"recConHead =" Doc -> Doc -> Doc
<?> ConHead -> Doc
forall a. Pretty a => a -> Doc
pretty ConHead
recConHead
, Doc
"recNamedCon =" Doc -> Doc -> Doc
<?> Bool -> Doc
forall a. Pretty a => a -> Doc
pretty Bool
recNamedCon
, Doc
"recFields =" Doc -> Doc -> Doc
<?> [Dom QName] -> Doc
forall a. Pretty a => a -> Doc
pretty [Dom QName]
recFields
, Doc
"recTel =" Doc -> Doc -> Doc
<?> Telescope -> Doc
forall a. Pretty a => a -> Doc
pretty Telescope
recTel
, Doc
"recMutual =" Doc -> Doc -> Doc
<?> Maybe [QName] -> Doc
forall a. Show a => a -> Doc
pshow Maybe [QName]
recMutual
, Doc
"recEtaEquality' =" Doc -> Doc -> Doc
<?> EtaEquality -> Doc
forall a. Show a => a -> Doc
pshow EtaEquality
recEtaEquality'
, Doc
"recInduction =" Doc -> Doc -> Doc
<?> Maybe Induction -> Doc
forall a. Show a => a -> Doc
pshow Maybe Induction
recInduction
, Doc
"recAbstr =" Doc -> Doc -> Doc
<?> IsAbstract -> Doc
forall a. Show a => a -> Doc
pshow IsAbstract
recAbstr
] Doc -> Doc -> Doc
<?> Doc
"}"
instance Pretty ConstructorData where
pretty :: ConstructorData -> Doc
pretty (ConstructorData
Int
conPars
Int
conArity
ConHead
conSrcCon
QName
conData
IsAbstract
conAbstr
CompKit
_conComp
Maybe [QName]
_conProj
[IsForced]
_conForced
Maybe [Bool]
conErased
Bool
conErasure
Bool
conInline
) =
Doc
"Constructor {" Doc -> Doc -> Doc
<?> [Doc] -> Doc
forall (t :: * -> *). Foldable t => t Doc -> Doc
vcat
[ Doc
"conPars =" Doc -> Doc -> Doc
<?> Int -> Doc
forall a. Show a => a -> Doc
pshow Int
conPars
, Doc
"conArity =" Doc -> Doc -> Doc
<?> Int -> Doc
forall a. Show a => a -> Doc
pshow Int
conArity
, Doc
"conSrcCon =" Doc -> Doc -> Doc
<?> ConHead -> Doc
forall a. Pretty a => a -> Doc
pretty ConHead
conSrcCon
, Doc
"conData =" Doc -> Doc -> Doc
<?> QName -> Doc
forall a. Pretty a => a -> Doc
pretty QName
conData
, Doc
"conAbstr =" Doc -> Doc -> Doc
<?> IsAbstract -> Doc
forall a. Show a => a -> Doc
pshow IsAbstract
conAbstr
, Doc
"conErased =" Doc -> Doc -> Doc
<?> Maybe [Bool] -> Doc
forall a. Show a => a -> Doc
pshow Maybe [Bool]
conErased
, Doc
"conErasure =" Doc -> Doc -> Doc
<?> Bool -> Doc
forall a. Show a => a -> Doc
pshow Bool
conErasure
, Doc
"conInline =" Doc -> Doc -> Doc
<?> Bool -> Doc
forall a. Show a => a -> Doc
pshow Bool
conInline
] Doc -> Doc -> Doc
<?> Doc
"}"
instance Pretty PrimitiveData where
pretty :: PrimitiveData -> Doc
pretty (PrimitiveData
IsAbstract
primAbstr
PrimitiveId
primName
[Clause]
primClauses
FunctionInverse
_primInv
Maybe CompiledClauses
primCompiled
IsOpaque
primOpaque
) =
Doc
"Primitive {" Doc -> Doc -> Doc
<?> [Doc] -> Doc
forall (t :: * -> *). Foldable t => t Doc -> Doc
vcat
[ Doc
"primAbstr =" Doc -> Doc -> Doc
<?> IsAbstract -> Doc
forall a. Show a => a -> Doc
pshow IsAbstract
primAbstr
, Doc
"primName =" Doc -> Doc -> Doc
<?> PrimitiveId -> Doc
forall a. Show a => a -> Doc
pshow PrimitiveId
primName
, Doc
"primClauses =" Doc -> Doc -> Doc
<?> [Clause] -> Doc
forall a. Show a => a -> Doc
pshow [Clause]
primClauses
, Doc
"primCompiled =" Doc -> Doc -> Doc
<?> Maybe CompiledClauses -> Doc
forall a. Show a => a -> Doc
pshow Maybe CompiledClauses
primCompiled
, Doc
"primOpaque =" Doc -> Doc -> Doc
<?> IsOpaque -> Doc
forall a. Show a => a -> Doc
pshow IsOpaque
primOpaque
] Doc -> Doc -> Doc
<?> Doc
"}"
instance Pretty PrimitiveSortData where
pretty :: PrimitiveSortData -> Doc
pretty (PrimitiveSortData BuiltinSort
primSortName Sort
primSortSort) =
Doc
"PrimitiveSort {" Doc -> Doc -> Doc
<?> [Doc] -> Doc
forall (t :: * -> *). Foldable t => t Doc -> Doc
vcat
[ Doc
"primSortName =" Doc -> Doc -> Doc
<?> BuiltinSort -> Doc
forall a. Show a => a -> Doc
pshow BuiltinSort
primSortName
, Doc
"primSortSort =" Doc -> Doc -> Doc
<?> Sort -> Doc
forall a. Show a => a -> Doc
pshow Sort
primSortSort
] Doc -> Doc -> Doc
<?> Doc
"}"
instance Pretty Projection where
pretty :: Projection -> Doc
pretty Projection{Int
Maybe QName
Arg QName
QName
ProjLams
projProper :: Projection -> Maybe QName
projOrig :: Projection -> QName
projFromType :: Projection -> Arg QName
projIndex :: Projection -> Int
projLams :: Projection -> ProjLams
projProper :: Maybe QName
projOrig :: QName
projFromType :: Arg QName
projIndex :: Int
projLams :: ProjLams
..} =
Doc
"Projection {" Doc -> Doc -> Doc
<?> [Doc] -> Doc
forall (t :: * -> *). Foldable t => t Doc -> Doc
vcat
[ Doc
"projProper =" Doc -> Doc -> Doc
<?> Maybe QName -> Doc
forall a. Pretty a => a -> Doc
pretty Maybe QName
projProper
, Doc
"projOrig =" Doc -> Doc -> Doc
<?> QName -> Doc
forall a. Pretty a => a -> Doc
pretty QName
projOrig
, Doc
"projFromType =" Doc -> Doc -> Doc
<?> Arg QName -> Doc
forall a. Pretty a => a -> Doc
pretty Arg QName
projFromType
, Doc
"projIndex =" Doc -> Doc -> Doc
<?> Int -> Doc
forall a. Show a => a -> Doc
pshow Int
projIndex
, Doc
"projLams =" Doc -> Doc -> Doc
<?> ProjLams -> Doc
forall a. Pretty a => a -> Doc
pretty ProjLams
projLams
]
instance Pretty c => Pretty (FunctionInverse' c) where
pretty :: FunctionInverse' c -> Doc
pretty FunctionInverse' c
NotInjective = Doc
"NotInjective"
pretty (Inverse InversionMap c
inv) = Doc
"Inverse" Doc -> Doc -> Doc
<?>
[Doc] -> Doc
forall (t :: * -> *). Foldable t => t Doc -> Doc
vcat [ TermHead -> Doc
forall a. Pretty a => a -> Doc
pretty TermHead
h Doc -> Doc -> Doc
forall a. Doc a -> Doc a -> Doc a
<+> Doc
"->" Doc -> Doc -> Doc
<?> [c] -> Doc
forall a. Pretty a => a -> Doc
pretty [c]
cs
| (TermHead
h, [c]
cs) <- InversionMap c -> [(TermHead, [c])]
forall k a. Map k a -> [(k, a)]
Map.toList InversionMap c
inv ]
instance Pretty ProjLams where
pretty :: ProjLams -> Doc
pretty (ProjLams [Arg String]
args) = [Arg String] -> Doc
forall a. Pretty a => a -> Doc
pretty [Arg String]
args
recRecursive :: Defn -> Bool
recRecursive :: Defn -> Bool
recRecursive (RecordDefn RecordData
d) = RecordData -> Bool
recRecursive_ RecordData
d
recRecursive Defn
_ = Bool
forall a. HasCallStack => a
__IMPOSSIBLE__
recRecursive_ :: RecordData -> Bool
recRecursive_ :: RecordData -> Bool
recRecursive_ RecordData{ _recMutual :: RecordData -> Maybe [QName]
_recMutual = Just [QName]
qs } = Bool -> Bool
not (Bool -> Bool) -> Bool -> Bool
forall a b. (a -> b) -> a -> b
$ [QName] -> Bool
forall a. Null a => a -> Bool
null [QName]
qs
recRecursive_ RecordData
_ = Bool
forall a. HasCallStack => a
__IMPOSSIBLE__
recEtaEquality :: Defn -> HasEta
recEtaEquality :: Defn -> HasEta
recEtaEquality = EtaEquality -> HasEta
theEtaEquality (EtaEquality -> HasEta) -> (Defn -> EtaEquality) -> Defn -> HasEta
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Defn -> EtaEquality
recEtaEquality'
_recEtaEquality :: RecordData -> HasEta
_recEtaEquality :: RecordData -> HasEta
_recEtaEquality = EtaEquality -> HasEta
theEtaEquality (EtaEquality -> HasEta)
-> (RecordData -> EtaEquality) -> RecordData -> HasEta
forall b c a. (b -> c) -> (a -> b) -> a -> c
. RecordData -> EtaEquality
_recEtaEquality'
emptyFunctionData :: HasOptions m => m FunctionData
emptyFunctionData :: forall (m :: * -> *). HasOptions m => m FunctionData
emptyFunctionData = Bool -> FunctionData
emptyFunctionData_ (Bool -> FunctionData)
-> (PragmaOptions -> Bool) -> PragmaOptions -> FunctionData
forall b c a. (b -> c) -> (a -> b) -> a -> c
. PragmaOptions -> Bool
optErasure (PragmaOptions -> FunctionData)
-> m PragmaOptions -> m FunctionData
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> m PragmaOptions
forall (m :: * -> *). HasOptions m => m PragmaOptions
pragmaOptions
emptyFunctionData_ :: Bool -> FunctionData
emptyFunctionData_ :: Bool -> FunctionData
emptyFunctionData_ Bool
erasure = FunctionData
{ _funClauses :: [Clause]
_funClauses = []
, _funCompiled :: Maybe CompiledClauses
_funCompiled = Maybe CompiledClauses
forall a. Maybe a
Nothing
, _funSplitTree :: Maybe SplitTree
_funSplitTree = Maybe SplitTree
forall a. Maybe a
Nothing
, _funTreeless :: Maybe Compiled
_funTreeless = Maybe Compiled
forall a. Maybe a
Nothing
, _funInv :: FunctionInverse
_funInv = FunctionInverse
forall c. FunctionInverse' c
NotInjective
, _funMutual :: Maybe [QName]
_funMutual = Maybe [QName]
forall a. Maybe a
Nothing
, _funProjection :: Either ProjectionLikenessMissing Projection
_funProjection = ProjectionLikenessMissing
-> Either ProjectionLikenessMissing Projection
forall a b. a -> Either a b
Left ProjectionLikenessMissing
MaybeProjection
, _funFlags :: SmallSet FunctionFlag
_funFlags = [FunctionFlag] -> SmallSet FunctionFlag
forall a. SmallSetElement a => [a] -> SmallSet a
SmallSet.fromList [ FunctionFlag
FunErasure | Bool
erasure ]
, _funTerminates :: Maybe Bool
_funTerminates = Maybe Bool
forall a. Maybe a
Nothing
, _funExtLam :: Maybe ExtLamInfo
_funExtLam = Maybe ExtLamInfo
forall a. Maybe a
Nothing
, _funWith :: Maybe QName
_funWith = Maybe QName
forall a. Maybe a
Nothing
, _funCovering :: [Clause]
_funCovering = []
, _funIsKanOp :: Maybe QName
_funIsKanOp = Maybe QName
forall a. Maybe a
Nothing
, _funOpaque :: IsOpaque
_funOpaque = IsOpaque
TransparentDef
}
emptyFunction :: HasOptions m => m Defn
emptyFunction :: forall (m :: * -> *). HasOptions m => m Defn
emptyFunction = FunctionData -> Defn
FunctionDefn (FunctionData -> Defn) -> m FunctionData -> m Defn
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> m FunctionData
forall (m :: * -> *). HasOptions m => m FunctionData
emptyFunctionData
emptyFunction_ :: Bool -> Defn
emptyFunction_ :: Bool -> Defn
emptyFunction_ = FunctionData -> Defn
FunctionDefn (FunctionData -> Defn) -> (Bool -> FunctionData) -> Bool -> Defn
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Bool -> FunctionData
emptyFunctionData_
funFlag_ :: FunctionFlag -> Lens' FunctionData Bool
funFlag_ :: FunctionFlag -> Lens' FunctionData Bool
funFlag_ FunctionFlag
flag Bool -> f Bool
f def :: FunctionData
def@FunctionData{ _funFlags :: FunctionData -> SmallSet FunctionFlag
_funFlags = SmallSet FunctionFlag
flags } =
Bool -> f Bool
f (FunctionFlag -> SmallSet FunctionFlag -> Bool
forall a. SmallSetElement a => a -> SmallSet a -> Bool
SmallSet.member FunctionFlag
flag SmallSet FunctionFlag
flags) f Bool -> (Bool -> FunctionData) -> f FunctionData
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&>
\ Bool
b -> FunctionData
def{ _funFlags = (if b then SmallSet.insert else SmallSet.delete) flag flags }
funFlag :: FunctionFlag -> Lens' Defn Bool
funFlag :: FunctionFlag -> Lens' Defn Bool
funFlag FunctionFlag
flag Bool -> f Bool
f = \case
FunctionDefn FunctionData
d -> FunctionData -> Defn
FunctionDefn (FunctionData -> Defn) -> f FunctionData -> f Defn
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> FunctionFlag -> Lens' FunctionData Bool
funFlag_ FunctionFlag
flag Bool -> f Bool
f FunctionData
d
Defn
def -> Bool -> f Bool
f Bool
False f Bool -> Defn -> f Defn
forall (f :: * -> *) a b. Functor f => f a -> b -> f b
$> Defn
def
funStatic, funInline, funMacro :: Lens' Defn Bool
funStatic :: Lens' Defn Bool
funStatic = FunctionFlag -> Lens' Defn Bool
funFlag FunctionFlag
FunStatic
funInline :: Lens' Defn Bool
funInline = FunctionFlag -> Lens' Defn Bool
funFlag FunctionFlag
FunInline
funMacro :: Lens' Defn Bool
funMacro = FunctionFlag -> Lens' Defn Bool
funFlag FunctionFlag
FunMacro
funMacro_ :: Lens' FunctionData Bool
funMacro_ :: Lens' FunctionData Bool
funMacro_ = FunctionFlag -> Lens' FunctionData Bool
funFlag_ FunctionFlag
FunMacro
funFirstOrder :: Lens' Defn Bool
funFirstOrder :: Lens' Defn Bool
funFirstOrder = FunctionFlag -> Lens' Defn Bool
funFlag FunctionFlag
FunFirstOrder
funErasure :: Lens' Defn Bool
funErasure :: Lens' Defn Bool
funErasure = FunctionFlag -> Lens' Defn Bool
funFlag FunctionFlag
FunErasure
funAbstract :: Lens' Defn Bool
funAbstract :: Lens' Defn Bool
funAbstract = FunctionFlag -> Lens' Defn Bool
funFlag FunctionFlag
FunAbstract
funAbstr :: Lens' Defn IsAbstract
funAbstr :: Lens' Defn IsAbstract
funAbstr = (Bool -> f Bool) -> Defn -> f Defn
Lens' Defn Bool
funAbstract ((Bool -> f Bool) -> Defn -> f Defn)
-> ((IsAbstract -> f IsAbstract) -> Bool -> f Bool)
-> (IsAbstract -> f IsAbstract)
-> Defn
-> f Defn
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (Bool -> IsAbstract)
-> (IsAbstract -> Bool) -> Lens' Bool IsAbstract
forall o i. (o -> i) -> (i -> o) -> Lens' o i
iso Bool -> IsAbstract
forall a. Boolean a => Bool -> a
fromBool IsAbstract -> Bool
forall a. IsBool a => a -> Bool
toBool
funAbstract_ :: Lens' FunctionData Bool
funAbstract_ :: Lens' FunctionData Bool
funAbstract_ = FunctionFlag -> Lens' FunctionData Bool
funFlag_ FunctionFlag
FunAbstract
funAbstr_ :: Lens' FunctionData IsAbstract
funAbstr_ :: Lens' FunctionData IsAbstract
funAbstr_ = (Bool -> f Bool) -> FunctionData -> f FunctionData
Lens' FunctionData Bool
funAbstract_ ((Bool -> f Bool) -> FunctionData -> f FunctionData)
-> ((IsAbstract -> f IsAbstract) -> Bool -> f Bool)
-> (IsAbstract -> f IsAbstract)
-> FunctionData
-> f FunctionData
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (Bool -> IsAbstract)
-> (IsAbstract -> Bool) -> Lens' Bool IsAbstract
forall o i. (o -> i) -> (i -> o) -> Lens' o i
iso Bool -> IsAbstract
forall a. Boolean a => Bool -> a
fromBool IsAbstract -> Bool
forall a. IsBool a => a -> Bool
toBool
funProj :: Lens' Defn Bool
funProj :: Lens' Defn Bool
funProj = FunctionFlag -> Lens' Defn Bool
funFlag FunctionFlag
FunProj
funProj_ :: Lens' FunctionData Bool
funProj_ :: Lens' FunctionData Bool
funProj_ = FunctionFlag -> Lens' FunctionData Bool
funFlag_ FunctionFlag
FunProj
isMacro :: Defn -> Bool
isMacro :: Defn -> Bool
isMacro = (Defn -> Lens' Defn Bool -> Bool
forall o i. o -> Lens' o i -> i
^. (Bool -> f Bool) -> Defn -> f Defn
Lens' Defn Bool
funMacro)
isEmptyFunction :: Defn -> Bool
isEmptyFunction :: Defn -> Bool
isEmptyFunction Defn
def =
case Defn
def of
Function { funClauses :: Defn -> [Clause]
funClauses = [] } -> Bool
True
Defn
_ -> Bool
False
isExtendedLambda :: Defn -> Bool
isExtendedLambda :: Defn -> Bool
isExtendedLambda Defn
def =
case Defn
def of
Function { funExtLam :: Defn -> Maybe ExtLamInfo
funExtLam = Just{} } -> Bool
True
Defn
_ -> Bool
False
isWithFunction :: Defn -> Bool
isWithFunction :: Defn -> Bool
isWithFunction Defn
def =
case Defn
def of
Function { funWith :: Defn -> Maybe QName
funWith = Just{} } -> Bool
True
Defn
_ -> Bool
False
isCopatternLHS :: [Clause] -> Bool
isCopatternLHS :: [Clause] -> Bool
isCopatternLHS = (Clause -> Bool) -> [Clause] -> Bool
forall (t :: * -> *) a. Foldable t => (a -> Bool) -> t a -> Bool
List.any ((NamedArg DeBruijnPattern -> Bool)
-> [NamedArg DeBruijnPattern] -> Bool
forall (t :: * -> *) a. Foldable t => (a -> Bool) -> t a -> Bool
List.any (Maybe (ProjOrigin, AmbiguousQName) -> Bool
forall a. Maybe a -> Bool
isJust (Maybe (ProjOrigin, AmbiguousQName) -> Bool)
-> (NamedArg DeBruijnPattern -> Maybe (ProjOrigin, AmbiguousQName))
-> NamedArg DeBruijnPattern
-> Bool
forall b c a. (b -> c) -> (a -> b) -> a -> c
. NamedArg DeBruijnPattern -> Maybe (ProjOrigin, AmbiguousQName)
forall a. IsProjP a => a -> Maybe (ProjOrigin, AmbiguousQName)
A.isProjP) ([NamedArg DeBruijnPattern] -> Bool)
-> (Clause -> [NamedArg DeBruijnPattern]) -> Clause -> Bool
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Clause -> [NamedArg DeBruijnPattern]
namedClausePats)
recCon :: Defn -> QName
recCon :: Defn -> QName
recCon Record{ ConHead
recConHead :: Defn -> ConHead
recConHead :: ConHead
recConHead } = ConHead -> QName
conName ConHead
recConHead
recCon Defn
_ = QName
forall a. HasCallStack => a
__IMPOSSIBLE__
defIsRecord :: Defn -> Bool
defIsRecord :: Defn -> Bool
defIsRecord Record{} = Bool
True
defIsRecord Defn
_ = Bool
False
defIsDataOrRecord :: Defn -> Bool
defIsDataOrRecord :: Defn -> Bool
defIsDataOrRecord Record{} = Bool
True
defIsDataOrRecord Datatype{} = Bool
True
defIsDataOrRecord Defn
_ = Bool
False
defConstructors :: Defn -> [QName]
defConstructors :: Defn -> [QName]
defConstructors Datatype{dataCons :: Defn -> [QName]
dataCons = [QName]
cs} = [QName]
cs
defConstructors Record{recConHead :: Defn -> ConHead
recConHead = ConHead
c} = [ConHead -> QName
conName ConHead
c]
defConstructors Defn
_ = [QName]
forall a. HasCallStack => a
__IMPOSSIBLE__
newtype Fields = Fields [(C.Name, Type)]
deriving Fields
Fields -> Bool
Fields -> (Fields -> Bool) -> Null Fields
forall a. a -> (a -> Bool) -> Null a
$cempty :: Fields
empty :: Fields
$cnull :: Fields -> Bool
null :: Fields -> Bool
Null
data Simplification = YesSimplification | NoSimplification
deriving (Simplification -> Simplification -> Bool
(Simplification -> Simplification -> Bool)
-> (Simplification -> Simplification -> Bool) -> Eq Simplification
forall a. (a -> a -> Bool) -> (a -> a -> Bool) -> Eq a
$c== :: Simplification -> Simplification -> Bool
== :: Simplification -> Simplification -> Bool
$c/= :: Simplification -> Simplification -> Bool
/= :: Simplification -> Simplification -> Bool
Eq, Int -> Simplification -> ShowS
[Simplification] -> ShowS
Simplification -> String
(Int -> Simplification -> ShowS)
-> (Simplification -> String)
-> ([Simplification] -> ShowS)
-> Show Simplification
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> Simplification -> ShowS
showsPrec :: Int -> Simplification -> ShowS
$cshow :: Simplification -> String
show :: Simplification -> String
$cshowList :: [Simplification] -> ShowS
showList :: [Simplification] -> ShowS
Show, (forall x. Simplification -> Rep Simplification x)
-> (forall x. Rep Simplification x -> Simplification)
-> Generic Simplification
forall x. Rep Simplification x -> Simplification
forall x. Simplification -> Rep Simplification x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. Simplification -> Rep Simplification x
from :: forall x. Simplification -> Rep Simplification x
$cto :: forall x. Rep Simplification x -> Simplification
to :: forall x. Rep Simplification x -> Simplification
Generic)
instance Null Simplification where
empty :: Simplification
empty = Simplification
NoSimplification
null :: Simplification -> Bool
null = (Simplification -> Simplification -> Bool
forall a. Eq a => a -> a -> Bool
== Simplification
NoSimplification)
instance Semigroup Simplification where
Simplification
YesSimplification <> :: Simplification -> Simplification -> Simplification
<> Simplification
_ = Simplification
YesSimplification
Simplification
NoSimplification <> Simplification
s = Simplification
s
instance Monoid Simplification where
mempty :: Simplification
mempty = Simplification
NoSimplification
mappend :: Simplification -> Simplification -> Simplification
mappend = Simplification -> Simplification -> Simplification
forall a. Semigroup a => a -> a -> a
(<>)
data Reduced no yes
= NoReduction no
| YesReduction Simplification yes
deriving (forall a b. (a -> b) -> Reduced no a -> Reduced no b)
-> (forall a b. a -> Reduced no b -> Reduced no a)
-> Functor (Reduced no)
forall a b. a -> Reduced no b -> Reduced no a
forall a b. (a -> b) -> Reduced no a -> Reduced no b
forall no a b. a -> Reduced no b -> Reduced no a
forall no a b. (a -> b) -> Reduced no a -> Reduced no b
forall (f :: * -> *).
(forall a b. (a -> b) -> f a -> f b)
-> (forall a b. a -> f b -> f a) -> Functor f
$cfmap :: forall no a b. (a -> b) -> Reduced no a -> Reduced no b
fmap :: forall a b. (a -> b) -> Reduced no a -> Reduced no b
$c<$ :: forall no a b. a -> Reduced no b -> Reduced no a
<$ :: forall a b. a -> Reduced no b -> Reduced no a
Functor
redReturn :: a -> ReduceM (Reduced a' a)
redReturn :: forall a a'. a -> ReduceM (Reduced a' a)
redReturn = Reduced a' a -> ReduceM (Reduced a' a)
forall a. a -> ReduceM a
forall (m :: * -> *) a. Monad m => a -> m a
return (Reduced a' a -> ReduceM (Reduced a' a))
-> (a -> Reduced a' a) -> a -> ReduceM (Reduced a' a)
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Simplification -> a -> Reduced a' a
forall no yes. Simplification -> yes -> Reduced no yes
YesReduction Simplification
YesSimplification
redBind :: ReduceM (Reduced a a') -> (a -> b) ->
(a' -> ReduceM (Reduced b b')) -> ReduceM (Reduced b b')
redBind :: forall a a' b b'.
ReduceM (Reduced a a')
-> (a -> b)
-> (a' -> ReduceM (Reduced b b'))
-> ReduceM (Reduced b b')
redBind ReduceM (Reduced a a')
ma a -> b
f a' -> ReduceM (Reduced b b')
k = do
r <- ReduceM (Reduced a a')
ma
case r of
NoReduction a
x -> Reduced b b' -> ReduceM (Reduced b b')
forall a. a -> ReduceM a
forall (m :: * -> *) a. Monad m => a -> m a
return (Reduced b b' -> ReduceM (Reduced b b'))
-> Reduced b b' -> ReduceM (Reduced b b')
forall a b. (a -> b) -> a -> b
$ b -> Reduced b b'
forall no yes. no -> Reduced no yes
NoReduction (b -> Reduced b b') -> b -> Reduced b b'
forall a b. (a -> b) -> a -> b
$ a -> b
f a
x
YesReduction Simplification
_ a'
y -> a' -> ReduceM (Reduced b b')
k a'
y
data IsReduced
= NotReduced
| Reduced (Blocked ())
data MaybeReduced a = MaybeRed
{ forall a. MaybeReduced a -> IsReduced
isReduced :: IsReduced
, forall a. MaybeReduced a -> a
ignoreReduced :: a
}
deriving ((forall a b. (a -> b) -> MaybeReduced a -> MaybeReduced b)
-> (forall a b. a -> MaybeReduced b -> MaybeReduced a)
-> Functor MaybeReduced
forall a b. a -> MaybeReduced b -> MaybeReduced a
forall a b. (a -> b) -> MaybeReduced a -> MaybeReduced b
forall (f :: * -> *).
(forall a b. (a -> b) -> f a -> f b)
-> (forall a b. a -> f b -> f a) -> Functor f
$cfmap :: forall a b. (a -> b) -> MaybeReduced a -> MaybeReduced b
fmap :: forall a b. (a -> b) -> MaybeReduced a -> MaybeReduced b
$c<$ :: forall a b. a -> MaybeReduced b -> MaybeReduced a
<$ :: forall a b. a -> MaybeReduced b -> MaybeReduced a
Functor)
instance IsProjElim e => IsProjElim (MaybeReduced e) where
isProjElim :: MaybeReduced e -> Maybe (ProjOrigin, QName)
isProjElim = e -> Maybe (ProjOrigin, QName)
forall e. IsProjElim e => e -> Maybe (ProjOrigin, QName)
isProjElim (e -> Maybe (ProjOrigin, QName))
-> (MaybeReduced e -> e)
-> MaybeReduced e
-> Maybe (ProjOrigin, QName)
forall b c a. (b -> c) -> (a -> b) -> a -> c
. MaybeReduced e -> e
forall a. MaybeReduced a -> a
ignoreReduced
type MaybeReducedArgs = [MaybeReduced (Arg Term)]
type MaybeReducedElims = [MaybeReduced Elim]
notReduced :: a -> MaybeReduced a
notReduced :: forall a. a -> MaybeReduced a
notReduced a
x = IsReduced -> a -> MaybeReduced a
forall a. IsReduced -> a -> MaybeReduced a
MaybeRed IsReduced
NotReduced a
x
reduced :: Blocked (Arg Term) -> MaybeReduced (Arg Term)
reduced :: Blocked (Arg Term) -> MaybeReduced (Arg Term)
reduced Blocked (Arg Term)
b = IsReduced -> Arg Term -> MaybeReduced (Arg Term)
forall a. IsReduced -> a -> MaybeReduced a
MaybeRed (Blocked_ -> IsReduced
Reduced (Blocked_ -> IsReduced) -> Blocked_ -> IsReduced
forall a b. (a -> b) -> a -> b
$ () () -> Blocked (Arg Term) -> Blocked_
forall a b. a -> Blocked' Term b -> Blocked' Term a
forall (f :: * -> *) a b. Functor f => a -> f b -> f a
<$ Blocked (Arg Term)
b) (Arg Term -> MaybeReduced (Arg Term))
-> Arg Term -> MaybeReduced (Arg Term)
forall a b. (a -> b) -> a -> b
$ Blocked (Arg Term) -> Arg Term
forall t a. Blocked' t a -> a
ignoreBlocking Blocked (Arg Term)
b
data AllowedReduction
= ProjectionReductions
| InlineReductions
| CopatternReductions
| FunctionReductions
| RecursiveReductions
| LevelReductions
| TypeLevelReductions
| UnconfirmedReductions
| NonTerminatingReductions
deriving (Int -> AllowedReduction -> ShowS
[AllowedReduction] -> ShowS
AllowedReduction -> String
(Int -> AllowedReduction -> ShowS)
-> (AllowedReduction -> String)
-> ([AllowedReduction] -> ShowS)
-> Show AllowedReduction
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> AllowedReduction -> ShowS
showsPrec :: Int -> AllowedReduction -> ShowS
$cshow :: AllowedReduction -> String
show :: AllowedReduction -> String
$cshowList :: [AllowedReduction] -> ShowS
showList :: [AllowedReduction] -> ShowS
Show, AllowedReduction -> AllowedReduction -> Bool
(AllowedReduction -> AllowedReduction -> Bool)
-> (AllowedReduction -> AllowedReduction -> Bool)
-> Eq AllowedReduction
forall a. (a -> a -> Bool) -> (a -> a -> Bool) -> Eq a
$c== :: AllowedReduction -> AllowedReduction -> Bool
== :: AllowedReduction -> AllowedReduction -> Bool
$c/= :: AllowedReduction -> AllowedReduction -> Bool
/= :: AllowedReduction -> AllowedReduction -> Bool
Eq, Eq AllowedReduction
Eq AllowedReduction =>
(AllowedReduction -> AllowedReduction -> Ordering)
-> (AllowedReduction -> AllowedReduction -> Bool)
-> (AllowedReduction -> AllowedReduction -> Bool)
-> (AllowedReduction -> AllowedReduction -> Bool)
-> (AllowedReduction -> AllowedReduction -> Bool)
-> (AllowedReduction -> AllowedReduction -> AllowedReduction)
-> (AllowedReduction -> AllowedReduction -> AllowedReduction)
-> Ord AllowedReduction
AllowedReduction -> AllowedReduction -> Bool
AllowedReduction -> AllowedReduction -> Ordering
AllowedReduction -> AllowedReduction -> AllowedReduction
forall a.
Eq a =>
(a -> a -> Ordering)
-> (a -> a -> Bool)
-> (a -> a -> Bool)
-> (a -> a -> Bool)
-> (a -> a -> Bool)
-> (a -> a -> a)
-> (a -> a -> a)
-> Ord a
$ccompare :: AllowedReduction -> AllowedReduction -> Ordering
compare :: AllowedReduction -> AllowedReduction -> Ordering
$c< :: AllowedReduction -> AllowedReduction -> Bool
< :: AllowedReduction -> AllowedReduction -> Bool
$c<= :: AllowedReduction -> AllowedReduction -> Bool
<= :: AllowedReduction -> AllowedReduction -> Bool
$c> :: AllowedReduction -> AllowedReduction -> Bool
> :: AllowedReduction -> AllowedReduction -> Bool
$c>= :: AllowedReduction -> AllowedReduction -> Bool
>= :: AllowedReduction -> AllowedReduction -> Bool
$cmax :: AllowedReduction -> AllowedReduction -> AllowedReduction
max :: AllowedReduction -> AllowedReduction -> AllowedReduction
$cmin :: AllowedReduction -> AllowedReduction -> AllowedReduction
min :: AllowedReduction -> AllowedReduction -> AllowedReduction
Ord, Int -> AllowedReduction
AllowedReduction -> Int
AllowedReduction -> [AllowedReduction]
AllowedReduction -> AllowedReduction
AllowedReduction -> AllowedReduction -> [AllowedReduction]
AllowedReduction
-> AllowedReduction -> AllowedReduction -> [AllowedReduction]
(AllowedReduction -> AllowedReduction)
-> (AllowedReduction -> AllowedReduction)
-> (Int -> AllowedReduction)
-> (AllowedReduction -> Int)
-> (AllowedReduction -> [AllowedReduction])
-> (AllowedReduction -> AllowedReduction -> [AllowedReduction])
-> (AllowedReduction -> AllowedReduction -> [AllowedReduction])
-> (AllowedReduction
-> AllowedReduction -> AllowedReduction -> [AllowedReduction])
-> Enum AllowedReduction
forall a.
(a -> a)
-> (a -> a)
-> (Int -> a)
-> (a -> Int)
-> (a -> [a])
-> (a -> a -> [a])
-> (a -> a -> [a])
-> (a -> a -> a -> [a])
-> Enum a
$csucc :: AllowedReduction -> AllowedReduction
succ :: AllowedReduction -> AllowedReduction
$cpred :: AllowedReduction -> AllowedReduction
pred :: AllowedReduction -> AllowedReduction
$ctoEnum :: Int -> AllowedReduction
toEnum :: Int -> AllowedReduction
$cfromEnum :: AllowedReduction -> Int
fromEnum :: AllowedReduction -> Int
$cenumFrom :: AllowedReduction -> [AllowedReduction]
enumFrom :: AllowedReduction -> [AllowedReduction]
$cenumFromThen :: AllowedReduction -> AllowedReduction -> [AllowedReduction]
enumFromThen :: AllowedReduction -> AllowedReduction -> [AllowedReduction]
$cenumFromTo :: AllowedReduction -> AllowedReduction -> [AllowedReduction]
enumFromTo :: AllowedReduction -> AllowedReduction -> [AllowedReduction]
$cenumFromThenTo :: AllowedReduction
-> AllowedReduction -> AllowedReduction -> [AllowedReduction]
enumFromThenTo :: AllowedReduction
-> AllowedReduction -> AllowedReduction -> [AllowedReduction]
Enum, AllowedReduction
AllowedReduction -> AllowedReduction -> Bounded AllowedReduction
forall a. a -> a -> Bounded a
$cminBound :: AllowedReduction
minBound :: AllowedReduction
$cmaxBound :: AllowedReduction
maxBound :: AllowedReduction
Bounded, Ord AllowedReduction
Ord AllowedReduction =>
((AllowedReduction, AllowedReduction) -> [AllowedReduction])
-> ((AllowedReduction, AllowedReduction)
-> AllowedReduction -> Int)
-> ((AllowedReduction, AllowedReduction)
-> AllowedReduction -> Int)
-> ((AllowedReduction, AllowedReduction)
-> AllowedReduction -> Bool)
-> ((AllowedReduction, AllowedReduction) -> Int)
-> ((AllowedReduction, AllowedReduction) -> Int)
-> Ix AllowedReduction
(AllowedReduction, AllowedReduction) -> Int
(AllowedReduction, AllowedReduction) -> [AllowedReduction]
(AllowedReduction, AllowedReduction) -> AllowedReduction -> Bool
(AllowedReduction, AllowedReduction) -> AllowedReduction -> Int
forall a.
Ord a =>
((a, a) -> [a])
-> ((a, a) -> a -> Int)
-> ((a, a) -> a -> Int)
-> ((a, a) -> a -> Bool)
-> ((a, a) -> Int)
-> ((a, a) -> Int)
-> Ix a
$crange :: (AllowedReduction, AllowedReduction) -> [AllowedReduction]
range :: (AllowedReduction, AllowedReduction) -> [AllowedReduction]
$cindex :: (AllowedReduction, AllowedReduction) -> AllowedReduction -> Int
index :: (AllowedReduction, AllowedReduction) -> AllowedReduction -> Int
$cunsafeIndex :: (AllowedReduction, AllowedReduction) -> AllowedReduction -> Int
unsafeIndex :: (AllowedReduction, AllowedReduction) -> AllowedReduction -> Int
$cinRange :: (AllowedReduction, AllowedReduction) -> AllowedReduction -> Bool
inRange :: (AllowedReduction, AllowedReduction) -> AllowedReduction -> Bool
$crangeSize :: (AllowedReduction, AllowedReduction) -> Int
rangeSize :: (AllowedReduction, AllowedReduction) -> Int
$cunsafeRangeSize :: (AllowedReduction, AllowedReduction) -> Int
unsafeRangeSize :: (AllowedReduction, AllowedReduction) -> Int
Ix, (forall x. AllowedReduction -> Rep AllowedReduction x)
-> (forall x. Rep AllowedReduction x -> AllowedReduction)
-> Generic AllowedReduction
forall x. Rep AllowedReduction x -> AllowedReduction
forall x. AllowedReduction -> Rep AllowedReduction x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. AllowedReduction -> Rep AllowedReduction x
from :: forall x. AllowedReduction -> Rep AllowedReduction x
$cto :: forall x. Rep AllowedReduction x -> AllowedReduction
to :: forall x. Rep AllowedReduction x -> AllowedReduction
Generic)
instance SmallSet.SmallSetElement AllowedReduction
type AllowedReductions = SmallSet AllowedReduction
allReductions :: AllowedReductions
allReductions :: AllowedReductions
allReductions = AllowedReduction -> AllowedReductions -> AllowedReductions
forall a. SmallSetElement a => a -> SmallSet a -> SmallSet a
SmallSet.delete AllowedReduction
NonTerminatingReductions AllowedReductions
reallyAllReductions
reallyAllReductions :: AllowedReductions
reallyAllReductions :: AllowedReductions
reallyAllReductions = AllowedReductions
forall a. SmallSetElement a => SmallSet a
SmallSet.total
data ReduceDefs
= OnlyReduceDefs (Set QName)
| DontReduceDefs (Set QName)
deriving (forall x. ReduceDefs -> Rep ReduceDefs x)
-> (forall x. Rep ReduceDefs x -> ReduceDefs) -> Generic ReduceDefs
forall x. Rep ReduceDefs x -> ReduceDefs
forall x. ReduceDefs -> Rep ReduceDefs x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. ReduceDefs -> Rep ReduceDefs x
from :: forall x. ReduceDefs -> Rep ReduceDefs x
$cto :: forall x. Rep ReduceDefs x -> ReduceDefs
to :: forall x. Rep ReduceDefs x -> ReduceDefs
Generic
reduceAllDefs :: ReduceDefs
reduceAllDefs :: ReduceDefs
reduceAllDefs = Set QName -> ReduceDefs
DontReduceDefs Set QName
forall a. Null a => a
empty
locallyReduceDefs :: MonadTCEnv m => ReduceDefs -> m a -> m a
locallyReduceDefs :: forall (m :: * -> *) a. MonadTCEnv m => ReduceDefs -> m a -> m a
locallyReduceDefs = Lens' TCEnv ReduceDefs -> (ReduceDefs -> ReduceDefs) -> m a -> m a
forall (m :: * -> *) a b.
MonadTCEnv m =>
Lens' TCEnv a -> (a -> a) -> m b -> m b
locallyTC (ReduceDefs -> f ReduceDefs) -> TCEnv -> f TCEnv
Lens' TCEnv ReduceDefs
eReduceDefs ((ReduceDefs -> ReduceDefs) -> m a -> m a)
-> (ReduceDefs -> ReduceDefs -> ReduceDefs)
-> ReduceDefs
-> m a
-> m a
forall b c a. (b -> c) -> (a -> b) -> a -> c
. ReduceDefs -> ReduceDefs -> ReduceDefs
forall a b. a -> b -> a
const
locallyReduceAllDefs :: MonadTCEnv m => m a -> m a
locallyReduceAllDefs :: forall (m :: * -> *) a. MonadTCEnv m => m a -> m a
locallyReduceAllDefs = ReduceDefs -> m a -> m a
forall (m :: * -> *) a. MonadTCEnv m => ReduceDefs -> m a -> m a
locallyReduceDefs ReduceDefs
reduceAllDefs
shouldReduceDef :: (MonadTCEnv m) => QName -> m Bool
shouldReduceDef :: forall (m :: * -> *). MonadTCEnv m => QName -> m Bool
shouldReduceDef QName
f = (TCEnv -> ReduceDefs) -> m ReduceDefs
forall (m :: * -> *) a. MonadTCEnv m => (TCEnv -> a) -> m a
asksTC TCEnv -> ReduceDefs
envReduceDefs m ReduceDefs -> (ReduceDefs -> Bool) -> m Bool
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \case
OnlyReduceDefs Set QName
defs -> QName
f QName -> Set QName -> Bool
forall a. Ord a => a -> Set a -> Bool
`Set.member` Set QName
defs
DontReduceDefs Set QName
defs -> Bool -> Bool
not (Bool -> Bool) -> Bool -> Bool
forall a b. (a -> b) -> a -> b
$ QName
f QName -> Set QName -> Bool
forall a. Ord a => a -> Set a -> Bool
`Set.member` Set QName
defs
toReduceDefs :: (Bool, [QName]) -> ReduceDefs
toReduceDefs :: (Bool, [QName]) -> ReduceDefs
toReduceDefs (Bool
True, [QName]
ns) = Set QName -> ReduceDefs
OnlyReduceDefs ([QName] -> Set QName
forall a. Ord a => [a] -> Set a
Data.Set.fromList [QName]
ns)
toReduceDefs (Bool
False, [QName]
ns) = Set QName -> ReduceDefs
DontReduceDefs ([QName] -> Set QName
forall a. Ord a => [a] -> Set a
Data.Set.fromList [QName]
ns)
fromReduceDefs :: ReduceDefs -> (Bool, [QName])
fromReduceDefs :: ReduceDefs -> (Bool, [QName])
fromReduceDefs (OnlyReduceDefs Set QName
ns) = (Bool
True, Set QName -> [QName]
forall a. Set a -> [a]
toList Set QName
ns)
fromReduceDefs (DontReduceDefs Set QName
ns) = (Bool
False, Set QName -> [QName]
forall a. Set a -> [a]
toList Set QName
ns)
locallyReconstructed :: MonadTCEnv m => m a -> m a
locallyReconstructed :: forall (m :: * -> *) a. MonadTCEnv m => m a -> m a
locallyReconstructed = Lens' TCEnv Bool -> (Bool -> Bool) -> m a -> m a
forall (m :: * -> *) a b.
MonadTCEnv m =>
Lens' TCEnv a -> (a -> a) -> m b -> m b
locallyTC (Bool -> f Bool) -> TCEnv -> f TCEnv
Lens' TCEnv Bool
eReconstructed ((Bool -> Bool) -> m a -> m a)
-> (Bool -> Bool -> Bool) -> Bool -> m a -> m a
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Bool -> Bool -> Bool
forall a b. a -> b -> a
const (Bool -> m a -> m a) -> Bool -> m a -> m a
forall a b. (a -> b) -> a -> b
$ Bool
True
isReconstructed :: (MonadTCEnv m) => m Bool
isReconstructed :: forall (m :: * -> *). MonadTCEnv m => m Bool
isReconstructed = (TCEnv -> Bool) -> m Bool
forall (m :: * -> *) a. MonadTCEnv m => (TCEnv -> a) -> m a
asksTC TCEnv -> Bool
envReconstructed
data PrimitiveImpl = PrimImpl Type PrimFun
data PrimFun = PrimFun
{ PrimFun -> QName
primFunName :: QName
, PrimFun -> Int
primFunArity :: Arity
, PrimFun -> [Occurrence]
primFunArgOccurrences :: [Occurrence]
, PrimFun -> Args -> Int -> ReduceM (Reduced MaybeReducedArgs Term)
primFunImplementation :: [Arg Term] -> Int -> ReduceM (Reduced MaybeReducedArgs Term)
}
deriving (forall x. PrimFun -> Rep PrimFun x)
-> (forall x. Rep PrimFun x -> PrimFun) -> Generic PrimFun
forall x. Rep PrimFun x -> PrimFun
forall x. PrimFun -> Rep PrimFun x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. PrimFun -> Rep PrimFun x
from :: forall x. PrimFun -> Rep PrimFun x
$cto :: forall x. Rep PrimFun x -> PrimFun
to :: forall x. Rep PrimFun x -> PrimFun
Generic
primFun :: QName -> Arity -> ([Arg Term] -> ReduceM (Reduced MaybeReducedArgs Term)) -> PrimFun
primFun :: QName
-> Int
-> (Args -> ReduceM (Reduced MaybeReducedArgs Term))
-> PrimFun
primFun QName
q Int
ar Args -> ReduceM (Reduced MaybeReducedArgs Term)
imp = QName
-> Int
-> [Occurrence]
-> (Args -> Int -> ReduceM (Reduced MaybeReducedArgs Term))
-> PrimFun
PrimFun QName
q Int
ar [] (\Args
args Int
_ -> Args -> ReduceM (Reduced MaybeReducedArgs Term)
imp Args
args)
defClauses :: Definition -> [Clause]
defClauses :: Definition -> [Clause]
defClauses Defn{theDef :: Definition -> Defn
theDef = Function{funClauses :: Defn -> [Clause]
funClauses = [Clause]
cs}} = [Clause]
cs
defClauses Defn{theDef :: Definition -> Defn
theDef = Primitive{primClauses :: Defn -> [Clause]
primClauses = [Clause]
cs}} = [Clause]
cs
defClauses Defn{theDef :: Definition -> Defn
theDef = Datatype{dataClause :: Defn -> Maybe Clause
dataClause = Just Clause
c}} = [Clause
c]
defClauses Defn{theDef :: Definition -> Defn
theDef = Record{recClause :: Defn -> Maybe Clause
recClause = Just Clause
c}} = [Clause
c]
defClauses Definition
_ = []
defCompiled :: Definition -> Maybe CompiledClauses
defCompiled :: Definition -> Maybe CompiledClauses
defCompiled Defn{theDef :: Definition -> Defn
theDef = Function {funCompiled :: Defn -> Maybe CompiledClauses
funCompiled = Maybe CompiledClauses
mcc}} = Maybe CompiledClauses
mcc
defCompiled Defn{theDef :: Definition -> Defn
theDef = Primitive{primCompiled :: Defn -> Maybe CompiledClauses
primCompiled = Maybe CompiledClauses
mcc}} = Maybe CompiledClauses
mcc
defCompiled Definition
_ = Maybe CompiledClauses
forall a. Maybe a
Nothing
defParameters :: Definition -> Maybe Nat
defParameters :: Definition -> Maybe Int
defParameters Defn{theDef :: Definition -> Defn
theDef = Datatype{dataPars :: Defn -> Int
dataPars = Int
n}} = Int -> Maybe Int
forall a. a -> Maybe a
Just Int
n
defParameters Defn{theDef :: Definition -> Defn
theDef = Record {recPars :: Defn -> Int
recPars = Int
n}} = Int -> Maybe Int
forall a. a -> Maybe a
Just Int
n
defParameters Definition
_ = Maybe Int
forall a. Maybe a
Nothing
defInverse :: Definition -> FunctionInverse
defInverse :: Definition -> FunctionInverse
defInverse Defn{theDef :: Definition -> Defn
theDef = Function { funInv :: Defn -> FunctionInverse
funInv = FunctionInverse
inv }} = FunctionInverse
inv
defInverse Defn{theDef :: Definition -> Defn
theDef = Primitive{ primInv :: Defn -> FunctionInverse
primInv = FunctionInverse
inv }} = FunctionInverse
inv
defInverse Definition
_ = FunctionInverse
forall c. FunctionInverse' c
NotInjective
defCompilerPragmas :: BackendName -> Definition -> [CompilerPragma]
defCompilerPragmas :: BackendName -> Definition -> [CompilerPragma]
defCompilerPragmas BackendName
b = [CompilerPragma] -> [CompilerPragma]
forall a. [a] -> [a]
reverse ([CompilerPragma] -> [CompilerPragma])
-> (Definition -> [CompilerPragma])
-> Definition
-> [CompilerPragma]
forall b c a. (b -> c) -> (a -> b) -> a -> c
. [CompilerPragma] -> Maybe [CompilerPragma] -> [CompilerPragma]
forall a. a -> Maybe a -> a
fromMaybe [] (Maybe [CompilerPragma] -> [CompilerPragma])
-> (Definition -> Maybe [CompilerPragma])
-> Definition
-> [CompilerPragma]
forall b c a. (b -> c) -> (a -> b) -> a -> c
. BackendName -> CompiledRepresentation -> Maybe [CompilerPragma]
forall k a. Ord k => k -> Map k a -> Maybe a
Map.lookup BackendName
b (CompiledRepresentation -> Maybe [CompilerPragma])
-> (Definition -> CompiledRepresentation)
-> Definition
-> Maybe [CompilerPragma]
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Definition -> CompiledRepresentation
defCompiledRep
defNonterminating :: Definition -> Bool
defNonterminating :: Definition -> Bool
defNonterminating Defn{theDef :: Definition -> Defn
theDef = Function{funTerminates :: Defn -> Maybe Bool
funTerminates = Just Bool
False}} = Bool
True
defNonterminating Definition
_ = Bool
False
defTerminationUnconfirmed :: Definition -> Bool
defTerminationUnconfirmed :: Definition -> Bool
defTerminationUnconfirmed Defn{theDef :: Definition -> Defn
theDef = Function{funTerminates :: Defn -> Maybe Bool
funTerminates = Just Bool
True}} = Bool
False
defTerminationUnconfirmed Defn{theDef :: Definition -> Defn
theDef = Function{funTerminates :: Defn -> Maybe Bool
funTerminates = Maybe Bool
_ }} = Bool
True
defTerminationUnconfirmed Definition
_ = Bool
False
defAbstract :: Definition -> IsAbstract
defAbstract :: Definition -> IsAbstract
defAbstract Definition
def = case Definition -> Defn
theDef Definition
def of
AxiomDefn AxiomData
_ -> IsAbstract
ConcreteDef
DataOrRecSigDefn DataOrRecSigData
_ -> IsAbstract
ConcreteDef
GeneralizableVar NumGeneralizableArgs
_ -> IsAbstract
ConcreteDef
AbstractDefn Defn
_ -> IsAbstract
AbstractDef
FunctionDefn FunctionData
d -> FunctionData
d FunctionData -> Lens' FunctionData IsAbstract -> IsAbstract
forall o i. o -> Lens' o i -> i
^. (IsAbstract -> f IsAbstract) -> FunctionData -> f FunctionData
Lens' FunctionData IsAbstract
funAbstr_
DatatypeDefn DatatypeData
d -> DatatypeData -> IsAbstract
_dataAbstr DatatypeData
d
RecordDefn RecordData
d -> RecordData -> IsAbstract
_recAbstr RecordData
d
ConstructorDefn ConstructorData
d -> ConstructorData -> IsAbstract
_conAbstr ConstructorData
d
PrimitiveDefn PrimitiveData
d -> PrimitiveData -> IsAbstract
_primAbstr PrimitiveData
d
PrimitiveSortDefn PrimitiveSortData
_ -> IsAbstract
ConcreteDef
defOpaque :: Definition -> IsOpaque
defOpaque :: Definition -> IsOpaque
defOpaque Definition
d = case Definition -> Defn
theDef Definition
d of
Function{funOpaque :: Defn -> IsOpaque
funOpaque=IsOpaque
o} -> IsOpaque
o
Primitive{primOpaque :: Defn -> IsOpaque
primOpaque=IsOpaque
o} -> IsOpaque
o
Axiom{} -> IsOpaque
TransparentDef
AbstractDefn{} -> IsOpaque
TransparentDef
DataOrRecSig{} -> IsOpaque
TransparentDef
GeneralizableVar{} -> IsOpaque
TransparentDef
Datatype{} -> IsOpaque
TransparentDef
Record{} -> IsOpaque
TransparentDef
Constructor{} -> IsOpaque
TransparentDef
PrimitiveSort{} -> IsOpaque
TransparentDef
defForced :: Definition -> [IsForced]
defForced :: Definition -> [IsForced]
defForced Definition
d = case Definition -> Defn
theDef Definition
d of
Constructor{conForced :: Defn -> [IsForced]
conForced = [IsForced]
fs} -> [IsForced]
fs
Axiom{} -> []
DataOrRecSig{} -> []
GeneralizableVar{} -> []
AbstractDefn{} -> []
Function{} -> []
Datatype{} -> []
Record{} -> []
Primitive{} -> []
PrimitiveSort{} -> []
type FunctionInverse = FunctionInverse' Clause
type InversionMap c = Map TermHead [c]
data FunctionInverse' c
= NotInjective
| Inverse (InversionMap c)
deriving (Int -> FunctionInverse' c -> ShowS
[FunctionInverse' c] -> ShowS
FunctionInverse' c -> String
(Int -> FunctionInverse' c -> ShowS)
-> (FunctionInverse' c -> String)
-> ([FunctionInverse' c] -> ShowS)
-> Show (FunctionInverse' c)
forall c. Show c => Int -> FunctionInverse' c -> ShowS
forall c. Show c => [FunctionInverse' c] -> ShowS
forall c. Show c => FunctionInverse' c -> String
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: forall c. Show c => Int -> FunctionInverse' c -> ShowS
showsPrec :: Int -> FunctionInverse' c -> ShowS
$cshow :: forall c. Show c => FunctionInverse' c -> String
show :: FunctionInverse' c -> String
$cshowList :: forall c. Show c => [FunctionInverse' c] -> ShowS
showList :: [FunctionInverse' c] -> ShowS
Show, (forall a b. (a -> b) -> FunctionInverse' a -> FunctionInverse' b)
-> (forall a b. a -> FunctionInverse' b -> FunctionInverse' a)
-> Functor FunctionInverse'
forall a b. a -> FunctionInverse' b -> FunctionInverse' a
forall a b. (a -> b) -> FunctionInverse' a -> FunctionInverse' b
forall (f :: * -> *).
(forall a b. (a -> b) -> f a -> f b)
-> (forall a b. a -> f b -> f a) -> Functor f
$cfmap :: forall a b. (a -> b) -> FunctionInverse' a -> FunctionInverse' b
fmap :: forall a b. (a -> b) -> FunctionInverse' a -> FunctionInverse' b
$c<$ :: forall a b. a -> FunctionInverse' b -> FunctionInverse' a
<$ :: forall a b. a -> FunctionInverse' b -> FunctionInverse' a
Functor, (forall x. FunctionInverse' c -> Rep (FunctionInverse' c) x)
-> (forall x. Rep (FunctionInverse' c) x -> FunctionInverse' c)
-> Generic (FunctionInverse' c)
forall x. Rep (FunctionInverse' c) x -> FunctionInverse' c
forall x. FunctionInverse' c -> Rep (FunctionInverse' c) x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
forall c x. Rep (FunctionInverse' c) x -> FunctionInverse' c
forall c x. FunctionInverse' c -> Rep (FunctionInverse' c) x
$cfrom :: forall c x. FunctionInverse' c -> Rep (FunctionInverse' c) x
from :: forall x. FunctionInverse' c -> Rep (FunctionInverse' c) x
$cto :: forall c x. Rep (FunctionInverse' c) x -> FunctionInverse' c
to :: forall x. Rep (FunctionInverse' c) x -> FunctionInverse' c
Generic)
data TermHead = SortHead
| PiHead
| ConsHead QName
| VarHead Nat
| UnknownHead
deriving (TermHead -> TermHead -> Bool
(TermHead -> TermHead -> Bool)
-> (TermHead -> TermHead -> Bool) -> Eq TermHead
forall a. (a -> a -> Bool) -> (a -> a -> Bool) -> Eq a
$c== :: TermHead -> TermHead -> Bool
== :: TermHead -> TermHead -> Bool
$c/= :: TermHead -> TermHead -> Bool
/= :: TermHead -> TermHead -> Bool
Eq, Eq TermHead
Eq TermHead =>
(TermHead -> TermHead -> Ordering)
-> (TermHead -> TermHead -> Bool)
-> (TermHead -> TermHead -> Bool)
-> (TermHead -> TermHead -> Bool)
-> (TermHead -> TermHead -> Bool)
-> (TermHead -> TermHead -> TermHead)
-> (TermHead -> TermHead -> TermHead)
-> Ord TermHead
TermHead -> TermHead -> Bool
TermHead -> TermHead -> Ordering
TermHead -> TermHead -> TermHead
forall a.
Eq a =>
(a -> a -> Ordering)
-> (a -> a -> Bool)
-> (a -> a -> Bool)
-> (a -> a -> Bool)
-> (a -> a -> Bool)
-> (a -> a -> a)
-> (a -> a -> a)
-> Ord a
$ccompare :: TermHead -> TermHead -> Ordering
compare :: TermHead -> TermHead -> Ordering
$c< :: TermHead -> TermHead -> Bool
< :: TermHead -> TermHead -> Bool
$c<= :: TermHead -> TermHead -> Bool
<= :: TermHead -> TermHead -> Bool
$c> :: TermHead -> TermHead -> Bool
> :: TermHead -> TermHead -> Bool
$c>= :: TermHead -> TermHead -> Bool
>= :: TermHead -> TermHead -> Bool
$cmax :: TermHead -> TermHead -> TermHead
max :: TermHead -> TermHead -> TermHead
$cmin :: TermHead -> TermHead -> TermHead
min :: TermHead -> TermHead -> TermHead
Ord, Int -> TermHead -> ShowS
[TermHead] -> ShowS
TermHead -> String
(Int -> TermHead -> ShowS)
-> (TermHead -> String) -> ([TermHead] -> ShowS) -> Show TermHead
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> TermHead -> ShowS
showsPrec :: Int -> TermHead -> ShowS
$cshow :: TermHead -> String
show :: TermHead -> String
$cshowList :: [TermHead] -> ShowS
showList :: [TermHead] -> ShowS
Show, (forall x. TermHead -> Rep TermHead x)
-> (forall x. Rep TermHead x -> TermHead) -> Generic TermHead
forall x. Rep TermHead x -> TermHead
forall x. TermHead -> Rep TermHead x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. TermHead -> Rep TermHead x
from :: forall x. TermHead -> Rep TermHead x
$cto :: forall x. Rep TermHead x -> TermHead
to :: forall x. Rep TermHead x -> TermHead
Generic)
instance Pretty TermHead where
pretty :: TermHead -> Doc
pretty = \ case
TermHead
SortHead -> Doc
"SortHead"
TermHead
PiHead -> Doc
"PiHead"
ConsHead QName
q -> Doc
"ConsHead" Doc -> Doc -> Doc
forall a. Doc a -> Doc a -> Doc a
<+> QName -> Doc
forall a. Pretty a => a -> Doc
pretty QName
q
VarHead Int
i -> String -> Doc
forall a. String -> Doc a
text (String
"VarHead " String -> ShowS
forall a. [a] -> [a] -> [a]
++ Int -> String
forall a. Show a => a -> String
show Int
i)
TermHead
UnknownHead -> Doc
"UnknownHead"
newtype MutualId = MutualId Word32
deriving (MutualId -> MutualId -> Bool
(MutualId -> MutualId -> Bool)
-> (MutualId -> MutualId -> Bool) -> Eq MutualId
forall a. (a -> a -> Bool) -> (a -> a -> Bool) -> Eq a
$c== :: MutualId -> MutualId -> Bool
== :: MutualId -> MutualId -> Bool
$c/= :: MutualId -> MutualId -> Bool
/= :: MutualId -> MutualId -> Bool
Eq, Eq MutualId
Eq MutualId =>
(MutualId -> MutualId -> Ordering)
-> (MutualId -> MutualId -> Bool)
-> (MutualId -> MutualId -> Bool)
-> (MutualId -> MutualId -> Bool)
-> (MutualId -> MutualId -> Bool)
-> (MutualId -> MutualId -> MutualId)
-> (MutualId -> MutualId -> MutualId)
-> Ord MutualId
MutualId -> MutualId -> Bool
MutualId -> MutualId -> Ordering
MutualId -> MutualId -> MutualId
forall a.
Eq a =>
(a -> a -> Ordering)
-> (a -> a -> Bool)
-> (a -> a -> Bool)
-> (a -> a -> Bool)
-> (a -> a -> Bool)
-> (a -> a -> a)
-> (a -> a -> a)
-> Ord a
$ccompare :: MutualId -> MutualId -> Ordering
compare :: MutualId -> MutualId -> Ordering
$c< :: MutualId -> MutualId -> Bool
< :: MutualId -> MutualId -> Bool
$c<= :: MutualId -> MutualId -> Bool
<= :: MutualId -> MutualId -> Bool
$c> :: MutualId -> MutualId -> Bool
> :: MutualId -> MutualId -> Bool
$c>= :: MutualId -> MutualId -> Bool
>= :: MutualId -> MutualId -> Bool
$cmax :: MutualId -> MutualId -> MutualId
max :: MutualId -> MutualId -> MutualId
$cmin :: MutualId -> MutualId -> MutualId
min :: MutualId -> MutualId -> MutualId
Ord, Int -> MutualId -> ShowS
[MutualId] -> ShowS
MutualId -> String
(Int -> MutualId -> ShowS)
-> (MutualId -> String) -> ([MutualId] -> ShowS) -> Show MutualId
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> MutualId -> ShowS
showsPrec :: Int -> MutualId -> ShowS
$cshow :: MutualId -> String
show :: MutualId -> String
$cshowList :: [MutualId] -> ShowS
showList :: [MutualId] -> ShowS
Show, Integer -> MutualId
MutualId -> MutualId
MutualId -> MutualId -> MutualId
(MutualId -> MutualId -> MutualId)
-> (MutualId -> MutualId -> MutualId)
-> (MutualId -> MutualId -> MutualId)
-> (MutualId -> MutualId)
-> (MutualId -> MutualId)
-> (MutualId -> MutualId)
-> (Integer -> MutualId)
-> Num MutualId
forall a.
(a -> a -> a)
-> (a -> a -> a)
-> (a -> a -> a)
-> (a -> a)
-> (a -> a)
-> (a -> a)
-> (Integer -> a)
-> Num a
$c+ :: MutualId -> MutualId -> MutualId
+ :: MutualId -> MutualId -> MutualId
$c- :: MutualId -> MutualId -> MutualId
- :: MutualId -> MutualId -> MutualId
$c* :: MutualId -> MutualId -> MutualId
* :: MutualId -> MutualId -> MutualId
$cnegate :: MutualId -> MutualId
negate :: MutualId -> MutualId
$cabs :: MutualId -> MutualId
abs :: MutualId -> MutualId
$csignum :: MutualId -> MutualId
signum :: MutualId -> MutualId
$cfromInteger :: Integer -> MutualId
fromInteger :: Integer -> MutualId
Num, Int -> MutualId
MutualId -> Int
MutualId -> [MutualId]
MutualId -> MutualId
MutualId -> MutualId -> [MutualId]
MutualId -> MutualId -> MutualId -> [MutualId]
(MutualId -> MutualId)
-> (MutualId -> MutualId)
-> (Int -> MutualId)
-> (MutualId -> Int)
-> (MutualId -> [MutualId])
-> (MutualId -> MutualId -> [MutualId])
-> (MutualId -> MutualId -> [MutualId])
-> (MutualId -> MutualId -> MutualId -> [MutualId])
-> Enum MutualId
forall a.
(a -> a)
-> (a -> a)
-> (Int -> a)
-> (a -> Int)
-> (a -> [a])
-> (a -> a -> [a])
-> (a -> a -> [a])
-> (a -> a -> a -> [a])
-> Enum a
$csucc :: MutualId -> MutualId
succ :: MutualId -> MutualId
$cpred :: MutualId -> MutualId
pred :: MutualId -> MutualId
$ctoEnum :: Int -> MutualId
toEnum :: Int -> MutualId
$cfromEnum :: MutualId -> Int
fromEnum :: MutualId -> Int
$cenumFrom :: MutualId -> [MutualId]
enumFrom :: MutualId -> [MutualId]
$cenumFromThen :: MutualId -> MutualId -> [MutualId]
enumFromThen :: MutualId -> MutualId -> [MutualId]
$cenumFromTo :: MutualId -> MutualId -> [MutualId]
enumFromTo :: MutualId -> MutualId -> [MutualId]
$cenumFromThenTo :: MutualId -> MutualId -> MutualId -> [MutualId]
enumFromThenTo :: MutualId -> MutualId -> MutualId -> [MutualId]
Enum, MutualId -> ()
(MutualId -> ()) -> NFData MutualId
forall a. (a -> ()) -> NFData a
$crnf :: MutualId -> ()
rnf :: MutualId -> ()
NFData)
instance Pretty MutualId where
pretty :: MutualId -> Doc
pretty (MutualId Word32
i) = Word32 -> Doc
forall a. Pretty a => a -> Doc
pretty Word32
i
type MutualBlocks = IntMap MutualBlock
data MutualBlock = MutualBlock
{ MutualBlock -> MutualInfo
mutualInfo :: MutualInfo
, MutualBlock -> Set QName
mutualNames :: Set QName
} deriving (Int -> MutualBlock -> ShowS
[MutualBlock] -> ShowS
MutualBlock -> String
(Int -> MutualBlock -> ShowS)
-> (MutualBlock -> String)
-> ([MutualBlock] -> ShowS)
-> Show MutualBlock
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> MutualBlock -> ShowS
showsPrec :: Int -> MutualBlock -> ShowS
$cshow :: MutualBlock -> String
show :: MutualBlock -> String
$cshowList :: [MutualBlock] -> ShowS
showList :: [MutualBlock] -> ShowS
Show, MutualBlock -> MutualBlock -> Bool
(MutualBlock -> MutualBlock -> Bool)
-> (MutualBlock -> MutualBlock -> Bool) -> Eq MutualBlock
forall a. (a -> a -> Bool) -> (a -> a -> Bool) -> Eq a
$c== :: MutualBlock -> MutualBlock -> Bool
== :: MutualBlock -> MutualBlock -> Bool
$c/= :: MutualBlock -> MutualBlock -> Bool
/= :: MutualBlock -> MutualBlock -> Bool
Eq, (forall x. MutualBlock -> Rep MutualBlock x)
-> (forall x. Rep MutualBlock x -> MutualBlock)
-> Generic MutualBlock
forall x. Rep MutualBlock x -> MutualBlock
forall x. MutualBlock -> Rep MutualBlock x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. MutualBlock -> Rep MutualBlock x
from :: forall x. MutualBlock -> Rep MutualBlock x
$cto :: forall x. Rep MutualBlock x -> MutualBlock
to :: forall x. Rep MutualBlock x -> MutualBlock
Generic)
instance Null MutualBlock where
empty :: MutualBlock
empty = MutualInfo -> Set QName -> MutualBlock
MutualBlock MutualInfo
forall a. Null a => a
empty Set QName
forall a. Null a => a
empty
type Statistics = Map String Integer
data Call
= CheckClause Type A.SpineClause
| CheckLHS A.SpineLHS
| CheckPattern A.Pattern Telescope Type
| CheckPatternLinearityType C.Name
| CheckPatternLinearityValue C.Name
| CheckLetBinding A.LetBinding
| InferExpr A.Expr
| CheckExprCall Comparison A.Expr Type
| CheckDotPattern A.Expr Term
| CheckProjection Range QName Type
| IsTypeCall Comparison A.Expr Sort
| IsType_ A.Expr
| InferVar Name
| InferDef QName
| CheckArguments Range [NamedArg A.Expr] Type (Maybe Type)
| CheckMetaSolution Range MetaId Type Term
| CheckTargetType Range Type Type
| CheckDataDef Range QName [A.LamBinding] [A.Constructor]
| CheckRecDef Range QName [A.LamBinding] [A.Constructor]
| CheckConstructor QName Telescope Sort A.Constructor
| CheckConArgFitsIn QName Bool Type Sort
| CheckFunDefCall Range QName [A.Clause] Bool
| CheckPragma Range A.Pragma
| CheckPrimitive Range QName A.Expr
| CheckIsEmpty Range Type
| CheckConfluence QName QName
| CheckModuleParameters ModuleName A.Telescope
| CheckWithFunctionType Type
| CheckSectionApplication Range Erased ModuleName A.ModuleApplication
| CheckNamedWhere ModuleName
| CheckIApplyConfluence
Range
QName
Term
Term
Term
Type
| ScopeCheckExpr C.Expr
| ScopeCheckDeclaration NiceDeclaration
| ScopeCheckLHS C.QName C.Pattern
| NoHighlighting
| ModuleContents
| SetRange Range
deriving (forall x. Call -> Rep Call x)
-> (forall x. Rep Call x -> Call) -> Generic Call
forall x. Rep Call x -> Call
forall x. Call -> Rep Call x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. Call -> Rep Call x
from :: forall x. Call -> Rep Call x
$cto :: forall x. Rep Call x -> Call
to :: forall x. Rep Call x -> Call
Generic
instance Pretty Call where
pretty :: Call -> Doc
pretty CheckClause{} = Doc
"CheckClause"
pretty CheckLHS{} = Doc
"CheckLHS"
pretty CheckPattern{} = Doc
"CheckPattern"
pretty CheckPatternLinearityType{} = Doc
"CheckPatternLinearityType"
pretty CheckPatternLinearityValue{} = Doc
"CheckPatternLinearityValue"
pretty InferExpr{} = Doc
"InferExpr"
pretty CheckExprCall{} = Doc
"CheckExprCall"
pretty CheckLetBinding{} = Doc
"CheckLetBinding"
pretty CheckProjection{} = Doc
"CheckProjection"
pretty IsTypeCall{} = Doc
"IsTypeCall"
pretty IsType_{} = Doc
"IsType_"
pretty InferVar{} = Doc
"InferVar"
pretty InferDef{} = Doc
"InferDef"
pretty CheckArguments{} = Doc
"CheckArguments"
pretty CheckMetaSolution{} = Doc
"CheckMetaSolution"
pretty CheckTargetType{} = Doc
"CheckTargetType"
pretty CheckDataDef{} = Doc
"CheckDataDef"
pretty CheckRecDef{} = Doc
"CheckRecDef"
pretty CheckConstructor{} = Doc
"CheckConstructor"
pretty CheckConArgFitsIn{} = Doc
"CheckConArgFitsIn"
pretty CheckFunDefCall{} = Doc
"CheckFunDefCall"
pretty CheckPragma{} = Doc
"CheckPragma"
pretty CheckPrimitive{} = Doc
"CheckPrimitive"
pretty CheckModuleParameters{} = Doc
"CheckModuleParameters"
pretty CheckWithFunctionType{} = Doc
"CheckWithFunctionType"
pretty CheckNamedWhere{} = Doc
"CheckNamedWhere"
pretty ScopeCheckExpr{} = Doc
"ScopeCheckExpr"
pretty ScopeCheckDeclaration{} = Doc
"ScopeCheckDeclaration"
pretty ScopeCheckLHS{} = Doc
"ScopeCheckLHS"
pretty CheckDotPattern{} = Doc
"CheckDotPattern"
pretty SetRange{} = Doc
"SetRange"
pretty CheckSectionApplication{} = Doc
"CheckSectionApplication"
pretty CheckIsEmpty{} = Doc
"CheckIsEmpty"
pretty CheckConfluence{} = Doc
"CheckConfluence"
pretty NoHighlighting{} = Doc
"NoHighlighting"
pretty ModuleContents{} = Doc
"ModuleContents"
pretty CheckIApplyConfluence{} = Doc
"ModuleContents"
instance HasRange Call where
getRange :: Call -> Range
getRange (CheckClause Type
_ SpineClause
c) = SpineClause -> Range
forall a. HasRange a => a -> Range
getRange SpineClause
c
getRange (CheckLHS SpineLHS
lhs) = SpineLHS -> Range
forall a. HasRange a => a -> Range
getRange SpineLHS
lhs
getRange (CheckPattern Pattern
p Telescope
_ Type
_) = Pattern -> Range
forall a. HasRange a => a -> Range
getRange Pattern
p
getRange (CheckPatternLinearityType Name
x) = Name -> Range
forall a. HasRange a => a -> Range
getRange Name
x
getRange (CheckPatternLinearityValue Name
x) = Name -> Range
forall a. HasRange a => a -> Range
getRange Name
x
getRange (InferExpr Expr
e) = Expr -> Range
forall a. HasRange a => a -> Range
getRange Expr
e
getRange (CheckExprCall Comparison
_ Expr
e Type
_) = Expr -> Range
forall a. HasRange a => a -> Range
getRange Expr
e
getRange (CheckLetBinding LetBinding
b) = LetBinding -> Range
forall a. HasRange a => a -> Range
getRange LetBinding
b
getRange (CheckProjection Range
r QName
_ Type
_) = Range
r
getRange (IsTypeCall Comparison
cmp Expr
e Sort
s) = Expr -> Range
forall a. HasRange a => a -> Range
getRange Expr
e
getRange (IsType_ Expr
e) = Expr -> Range
forall a. HasRange a => a -> Range
getRange Expr
e
getRange (InferVar Name
x) = Name -> Range
forall a. HasRange a => a -> Range
getRange Name
x
getRange (InferDef QName
f) = QName -> Range
forall a. HasRange a => a -> Range
getRange QName
f
getRange (CheckArguments Range
r [NamedArg Expr]
_ Type
_ Maybe Type
_) = Range
r
getRange (CheckMetaSolution Range
r MetaId
_ Type
_ Term
_) = Range
r
getRange (CheckTargetType Range
r Type
_ Type
_) = Range
r
getRange (CheckDataDef Range
i QName
_ [LamBinding]
_ [Constructor]
_) = Range -> Range
forall a. HasRange a => a -> Range
getRange Range
i
getRange (CheckRecDef Range
i QName
_ [LamBinding]
_ [Constructor]
_) = Range -> Range
forall a. HasRange a => a -> Range
getRange Range
i
getRange (CheckConstructor QName
_ Telescope
_ Sort
_ Constructor
c) = Constructor -> Range
forall a. HasRange a => a -> Range
getRange Constructor
c
getRange (CheckConArgFitsIn QName
c Bool
_ Type
_ Sort
_) = QName -> Range
forall a. HasRange a => a -> Range
getRange QName
c
getRange (CheckFunDefCall Range
i QName
_ [Clause]
_ Bool
_) = Range -> Range
forall a. HasRange a => a -> Range
getRange Range
i
getRange (CheckPragma Range
r Pragma
_) = Range
r
getRange (CheckPrimitive Range
i QName
_ Expr
_) = Range -> Range
forall a. HasRange a => a -> Range
getRange Range
i
getRange (CheckModuleParameters ModuleName
_ Telescope
tel) = Telescope -> Range
forall a. HasRange a => a -> Range
getRange Telescope
tel
getRange CheckWithFunctionType{} = Range
forall a. Range' a
noRange
getRange (CheckNamedWhere ModuleName
m) = ModuleName -> Range
forall a. HasRange a => a -> Range
getRange ModuleName
m
getRange (ScopeCheckExpr Expr
e) = Expr -> Range
forall a. HasRange a => a -> Range
getRange Expr
e
getRange (ScopeCheckDeclaration NiceDeclaration
d) = NiceDeclaration -> Range
forall a. HasRange a => a -> Range
getRange NiceDeclaration
d
getRange (ScopeCheckLHS QName
_ Pattern
p) = Pattern -> Range
forall a. HasRange a => a -> Range
getRange Pattern
p
getRange (CheckDotPattern Expr
e Term
_) = Expr -> Range
forall a. HasRange a => a -> Range
getRange Expr
e
getRange (SetRange Range
r) = Range
r
getRange (CheckSectionApplication Range
r Erased
_ ModuleName
_ ModuleApplication
_) = Range
r
getRange (CheckIsEmpty Range
r Type
_) = Range
r
getRange (CheckConfluence QName
rule1 QName
rule2) = Range -> Range -> Range
forall a. Ord a => a -> a -> a
max (QName -> Range
forall a. HasRange a => a -> Range
getRange QName
rule1) (QName -> Range
forall a. HasRange a => a -> Range
getRange QName
rule2)
getRange Call
NoHighlighting = Range
forall a. Range' a
noRange
getRange Call
ModuleContents = Range
forall a. Range' a
noRange
getRange (CheckIApplyConfluence Range
e QName
_ Term
_ Term
_ Term
_ Type
_) = Range -> Range
forall a. HasRange a => a -> Range
getRange Range
e
data InstanceTable = InstanceTable
{ InstanceTable -> DiscrimTree QName
_itableTree :: DiscrimTree QName
, InstanceTable -> Map QName Int
_itableCounts :: Map QName Int
}
deriving (Int -> InstanceTable -> ShowS
[InstanceTable] -> ShowS
InstanceTable -> String
(Int -> InstanceTable -> ShowS)
-> (InstanceTable -> String)
-> ([InstanceTable] -> ShowS)
-> Show InstanceTable
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> InstanceTable -> ShowS
showsPrec :: Int -> InstanceTable -> ShowS
$cshow :: InstanceTable -> String
show :: InstanceTable -> String
$cshowList :: [InstanceTable] -> ShowS
showList :: [InstanceTable] -> ShowS
Show, (forall x. InstanceTable -> Rep InstanceTable x)
-> (forall x. Rep InstanceTable x -> InstanceTable)
-> Generic InstanceTable
forall x. Rep InstanceTable x -> InstanceTable
forall x. InstanceTable -> Rep InstanceTable x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. InstanceTable -> Rep InstanceTable x
from :: forall x. InstanceTable -> Rep InstanceTable x
$cto :: forall x. Rep InstanceTable x -> InstanceTable
to :: forall x. Rep InstanceTable x -> InstanceTable
Generic)
instance Semigroup InstanceTable where
InstanceTable DiscrimTree QName
t Map QName Int
i <> :: InstanceTable -> InstanceTable -> InstanceTable
<> InstanceTable DiscrimTree QName
t' Map QName Int
i' = InstanceTable
{ _itableTree :: DiscrimTree QName
_itableTree = DiscrimTree QName
t DiscrimTree QName -> DiscrimTree QName -> DiscrimTree QName
forall a. Semigroup a => a -> a -> a
<> DiscrimTree QName
t'
, _itableCounts :: Map QName Int
_itableCounts = (Int -> Int -> Int)
-> Map QName Int -> Map QName Int -> Map QName Int
forall k a. Ord k => (a -> a -> a) -> Map k a -> Map k a -> Map k a
Map.unionWith Int -> Int -> Int
forall a. Num a => a -> a -> a
(+) Map QName Int
i Map QName Int
i'
}
instance Monoid InstanceTable where
mempty :: InstanceTable
mempty = DiscrimTree QName -> Map QName Int -> InstanceTable
InstanceTable DiscrimTree QName
forall a. Monoid a => a
mempty Map QName Int
forall a. Monoid a => a
mempty
itableTree :: Lens' InstanceTable (DiscrimTree QName)
itableTree :: Lens' InstanceTable (DiscrimTree QName)
itableTree DiscrimTree QName -> f (DiscrimTree QName)
f InstanceTable
s = DiscrimTree QName -> f (DiscrimTree QName)
f (InstanceTable -> DiscrimTree QName
_itableTree InstanceTable
s) f (DiscrimTree QName)
-> (DiscrimTree QName -> InstanceTable) -> f InstanceTable
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \DiscrimTree QName
x -> InstanceTable
s { _itableTree = x }
itableCounts :: Lens' InstanceTable (Map QName Int)
itableCounts :: Lens' InstanceTable (Map QName Int)
itableCounts Map QName Int -> f (Map QName Int)
f InstanceTable
s = Map QName Int -> f (Map QName Int)
f (InstanceTable -> Map QName Int
_itableCounts InstanceTable
s) f (Map QName Int)
-> (Map QName Int -> InstanceTable) -> f InstanceTable
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \Map QName Int
x -> InstanceTable
s { _itableCounts = x }
type TempInstanceTable = (InstanceTable , Set QName)
data BuiltinSort
= SortUniv Univ
| SortOmega Univ
| SortIntervalUniv
| SortLevelUniv
deriving (Int -> BuiltinSort -> ShowS
[BuiltinSort] -> ShowS
BuiltinSort -> String
(Int -> BuiltinSort -> ShowS)
-> (BuiltinSort -> String)
-> ([BuiltinSort] -> ShowS)
-> Show BuiltinSort
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> BuiltinSort -> ShowS
showsPrec :: Int -> BuiltinSort -> ShowS
$cshow :: BuiltinSort -> String
show :: BuiltinSort -> String
$cshowList :: [BuiltinSort] -> ShowS
showList :: [BuiltinSort] -> ShowS
Show, BuiltinSort -> BuiltinSort -> Bool
(BuiltinSort -> BuiltinSort -> Bool)
-> (BuiltinSort -> BuiltinSort -> Bool) -> Eq BuiltinSort
forall a. (a -> a -> Bool) -> (a -> a -> Bool) -> Eq a
$c== :: BuiltinSort -> BuiltinSort -> Bool
== :: BuiltinSort -> BuiltinSort -> Bool
$c/= :: BuiltinSort -> BuiltinSort -> Bool
/= :: BuiltinSort -> BuiltinSort -> Bool
Eq, (forall x. BuiltinSort -> Rep BuiltinSort x)
-> (forall x. Rep BuiltinSort x -> BuiltinSort)
-> Generic BuiltinSort
forall x. Rep BuiltinSort x -> BuiltinSort
forall x. BuiltinSort -> Rep BuiltinSort x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. BuiltinSort -> Rep BuiltinSort x
from :: forall x. BuiltinSort -> Rep BuiltinSort x
$cto :: forall x. Rep BuiltinSort x -> BuiltinSort
to :: forall x. Rep BuiltinSort x -> BuiltinSort
Generic)
pattern SortProp, SortSet, SortStrictSet, SortPropOmega, SortSetOmega, SortStrictSetOmega :: BuiltinSort
pattern $mSortProp :: forall {r}. BuiltinSort -> ((# #) -> r) -> ((# #) -> r) -> r
$bSortProp :: BuiltinSort
SortProp = SortUniv UProp
pattern $mSortSet :: forall {r}. BuiltinSort -> ((# #) -> r) -> ((# #) -> r) -> r
$bSortSet :: BuiltinSort
SortSet = SortUniv UType
pattern $mSortStrictSet :: forall {r}. BuiltinSort -> ((# #) -> r) -> ((# #) -> r) -> r
$bSortStrictSet :: BuiltinSort
SortStrictSet = SortUniv USSet
pattern $mSortPropOmega :: forall {r}. BuiltinSort -> ((# #) -> r) -> ((# #) -> r) -> r
$bSortPropOmega :: BuiltinSort
SortPropOmega = SortOmega UProp
pattern $mSortSetOmega :: forall {r}. BuiltinSort -> ((# #) -> r) -> ((# #) -> r) -> r
$bSortSetOmega :: BuiltinSort
SortSetOmega = SortOmega UType
pattern $mSortStrictSetOmega :: forall {r}. BuiltinSort -> ((# #) -> r) -> ((# #) -> r) -> r
$bSortStrictSetOmega :: BuiltinSort
SortStrictSetOmega = SortOmega USSet
{-# COMPLETE
SortProp, SortSet, SortStrictSet,
SortPropOmega, SortSetOmega, SortStrictSetOmega,
SortIntervalUniv, SortLevelUniv #-}
data BuiltinDescriptor
= BuiltinData (TCM Type) [BuiltinId]
| BuiltinDataCons (TCM Type)
| BuiltinPrim PrimitiveId (Term -> TCM ())
| BuiltinSort BuiltinSort
| BuiltinPostulate Relevance (TCM Type)
| BuiltinUnknown (Maybe (TCM Type)) (Term -> Type -> TCM ())
data BuiltinInfo =
BuiltinInfo { BuiltinInfo -> BuiltinId
builtinName :: BuiltinId
, BuiltinInfo -> BuiltinDescriptor
builtinDesc :: BuiltinDescriptor }
type BuiltinThings = BuiltinThings' PrimFun
type BuiltinThings' pf = Map SomeBuiltin (Builtin pf)
data Builtin pf
= Builtin Term
| Prim pf
| BuiltinRewriteRelations (Set QName)
deriving (Int -> Builtin pf -> ShowS
[Builtin pf] -> ShowS
Builtin pf -> String
(Int -> Builtin pf -> ShowS)
-> (Builtin pf -> String)
-> ([Builtin pf] -> ShowS)
-> Show (Builtin pf)
forall pf. Show pf => Int -> Builtin pf -> ShowS
forall pf. Show pf => [Builtin pf] -> ShowS
forall pf. Show pf => Builtin pf -> String
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: forall pf. Show pf => Int -> Builtin pf -> ShowS
showsPrec :: Int -> Builtin pf -> ShowS
$cshow :: forall pf. Show pf => Builtin pf -> String
show :: Builtin pf -> String
$cshowList :: forall pf. Show pf => [Builtin pf] -> ShowS
showList :: [Builtin pf] -> ShowS
Show, (forall a b. (a -> b) -> Builtin a -> Builtin b)
-> (forall a b. a -> Builtin b -> Builtin a) -> Functor Builtin
forall a b. a -> Builtin b -> Builtin a
forall a b. (a -> b) -> Builtin a -> Builtin b
forall (f :: * -> *).
(forall a b. (a -> b) -> f a -> f b)
-> (forall a b. a -> f b -> f a) -> Functor f
$cfmap :: forall a b. (a -> b) -> Builtin a -> Builtin b
fmap :: forall a b. (a -> b) -> Builtin a -> Builtin b
$c<$ :: forall a b. a -> Builtin b -> Builtin a
<$ :: forall a b. a -> Builtin b -> Builtin a
Functor, (forall m. Monoid m => Builtin m -> m)
-> (forall m a. Monoid m => (a -> m) -> Builtin a -> m)
-> (forall m a. Monoid m => (a -> m) -> Builtin a -> m)
-> (forall a b. (a -> b -> b) -> b -> Builtin a -> b)
-> (forall a b. (a -> b -> b) -> b -> Builtin a -> b)
-> (forall b a. (b -> a -> b) -> b -> Builtin a -> b)
-> (forall b a. (b -> a -> b) -> b -> Builtin a -> b)
-> (forall a. (a -> a -> a) -> Builtin a -> a)
-> (forall a. (a -> a -> a) -> Builtin a -> a)
-> (forall a. Builtin a -> [a])
-> (forall a. Builtin a -> Bool)
-> (forall a. Builtin a -> Int)
-> (forall a. Eq a => a -> Builtin a -> Bool)
-> (forall a. Ord a => Builtin a -> a)
-> (forall a. Ord a => Builtin a -> a)
-> (forall a. Num a => Builtin a -> a)
-> (forall a. Num a => Builtin a -> a)
-> Foldable Builtin
forall a. Eq a => a -> Builtin a -> Bool
forall a. Num a => Builtin a -> a
forall a. Ord a => Builtin a -> a
forall m. Monoid m => Builtin m -> m
forall a. Builtin a -> Bool
forall a. Builtin a -> Int
forall a. Builtin a -> [a]
forall a. (a -> a -> a) -> Builtin a -> a
forall m a. Monoid m => (a -> m) -> Builtin a -> m
forall b a. (b -> a -> b) -> b -> Builtin a -> b
forall a b. (a -> b -> b) -> b -> Builtin a -> b
forall (t :: * -> *).
(forall m. Monoid m => t m -> m)
-> (forall m a. Monoid m => (a -> m) -> t a -> m)
-> (forall m a. Monoid m => (a -> m) -> t a -> m)
-> (forall a b. (a -> b -> b) -> b -> t a -> b)
-> (forall a b. (a -> b -> b) -> b -> t a -> b)
-> (forall b a. (b -> a -> b) -> b -> t a -> b)
-> (forall b a. (b -> a -> b) -> b -> t a -> b)
-> (forall a. (a -> a -> a) -> t a -> a)
-> (forall a. (a -> a -> a) -> t a -> a)
-> (forall a. t a -> [a])
-> (forall a. t a -> Bool)
-> (forall a. t a -> Int)
-> (forall a. Eq a => a -> t a -> Bool)
-> (forall a. Ord a => t a -> a)
-> (forall a. Ord a => t a -> a)
-> (forall a. Num a => t a -> a)
-> (forall a. Num a => t a -> a)
-> Foldable t
$cfold :: forall m. Monoid m => Builtin m -> m
fold :: forall m. Monoid m => Builtin m -> m
$cfoldMap :: forall m a. Monoid m => (a -> m) -> Builtin a -> m
foldMap :: forall m a. Monoid m => (a -> m) -> Builtin a -> m
$cfoldMap' :: forall m a. Monoid m => (a -> m) -> Builtin a -> m
foldMap' :: forall m a. Monoid m => (a -> m) -> Builtin a -> m
$cfoldr :: forall a b. (a -> b -> b) -> b -> Builtin a -> b
foldr :: forall a b. (a -> b -> b) -> b -> Builtin a -> b
$cfoldr' :: forall a b. (a -> b -> b) -> b -> Builtin a -> b
foldr' :: forall a b. (a -> b -> b) -> b -> Builtin a -> b
$cfoldl :: forall b a. (b -> a -> b) -> b -> Builtin a -> b
foldl :: forall b a. (b -> a -> b) -> b -> Builtin a -> b
$cfoldl' :: forall b a. (b -> a -> b) -> b -> Builtin a -> b
foldl' :: forall b a. (b -> a -> b) -> b -> Builtin a -> b
$cfoldr1 :: forall a. (a -> a -> a) -> Builtin a -> a
foldr1 :: forall a. (a -> a -> a) -> Builtin a -> a
$cfoldl1 :: forall a. (a -> a -> a) -> Builtin a -> a
foldl1 :: forall a. (a -> a -> a) -> Builtin a -> a
$ctoList :: forall a. Builtin a -> [a]
toList :: forall a. Builtin a -> [a]
$cnull :: forall a. Builtin a -> Bool
null :: forall a. Builtin a -> Bool
$clength :: forall a. Builtin a -> Int
length :: forall a. Builtin a -> Int
$celem :: forall a. Eq a => a -> Builtin a -> Bool
elem :: forall a. Eq a => a -> Builtin a -> Bool
$cmaximum :: forall a. Ord a => Builtin a -> a
maximum :: forall a. Ord a => Builtin a -> a
$cminimum :: forall a. Ord a => Builtin a -> a
minimum :: forall a. Ord a => Builtin a -> a
$csum :: forall a. Num a => Builtin a -> a
sum :: forall a. Num a => Builtin a -> a
$cproduct :: forall a. Num a => Builtin a -> a
product :: forall a. Num a => Builtin a -> a
Foldable, Functor Builtin
Foldable Builtin
(Functor Builtin, Foldable Builtin) =>
(forall (f :: * -> *) a b.
Applicative f =>
(a -> f b) -> Builtin a -> f (Builtin b))
-> (forall (f :: * -> *) a.
Applicative f =>
Builtin (f a) -> f (Builtin a))
-> (forall (m :: * -> *) a b.
Monad m =>
(a -> m b) -> Builtin a -> m (Builtin b))
-> (forall (m :: * -> *) a.
Monad m =>
Builtin (m a) -> m (Builtin a))
-> Traversable Builtin
forall (t :: * -> *).
(Functor t, Foldable t) =>
(forall (f :: * -> *) a b.
Applicative f =>
(a -> f b) -> t a -> f (t b))
-> (forall (f :: * -> *) a. Applicative f => t (f a) -> f (t a))
-> (forall (m :: * -> *) a b.
Monad m =>
(a -> m b) -> t a -> m (t b))
-> (forall (m :: * -> *) a. Monad m => t (m a) -> m (t a))
-> Traversable t
forall (m :: * -> *) a. Monad m => Builtin (m a) -> m (Builtin a)
forall (f :: * -> *) a.
Applicative f =>
Builtin (f a) -> f (Builtin a)
forall (m :: * -> *) a b.
Monad m =>
(a -> m b) -> Builtin a -> m (Builtin b)
forall (f :: * -> *) a b.
Applicative f =>
(a -> f b) -> Builtin a -> f (Builtin b)
$ctraverse :: forall (f :: * -> *) a b.
Applicative f =>
(a -> f b) -> Builtin a -> f (Builtin b)
traverse :: forall (f :: * -> *) a b.
Applicative f =>
(a -> f b) -> Builtin a -> f (Builtin b)
$csequenceA :: forall (f :: * -> *) a.
Applicative f =>
Builtin (f a) -> f (Builtin a)
sequenceA :: forall (f :: * -> *) a.
Applicative f =>
Builtin (f a) -> f (Builtin a)
$cmapM :: forall (m :: * -> *) a b.
Monad m =>
(a -> m b) -> Builtin a -> m (Builtin b)
mapM :: forall (m :: * -> *) a b.
Monad m =>
(a -> m b) -> Builtin a -> m (Builtin b)
$csequence :: forall (m :: * -> *) a. Monad m => Builtin (m a) -> m (Builtin a)
sequence :: forall (m :: * -> *) a. Monad m => Builtin (m a) -> m (Builtin a)
Traversable, (forall x. Builtin pf -> Rep (Builtin pf) x)
-> (forall x. Rep (Builtin pf) x -> Builtin pf)
-> Generic (Builtin pf)
forall x. Rep (Builtin pf) x -> Builtin pf
forall x. Builtin pf -> Rep (Builtin pf) x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
forall pf x. Rep (Builtin pf) x -> Builtin pf
forall pf x. Builtin pf -> Rep (Builtin pf) x
$cfrom :: forall pf x. Builtin pf -> Rep (Builtin pf) x
from :: forall x. Builtin pf -> Rep (Builtin pf) x
$cto :: forall pf x. Rep (Builtin pf) x -> Builtin pf
to :: forall x. Rep (Builtin pf) x -> Builtin pf
Generic)
ifTopLevelAndHighlightingLevelIsOr ::
MonadTCEnv tcm => HighlightingLevel -> Bool -> tcm () -> tcm ()
ifTopLevelAndHighlightingLevelIsOr :: forall (tcm :: * -> *).
MonadTCEnv tcm =>
HighlightingLevel -> Bool -> tcm () -> tcm ()
ifTopLevelAndHighlightingLevelIsOr HighlightingLevel
l Bool
b tcm ()
m = do
e <- tcm TCEnv
forall (m :: * -> *). MonadTCEnv m => m TCEnv
askTC
when (envHighlightingLevel e >= l || b) $
case (envImportPath e) of
(TopLevelModuleName
_:TopLevelModuleName
_:[TopLevelModuleName]
_) -> () -> tcm ()
forall a. a -> tcm a
forall (f :: * -> *) a. Applicative f => a -> f a
pure ()
[TopLevelModuleName]
_ -> tcm ()
m
ifTopLevelAndHighlightingLevelIs ::
MonadTCEnv tcm => HighlightingLevel -> tcm () -> tcm ()
ifTopLevelAndHighlightingLevelIs :: forall (tcm :: * -> *).
MonadTCEnv tcm =>
HighlightingLevel -> tcm () -> tcm ()
ifTopLevelAndHighlightingLevelIs HighlightingLevel
l =
HighlightingLevel -> Bool -> tcm () -> tcm ()
forall (tcm :: * -> *).
MonadTCEnv tcm =>
HighlightingLevel -> Bool -> tcm () -> tcm ()
ifTopLevelAndHighlightingLevelIsOr HighlightingLevel
l Bool
False
data TCEnv =
TCEnv { TCEnv -> Context
envContext :: Context
, TCEnv -> LetBindings
envLetBindings :: LetBindings
, TCEnv -> ModuleName
envCurrentModule :: ModuleName
, TCEnv -> Maybe AbsolutePath
envCurrentPath :: Maybe AbsolutePath
, TCEnv -> [(ModuleName, Int)]
envAnonymousModules :: [(ModuleName, Nat)]
, TCEnv -> [TopLevelModuleName]
envImportPath :: [TopLevelModuleName]
, TCEnv -> Maybe MutualId
envMutualBlock :: Maybe MutualId
, TCEnv -> TerminationCheck ()
envTerminationCheck :: TerminationCheck ()
, TCEnv -> CoverageCheck
envCoverageCheck :: CoverageCheck
, TCEnv -> Bool
envMakeCase :: Bool
, TCEnv -> Bool
envSolvingConstraints :: Bool
, TCEnv -> Bool
envCheckingWhere :: Bool
, TCEnv -> Bool
envWorkingOnTypes :: Bool
, TCEnv -> Bool
envAssignMetas :: Bool
, TCEnv -> Set ProblemId
envActiveProblems :: Set ProblemId
, TCEnv -> Maybe ProblemId
envUnquoteProblem :: Maybe ProblemId
, TCEnv -> AbstractMode
envAbstractMode :: AbstractMode
, TCEnv -> Relevance
envRelevance :: Relevance
, TCEnv -> Quantity
envQuantity :: Quantity
, TCEnv -> Bool
envHardCompileTimeMode :: Bool
, TCEnv -> Bool
envSplitOnStrict :: Bool
, TCEnv -> Bool
envDisplayFormsEnabled :: Bool
, TCEnv -> Bool
envFoldLetBindings :: Bool
, TCEnv -> Range
envRange :: Range
, TCEnv -> Range
envHighlightingRange :: Range
, TCEnv -> IPClause
envClause :: IPClause
, TCEnv -> Maybe (Closure Call)
envCall :: Maybe (Closure Call)
, TCEnv -> HighlightingLevel
envHighlightingLevel :: HighlightingLevel
, TCEnv -> HighlightingMethod
envHighlightingMethod :: HighlightingMethod
, TCEnv -> ExpandHidden
envExpandLast :: ExpandHidden
, TCEnv -> Maybe QName
envAppDef :: Maybe QName
, TCEnv -> Simplification
envSimplification :: Simplification
, TCEnv -> AllowedReductions
envAllowedReductions :: AllowedReductions
, TCEnv -> ReduceDefs
envReduceDefs :: ReduceDefs
, TCEnv -> Bool
envReconstructed :: Bool
, TCEnv -> Int
envInjectivityDepth :: Int
, TCEnv -> Bool
envCompareBlocked :: Bool
, TCEnv -> Bool
envPrintDomainFreePi :: Bool
, TCEnv -> Bool
envPrintMetasBare :: Bool
, TCEnv -> Bool
envInsideDotPattern :: Bool
, TCEnv -> UnquoteFlags
envUnquoteFlags :: UnquoteFlags
, TCEnv -> Int
envInstanceDepth :: !Int
, TCEnv -> Bool
envIsDebugPrinting :: Bool
, TCEnv -> [QName]
envPrintingPatternLambdas :: [QName]
, TCEnv -> Bool
envCallByNeed :: Bool
, TCEnv -> CheckpointId
envCurrentCheckpoint :: CheckpointId
, TCEnv -> Map CheckpointId Substitution
envCheckpoints :: Map CheckpointId Substitution
, TCEnv -> DoGeneralize
envGeneralizeMetas :: DoGeneralize
, TCEnv -> Map QName GeneralizedValue
envGeneralizedVars :: Map QName GeneralizedValue
, TCEnv -> Maybe BackendName
envActiveBackendName :: Maybe BackendName
, TCEnv -> Bool
envConflComputingOverlap :: Bool
, TCEnv -> Bool
envCurrentlyElaborating :: Bool
, TCEnv -> Maybe Int
envSyntacticEqualityFuel :: !(Strict.Maybe Int)
, TCEnv -> Maybe OpaqueId
envCurrentOpaqueId :: !(Maybe OpaqueId)
, TCEnv -> Bool
envTermCheckReducing :: Bool
}
deriving ((forall x. TCEnv -> Rep TCEnv x)
-> (forall x. Rep TCEnv x -> TCEnv) -> Generic TCEnv
forall x. Rep TCEnv x -> TCEnv
forall x. TCEnv -> Rep TCEnv x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. TCEnv -> Rep TCEnv x
from :: forall x. TCEnv -> Rep TCEnv x
$cto :: forall x. Rep TCEnv x -> TCEnv
to :: forall x. Rep TCEnv x -> TCEnv
Generic)
initEnv :: TCEnv
initEnv :: TCEnv
initEnv = TCEnv { envContext :: Context
envContext = []
, envLetBindings :: LetBindings
envLetBindings = LetBindings
forall k a. Map k a
Map.empty
, envCurrentModule :: ModuleName
envCurrentModule = ModuleName
noModuleName
, envCurrentPath :: Maybe AbsolutePath
envCurrentPath = Maybe AbsolutePath
forall a. Maybe a
Nothing
, envAnonymousModules :: [(ModuleName, Int)]
envAnonymousModules = []
, envImportPath :: [TopLevelModuleName]
envImportPath = []
, envMutualBlock :: Maybe MutualId
envMutualBlock = Maybe MutualId
forall a. Maybe a
Nothing
, envTerminationCheck :: TerminationCheck ()
envTerminationCheck = TerminationCheck ()
forall m. TerminationCheck m
TerminationCheck
, envCoverageCheck :: CoverageCheck
envCoverageCheck = CoverageCheck
YesCoverageCheck
, envMakeCase :: Bool
envMakeCase = Bool
False
, envSolvingConstraints :: Bool
envSolvingConstraints = Bool
False
, envCheckingWhere :: Bool
envCheckingWhere = Bool
False
, envActiveProblems :: Set ProblemId
envActiveProblems = Set ProblemId
forall a. Set a
Set.empty
, envUnquoteProblem :: Maybe ProblemId
envUnquoteProblem = Maybe ProblemId
forall a. Maybe a
Nothing
, envWorkingOnTypes :: Bool
envWorkingOnTypes = Bool
False
, envAssignMetas :: Bool
envAssignMetas = Bool
True
, envAbstractMode :: AbstractMode
envAbstractMode = AbstractMode
ConcreteMode
, envRelevance :: Relevance
envRelevance = Relevance
unitRelevance
, envQuantity :: Quantity
envQuantity = Quantity
unitQuantity
, envHardCompileTimeMode :: Bool
envHardCompileTimeMode = Bool
False
, envSplitOnStrict :: Bool
envSplitOnStrict = Bool
False
, envDisplayFormsEnabled :: Bool
envDisplayFormsEnabled = Bool
True
, envFoldLetBindings :: Bool
envFoldLetBindings = Bool
True
, envRange :: Range
envRange = Range
forall a. Range' a
noRange
, envHighlightingRange :: Range
envHighlightingRange = Range
forall a. Range' a
noRange
, envClause :: IPClause
envClause = IPClause
IPNoClause
, envCall :: Maybe (Closure Call)
envCall = Maybe (Closure Call)
forall a. Maybe a
Nothing
, envHighlightingLevel :: HighlightingLevel
envHighlightingLevel = HighlightingLevel
None
, envHighlightingMethod :: HighlightingMethod
envHighlightingMethod = HighlightingMethod
Indirect
, envExpandLast :: ExpandHidden
envExpandLast = ExpandHidden
ExpandLast
, envAppDef :: Maybe QName
envAppDef = Maybe QName
forall a. Maybe a
Nothing
, envSimplification :: Simplification
envSimplification = Simplification
NoSimplification
, envAllowedReductions :: AllowedReductions
envAllowedReductions = AllowedReductions
allReductions
, envReduceDefs :: ReduceDefs
envReduceDefs = ReduceDefs
reduceAllDefs
, envReconstructed :: Bool
envReconstructed = Bool
False
, envInjectivityDepth :: Int
envInjectivityDepth = Int
0
, envCompareBlocked :: Bool
envCompareBlocked = Bool
False
, envPrintDomainFreePi :: Bool
envPrintDomainFreePi = Bool
False
, envPrintMetasBare :: Bool
envPrintMetasBare = Bool
False
, envInsideDotPattern :: Bool
envInsideDotPattern = Bool
False
, envUnquoteFlags :: UnquoteFlags
envUnquoteFlags = UnquoteFlags
defaultUnquoteFlags
, envInstanceDepth :: Int
envInstanceDepth = Int
0
, envIsDebugPrinting :: Bool
envIsDebugPrinting = Bool
False
, envPrintingPatternLambdas :: [QName]
envPrintingPatternLambdas = []
, envCallByNeed :: Bool
envCallByNeed = Bool
True
, envCurrentCheckpoint :: CheckpointId
envCurrentCheckpoint = CheckpointId
0
, envCheckpoints :: Map CheckpointId Substitution
envCheckpoints = CheckpointId -> Substitution -> Map CheckpointId Substitution
forall k a. k -> a -> Map k a
Map.singleton CheckpointId
0 Substitution
forall a. Substitution' a
IdS
, envGeneralizeMetas :: DoGeneralize
envGeneralizeMetas = DoGeneralize
NoGeneralize
, envGeneralizedVars :: Map QName GeneralizedValue
envGeneralizedVars = Map QName GeneralizedValue
forall k a. Map k a
Map.empty
, envActiveBackendName :: Maybe BackendName
envActiveBackendName = Maybe BackendName
forall a. Maybe a
Nothing
, envConflComputingOverlap :: Bool
envConflComputingOverlap = Bool
False
, envCurrentlyElaborating :: Bool
envCurrentlyElaborating = Bool
False
, envSyntacticEqualityFuel :: Maybe Int
envSyntacticEqualityFuel = Maybe Int
forall a. Maybe a
Strict.Nothing
, envCurrentOpaqueId :: Maybe OpaqueId
envCurrentOpaqueId = Maybe OpaqueId
forall a. Maybe a
Nothing
, envTermCheckReducing :: Bool
envTermCheckReducing = Bool
False
}
class LensTCEnv a where
lensTCEnv :: Lens' a TCEnv
instance LensTCEnv TCEnv where
lensTCEnv :: Lens' TCEnv TCEnv
lensTCEnv = (TCEnv -> f TCEnv) -> TCEnv -> f TCEnv
forall a. a -> a
id
data UnquoteFlags = UnquoteFlags
{ UnquoteFlags -> Bool
_unquoteNormalise :: Bool }
deriving (forall x. UnquoteFlags -> Rep UnquoteFlags x)
-> (forall x. Rep UnquoteFlags x -> UnquoteFlags)
-> Generic UnquoteFlags
forall x. Rep UnquoteFlags x -> UnquoteFlags
forall x. UnquoteFlags -> Rep UnquoteFlags x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. UnquoteFlags -> Rep UnquoteFlags x
from :: forall x. UnquoteFlags -> Rep UnquoteFlags x
$cto :: forall x. Rep UnquoteFlags x -> UnquoteFlags
to :: forall x. Rep UnquoteFlags x -> UnquoteFlags
Generic
defaultUnquoteFlags :: UnquoteFlags
defaultUnquoteFlags :: UnquoteFlags
defaultUnquoteFlags = UnquoteFlags
{ _unquoteNormalise :: Bool
_unquoteNormalise = Bool
False }
unquoteNormalise :: Lens' UnquoteFlags Bool
unquoteNormalise :: Lens' UnquoteFlags Bool
unquoteNormalise Bool -> f Bool
f UnquoteFlags
e = Bool -> f Bool
f (UnquoteFlags -> Bool
_unquoteNormalise UnquoteFlags
e) f Bool -> (Bool -> UnquoteFlags) -> f UnquoteFlags
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ Bool
x -> UnquoteFlags
e { _unquoteNormalise = x }
eUnquoteNormalise :: Lens' TCEnv Bool
eUnquoteNormalise :: Lens' TCEnv Bool
eUnquoteNormalise = (UnquoteFlags -> f UnquoteFlags) -> TCEnv -> f TCEnv
Lens' TCEnv UnquoteFlags
eUnquoteFlags ((UnquoteFlags -> f UnquoteFlags) -> TCEnv -> f TCEnv)
-> ((Bool -> f Bool) -> UnquoteFlags -> f UnquoteFlags)
-> (Bool -> f Bool)
-> TCEnv
-> f TCEnv
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (Bool -> f Bool) -> UnquoteFlags -> f UnquoteFlags
Lens' UnquoteFlags Bool
unquoteNormalise
eContext :: Lens' TCEnv Context
eContext :: Lens' TCEnv Context
eContext Context -> f Context
f TCEnv
e = Context -> f Context
f (TCEnv -> Context
envContext TCEnv
e) f Context -> (Context -> TCEnv) -> f TCEnv
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ Context
x -> TCEnv
e { envContext = x }
eLetBindings :: Lens' TCEnv LetBindings
eLetBindings :: Lens' TCEnv LetBindings
eLetBindings LetBindings -> f LetBindings
f TCEnv
e = LetBindings -> f LetBindings
f (TCEnv -> LetBindings
envLetBindings TCEnv
e) f LetBindings -> (LetBindings -> TCEnv) -> f TCEnv
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ LetBindings
x -> TCEnv
e { envLetBindings = x }
eCurrentModule :: Lens' TCEnv ModuleName
eCurrentModule :: Lens' TCEnv ModuleName
eCurrentModule ModuleName -> f ModuleName
f TCEnv
e = ModuleName -> f ModuleName
f (TCEnv -> ModuleName
envCurrentModule TCEnv
e) f ModuleName -> (ModuleName -> TCEnv) -> f TCEnv
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ ModuleName
x -> TCEnv
e { envCurrentModule = x }
eCurrentPath :: Lens' TCEnv (Maybe AbsolutePath)
eCurrentPath :: Lens' TCEnv (Maybe AbsolutePath)
eCurrentPath Maybe AbsolutePath -> f (Maybe AbsolutePath)
f TCEnv
e = Maybe AbsolutePath -> f (Maybe AbsolutePath)
f (TCEnv -> Maybe AbsolutePath
envCurrentPath TCEnv
e) f (Maybe AbsolutePath) -> (Maybe AbsolutePath -> TCEnv) -> f TCEnv
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ Maybe AbsolutePath
x -> TCEnv
e { envCurrentPath = x }
eAnonymousModules :: Lens' TCEnv [(ModuleName, Nat)]
eAnonymousModules :: Lens' TCEnv [(ModuleName, Int)]
eAnonymousModules [(ModuleName, Int)] -> f [(ModuleName, Int)]
f TCEnv
e = [(ModuleName, Int)] -> f [(ModuleName, Int)]
f (TCEnv -> [(ModuleName, Int)]
envAnonymousModules TCEnv
e) f [(ModuleName, Int)] -> ([(ModuleName, Int)] -> TCEnv) -> f TCEnv
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ [(ModuleName, Int)]
x -> TCEnv
e { envAnonymousModules = x }
eImportPath :: Lens' TCEnv [TopLevelModuleName]
eImportPath :: Lens' TCEnv [TopLevelModuleName]
eImportPath [TopLevelModuleName] -> f [TopLevelModuleName]
f TCEnv
e = [TopLevelModuleName] -> f [TopLevelModuleName]
f (TCEnv -> [TopLevelModuleName]
envImportPath TCEnv
e) f [TopLevelModuleName]
-> ([TopLevelModuleName] -> TCEnv) -> f TCEnv
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ [TopLevelModuleName]
x -> TCEnv
e { envImportPath = x }
eMutualBlock :: Lens' TCEnv (Maybe MutualId)
eMutualBlock :: Lens' TCEnv (Maybe MutualId)
eMutualBlock Maybe MutualId -> f (Maybe MutualId)
f TCEnv
e = Maybe MutualId -> f (Maybe MutualId)
f (TCEnv -> Maybe MutualId
envMutualBlock TCEnv
e) f (Maybe MutualId) -> (Maybe MutualId -> TCEnv) -> f TCEnv
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ Maybe MutualId
x -> TCEnv
e { envMutualBlock = x }
eTerminationCheck :: Lens' TCEnv (TerminationCheck ())
eTerminationCheck :: Lens' TCEnv (TerminationCheck ())
eTerminationCheck TerminationCheck () -> f (TerminationCheck ())
f TCEnv
e = TerminationCheck () -> f (TerminationCheck ())
f (TCEnv -> TerminationCheck ()
envTerminationCheck TCEnv
e) f (TerminationCheck ())
-> (TerminationCheck () -> TCEnv) -> f TCEnv
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ TerminationCheck ()
x -> TCEnv
e { envTerminationCheck = x }
eCoverageCheck :: Lens' TCEnv CoverageCheck
eCoverageCheck :: Lens' TCEnv CoverageCheck
eCoverageCheck CoverageCheck -> f CoverageCheck
f TCEnv
e = CoverageCheck -> f CoverageCheck
f (TCEnv -> CoverageCheck
envCoverageCheck TCEnv
e) f CoverageCheck -> (CoverageCheck -> TCEnv) -> f TCEnv
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ CoverageCheck
x -> TCEnv
e { envCoverageCheck = x }
eMakeCase :: Lens' TCEnv Bool
eMakeCase :: Lens' TCEnv Bool
eMakeCase Bool -> f Bool
f TCEnv
e = Bool -> f Bool
f (TCEnv -> Bool
envMakeCase TCEnv
e) f Bool -> (Bool -> TCEnv) -> f TCEnv
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ Bool
x -> TCEnv
e { envMakeCase = x }
eSolvingConstraints :: Lens' TCEnv Bool
eSolvingConstraints :: Lens' TCEnv Bool
eSolvingConstraints Bool -> f Bool
f TCEnv
e = Bool -> f Bool
f (TCEnv -> Bool
envSolvingConstraints TCEnv
e) f Bool -> (Bool -> TCEnv) -> f TCEnv
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ Bool
x -> TCEnv
e { envSolvingConstraints = x }
eCheckingWhere :: Lens' TCEnv Bool
eCheckingWhere :: Lens' TCEnv Bool
eCheckingWhere Bool -> f Bool
f TCEnv
e = Bool -> f Bool
f (TCEnv -> Bool
envCheckingWhere TCEnv
e) f Bool -> (Bool -> TCEnv) -> f TCEnv
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ Bool
x -> TCEnv
e { envCheckingWhere = x }
eWorkingOnTypes :: Lens' TCEnv Bool
eWorkingOnTypes :: Lens' TCEnv Bool
eWorkingOnTypes Bool -> f Bool
f TCEnv
e = Bool -> f Bool
f (TCEnv -> Bool
envWorkingOnTypes TCEnv
e) f Bool -> (Bool -> TCEnv) -> f TCEnv
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ Bool
x -> TCEnv
e { envWorkingOnTypes = x }
eAssignMetas :: Lens' TCEnv Bool
eAssignMetas :: Lens' TCEnv Bool
eAssignMetas Bool -> f Bool
f TCEnv
e = Bool -> f Bool
f (TCEnv -> Bool
envAssignMetas TCEnv
e) f Bool -> (Bool -> TCEnv) -> f TCEnv
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ Bool
x -> TCEnv
e { envAssignMetas = x }
eActiveProblems :: Lens' TCEnv (Set ProblemId)
eActiveProblems :: Lens' TCEnv (Set ProblemId)
eActiveProblems Set ProblemId -> f (Set ProblemId)
f TCEnv
e = Set ProblemId -> f (Set ProblemId)
f (TCEnv -> Set ProblemId
envActiveProblems TCEnv
e) f (Set ProblemId) -> (Set ProblemId -> TCEnv) -> f TCEnv
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ Set ProblemId
x -> TCEnv
e { envActiveProblems = x }
eAbstractMode :: Lens' TCEnv AbstractMode
eAbstractMode :: Lens' TCEnv AbstractMode
eAbstractMode AbstractMode -> f AbstractMode
f TCEnv
e = AbstractMode -> f AbstractMode
f (TCEnv -> AbstractMode
envAbstractMode TCEnv
e) f AbstractMode -> (AbstractMode -> TCEnv) -> f TCEnv
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ AbstractMode
x -> TCEnv
e { envAbstractMode = x }
eRelevance :: Lens' TCEnv Relevance
eRelevance :: Lens' TCEnv Relevance
eRelevance Relevance -> f Relevance
f TCEnv
e = Relevance -> f Relevance
f (TCEnv -> Relevance
envRelevance TCEnv
e) f Relevance -> (Relevance -> TCEnv) -> f TCEnv
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \Relevance
x -> TCEnv
e { envRelevance = x }
eQuantity :: Lens' TCEnv Quantity
eQuantity :: Lens' TCEnv Quantity
eQuantity Quantity -> f Quantity
f TCEnv
e =
if TCEnv -> Bool
envHardCompileTimeMode TCEnv
e
then Quantity -> f Quantity
f (Quantity -> Quantity
forall {a}. LensQuantity a => a -> a
check (TCEnv -> Quantity
envQuantity TCEnv
e)) f Quantity -> (Quantity -> TCEnv) -> f TCEnv
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&>
\Quantity
x -> TCEnv
e { envQuantity = check x }
else Quantity -> f Quantity
f (TCEnv -> Quantity
envQuantity TCEnv
e) f Quantity -> (Quantity -> TCEnv) -> f TCEnv
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \Quantity
x -> TCEnv
e { envQuantity = x }
where
check :: a -> a
check a
q
| a -> Bool
forall a. LensQuantity a => a -> Bool
hasQuantity0 a
q = a
q
| Bool
otherwise = a
forall a. HasCallStack => a
__IMPOSSIBLE__
eHardCompileTimeMode :: Lens' TCEnv Bool
eHardCompileTimeMode :: Lens' TCEnv Bool
eHardCompileTimeMode Bool -> f Bool
f TCEnv
e = Bool -> f Bool
f (TCEnv -> Bool
envHardCompileTimeMode TCEnv
e) f Bool -> (Bool -> TCEnv) -> f TCEnv
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \Bool
x -> TCEnv
e { envHardCompileTimeMode = x }
eSplitOnStrict :: Lens' TCEnv Bool
eSplitOnStrict :: Lens' TCEnv Bool
eSplitOnStrict Bool -> f Bool
f TCEnv
e = Bool -> f Bool
f (TCEnv -> Bool
envSplitOnStrict TCEnv
e) f Bool -> (Bool -> TCEnv) -> f TCEnv
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ Bool
x -> TCEnv
e { envSplitOnStrict = x }
eDisplayFormsEnabled :: Lens' TCEnv Bool
eDisplayFormsEnabled :: Lens' TCEnv Bool
eDisplayFormsEnabled Bool -> f Bool
f TCEnv
e = Bool -> f Bool
f (TCEnv -> Bool
envDisplayFormsEnabled TCEnv
e) f Bool -> (Bool -> TCEnv) -> f TCEnv
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ Bool
x -> TCEnv
e { envDisplayFormsEnabled = x }
eFoldLetBindings :: Lens' TCEnv Bool
eFoldLetBindings :: Lens' TCEnv Bool
eFoldLetBindings Bool -> f Bool
f TCEnv
e = Bool -> f Bool
f (TCEnv -> Bool
envFoldLetBindings TCEnv
e) f Bool -> (Bool -> TCEnv) -> f TCEnv
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ Bool
x -> TCEnv
e { envFoldLetBindings = x }
eRange :: Lens' TCEnv Range
eRange :: Lens' TCEnv Range
eRange Range -> f Range
f TCEnv
e = Range -> f Range
f (TCEnv -> Range
envRange TCEnv
e) f Range -> (Range -> TCEnv) -> f TCEnv
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ Range
x -> TCEnv
e { envRange = x }
eHighlightingRange :: Lens' TCEnv Range
eHighlightingRange :: Lens' TCEnv Range
eHighlightingRange Range -> f Range
f TCEnv
e = Range -> f Range
f (TCEnv -> Range
envHighlightingRange TCEnv
e) f Range -> (Range -> TCEnv) -> f TCEnv
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ Range
x -> TCEnv
e { envHighlightingRange = x }
eCall :: Lens' TCEnv (Maybe (Closure Call))
eCall :: Lens' TCEnv (Maybe (Closure Call))
eCall Maybe (Closure Call) -> f (Maybe (Closure Call))
f TCEnv
e = Maybe (Closure Call) -> f (Maybe (Closure Call))
f (TCEnv -> Maybe (Closure Call)
envCall TCEnv
e) f (Maybe (Closure Call))
-> (Maybe (Closure Call) -> TCEnv) -> f TCEnv
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ Maybe (Closure Call)
x -> TCEnv
e { envCall = x }
eHighlightingLevel :: Lens' TCEnv HighlightingLevel
eHighlightingLevel :: Lens' TCEnv HighlightingLevel
eHighlightingLevel HighlightingLevel -> f HighlightingLevel
f TCEnv
e = HighlightingLevel -> f HighlightingLevel
f (TCEnv -> HighlightingLevel
envHighlightingLevel TCEnv
e) f HighlightingLevel -> (HighlightingLevel -> TCEnv) -> f TCEnv
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ HighlightingLevel
x -> TCEnv
e { envHighlightingLevel = x }
eHighlightingMethod :: Lens' TCEnv HighlightingMethod
eHighlightingMethod :: Lens' TCEnv HighlightingMethod
eHighlightingMethod HighlightingMethod -> f HighlightingMethod
f TCEnv
e = HighlightingMethod -> f HighlightingMethod
f (TCEnv -> HighlightingMethod
envHighlightingMethod TCEnv
e) f HighlightingMethod -> (HighlightingMethod -> TCEnv) -> f TCEnv
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ HighlightingMethod
x -> TCEnv
e { envHighlightingMethod = x }
eExpandLast :: Lens' TCEnv ExpandHidden
eExpandLast :: Lens' TCEnv ExpandHidden
eExpandLast ExpandHidden -> f ExpandHidden
f TCEnv
e = ExpandHidden -> f ExpandHidden
f (TCEnv -> ExpandHidden
envExpandLast TCEnv
e) f ExpandHidden -> (ExpandHidden -> TCEnv) -> f TCEnv
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ ExpandHidden
x -> TCEnv
e { envExpandLast = x }
eExpandLastBool :: Lens' TCEnv Bool
eExpandLastBool :: Lens' TCEnv Bool
eExpandLastBool Bool -> f Bool
f TCEnv
e = Bool -> f Bool
f (ExpandHidden -> Bool
isExpandLast (ExpandHidden -> Bool) -> ExpandHidden -> Bool
forall a b. (a -> b) -> a -> b
$ TCEnv -> ExpandHidden
envExpandLast TCEnv
e) f Bool -> (Bool -> TCEnv) -> f TCEnv
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ Bool
x -> TCEnv
e { envExpandLast = toExpandLast x }
eAppDef :: Lens' TCEnv (Maybe QName)
eAppDef :: Lens' TCEnv (Maybe QName)
eAppDef Maybe QName -> f (Maybe QName)
f TCEnv
e = Maybe QName -> f (Maybe QName)
f (TCEnv -> Maybe QName
envAppDef TCEnv
e) f (Maybe QName) -> (Maybe QName -> TCEnv) -> f TCEnv
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ Maybe QName
x -> TCEnv
e { envAppDef = x }
eSimplification :: Lens' TCEnv Simplification
eSimplification :: Lens' TCEnv Simplification
eSimplification Simplification -> f Simplification
f TCEnv
e = Simplification -> f Simplification
f (TCEnv -> Simplification
envSimplification TCEnv
e) f Simplification -> (Simplification -> TCEnv) -> f TCEnv
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ Simplification
x -> TCEnv
e { envSimplification = x }
eAllowedReductions :: Lens' TCEnv AllowedReductions
eAllowedReductions :: Lens' TCEnv AllowedReductions
eAllowedReductions AllowedReductions -> f AllowedReductions
f TCEnv
e = AllowedReductions -> f AllowedReductions
f (TCEnv -> AllowedReductions
envAllowedReductions TCEnv
e) f AllowedReductions -> (AllowedReductions -> TCEnv) -> f TCEnv
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ AllowedReductions
x -> TCEnv
e { envAllowedReductions = x }
eReduceDefs :: Lens' TCEnv ReduceDefs
eReduceDefs :: Lens' TCEnv ReduceDefs
eReduceDefs ReduceDefs -> f ReduceDefs
f TCEnv
e = ReduceDefs -> f ReduceDefs
f (TCEnv -> ReduceDefs
envReduceDefs TCEnv
e) f ReduceDefs -> (ReduceDefs -> TCEnv) -> f TCEnv
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ ReduceDefs
x -> TCEnv
e { envReduceDefs = x }
eReduceDefsPair :: Lens' TCEnv (Bool, [QName])
eReduceDefsPair :: Lens' TCEnv (Bool, [QName])
eReduceDefsPair (Bool, [QName]) -> f (Bool, [QName])
f TCEnv
e = (Bool, [QName]) -> f (Bool, [QName])
f (ReduceDefs -> (Bool, [QName])
fromReduceDefs (ReduceDefs -> (Bool, [QName])) -> ReduceDefs -> (Bool, [QName])
forall a b. (a -> b) -> a -> b
$ TCEnv -> ReduceDefs
envReduceDefs TCEnv
e) f (Bool, [QName]) -> ((Bool, [QName]) -> TCEnv) -> f TCEnv
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ (Bool, [QName])
x -> TCEnv
e { envReduceDefs = toReduceDefs x }
eReconstructed :: Lens' TCEnv Bool
eReconstructed :: Lens' TCEnv Bool
eReconstructed Bool -> f Bool
f TCEnv
e = Bool -> f Bool
f (TCEnv -> Bool
envReconstructed TCEnv
e) f Bool -> (Bool -> TCEnv) -> f TCEnv
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ Bool
x -> TCEnv
e { envReconstructed = x }
eInjectivityDepth :: Lens' TCEnv Int
eInjectivityDepth :: Lens' TCEnv Int
eInjectivityDepth Int -> f Int
f TCEnv
e = Int -> f Int
f (TCEnv -> Int
envInjectivityDepth TCEnv
e) f Int -> (Int -> TCEnv) -> f TCEnv
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ Int
x -> TCEnv
e { envInjectivityDepth = x }
eCompareBlocked :: Lens' TCEnv Bool
eCompareBlocked :: Lens' TCEnv Bool
eCompareBlocked Bool -> f Bool
f TCEnv
e = Bool -> f Bool
f (TCEnv -> Bool
envCompareBlocked TCEnv
e) f Bool -> (Bool -> TCEnv) -> f TCEnv
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ Bool
x -> TCEnv
e { envCompareBlocked = x }
ePrintDomainFreePi :: Lens' TCEnv Bool
ePrintDomainFreePi :: Lens' TCEnv Bool
ePrintDomainFreePi Bool -> f Bool
f TCEnv
e = Bool -> f Bool
f (TCEnv -> Bool
envPrintDomainFreePi TCEnv
e) f Bool -> (Bool -> TCEnv) -> f TCEnv
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ Bool
x -> TCEnv
e { envPrintDomainFreePi = x }
ePrintMetasBare :: Lens' TCEnv Bool
ePrintMetasBare :: Lens' TCEnv Bool
ePrintMetasBare Bool -> f Bool
f TCEnv
e = Bool -> f Bool
f (TCEnv -> Bool
envPrintMetasBare TCEnv
e) f Bool -> (Bool -> TCEnv) -> f TCEnv
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ Bool
x -> TCEnv
e { envPrintMetasBare = x }
eInsideDotPattern :: Lens' TCEnv Bool
eInsideDotPattern :: Lens' TCEnv Bool
eInsideDotPattern Bool -> f Bool
f TCEnv
e = Bool -> f Bool
f (TCEnv -> Bool
envInsideDotPattern TCEnv
e) f Bool -> (Bool -> TCEnv) -> f TCEnv
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ Bool
x -> TCEnv
e { envInsideDotPattern = x }
eUnquoteFlags :: Lens' TCEnv UnquoteFlags
eUnquoteFlags :: Lens' TCEnv UnquoteFlags
eUnquoteFlags UnquoteFlags -> f UnquoteFlags
f TCEnv
e = UnquoteFlags -> f UnquoteFlags
f (TCEnv -> UnquoteFlags
envUnquoteFlags TCEnv
e) f UnquoteFlags -> (UnquoteFlags -> TCEnv) -> f TCEnv
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ UnquoteFlags
x -> TCEnv
e { envUnquoteFlags = x }
eInstanceDepth :: Lens' TCEnv Int
eInstanceDepth :: Lens' TCEnv Int
eInstanceDepth Int -> f Int
f TCEnv
e = Int -> f Int
f (TCEnv -> Int
envInstanceDepth TCEnv
e) f Int -> (Int -> TCEnv) -> f TCEnv
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ Int
x -> TCEnv
e { envInstanceDepth = x }
eIsDebugPrinting :: Lens' TCEnv Bool
eIsDebugPrinting :: Lens' TCEnv Bool
eIsDebugPrinting Bool -> f Bool
f TCEnv
e = Bool -> f Bool
f (TCEnv -> Bool
envIsDebugPrinting TCEnv
e) f Bool -> (Bool -> TCEnv) -> f TCEnv
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ Bool
x -> TCEnv
e { envIsDebugPrinting = x }
ePrintingPatternLambdas :: Lens' TCEnv [QName]
ePrintingPatternLambdas :: Lens' TCEnv [QName]
ePrintingPatternLambdas [QName] -> f [QName]
f TCEnv
e = [QName] -> f [QName]
f (TCEnv -> [QName]
envPrintingPatternLambdas TCEnv
e) f [QName] -> ([QName] -> TCEnv) -> f TCEnv
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ [QName]
x -> TCEnv
e { envPrintingPatternLambdas = x }
eCallByNeed :: Lens' TCEnv Bool
eCallByNeed :: Lens' TCEnv Bool
eCallByNeed Bool -> f Bool
f TCEnv
e = Bool -> f Bool
f (TCEnv -> Bool
envCallByNeed TCEnv
e) f Bool -> (Bool -> TCEnv) -> f TCEnv
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ Bool
x -> TCEnv
e { envCallByNeed = x }
eCurrentCheckpoint :: Lens' TCEnv CheckpointId
eCurrentCheckpoint :: Lens' TCEnv CheckpointId
eCurrentCheckpoint CheckpointId -> f CheckpointId
f TCEnv
e = CheckpointId -> f CheckpointId
f (TCEnv -> CheckpointId
envCurrentCheckpoint TCEnv
e) f CheckpointId -> (CheckpointId -> TCEnv) -> f TCEnv
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ CheckpointId
x -> TCEnv
e { envCurrentCheckpoint = x }
eCheckpoints :: Lens' TCEnv (Map CheckpointId Substitution)
eCheckpoints :: Lens' TCEnv (Map CheckpointId Substitution)
eCheckpoints Map CheckpointId Substitution -> f (Map CheckpointId Substitution)
f TCEnv
e = Map CheckpointId Substitution -> f (Map CheckpointId Substitution)
f (TCEnv -> Map CheckpointId Substitution
envCheckpoints TCEnv
e) f (Map CheckpointId Substitution)
-> (Map CheckpointId Substitution -> TCEnv) -> f TCEnv
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ Map CheckpointId Substitution
x -> TCEnv
e { envCheckpoints = x }
eGeneralizeMetas :: Lens' TCEnv DoGeneralize
eGeneralizeMetas :: Lens' TCEnv DoGeneralize
eGeneralizeMetas DoGeneralize -> f DoGeneralize
f TCEnv
e = DoGeneralize -> f DoGeneralize
f (TCEnv -> DoGeneralize
envGeneralizeMetas TCEnv
e) f DoGeneralize -> (DoGeneralize -> TCEnv) -> f TCEnv
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ DoGeneralize
x -> TCEnv
e { envGeneralizeMetas = x }
eGeneralizedVars :: Lens' TCEnv (Map QName GeneralizedValue)
eGeneralizedVars :: Lens' TCEnv (Map QName GeneralizedValue)
eGeneralizedVars Map QName GeneralizedValue -> f (Map QName GeneralizedValue)
f TCEnv
e = Map QName GeneralizedValue -> f (Map QName GeneralizedValue)
f (TCEnv -> Map QName GeneralizedValue
envGeneralizedVars TCEnv
e) f (Map QName GeneralizedValue)
-> (Map QName GeneralizedValue -> TCEnv) -> f TCEnv
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ Map QName GeneralizedValue
x -> TCEnv
e { envGeneralizedVars = x }
eActiveBackendName :: Lens' TCEnv (Maybe BackendName)
eActiveBackendName :: Lens' TCEnv (Maybe BackendName)
eActiveBackendName Maybe BackendName -> f (Maybe BackendName)
f TCEnv
e = Maybe BackendName -> f (Maybe BackendName)
f (TCEnv -> Maybe BackendName
envActiveBackendName TCEnv
e) f (Maybe BackendName) -> (Maybe BackendName -> TCEnv) -> f TCEnv
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ Maybe BackendName
x -> TCEnv
e { envActiveBackendName = x }
eConflComputingOverlap :: Lens' TCEnv Bool
eConflComputingOverlap :: Lens' TCEnv Bool
eConflComputingOverlap Bool -> f Bool
f TCEnv
e = Bool -> f Bool
f (TCEnv -> Bool
envConflComputingOverlap TCEnv
e) f Bool -> (Bool -> TCEnv) -> f TCEnv
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ Bool
x -> TCEnv
e { envConflComputingOverlap = x }
eCurrentlyElaborating :: Lens' TCEnv Bool
eCurrentlyElaborating :: Lens' TCEnv Bool
eCurrentlyElaborating Bool -> f Bool
f TCEnv
e = Bool -> f Bool
f (TCEnv -> Bool
envCurrentlyElaborating TCEnv
e) f Bool -> (Bool -> TCEnv) -> f TCEnv
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ Bool
x -> TCEnv
e { envCurrentlyElaborating = x }
{-# SPECIALISE currentModality :: TCM Modality #-}
currentModality :: MonadTCEnv m => m Modality
currentModality :: forall (m :: * -> *). MonadTCEnv m => m Modality
currentModality = do
r <- Lens' TCEnv Relevance -> m Relevance
forall (m :: * -> *) a. MonadTCEnv m => Lens' TCEnv a -> m a
viewTC (Relevance -> f Relevance) -> TCEnv -> f TCEnv
Lens' TCEnv Relevance
eRelevance
q <- viewTC eQuantity
return Modality
{ modRelevance = r
, modQuantity = q
, modCohesion = unitCohesion
}
type LetBindings = Map Name (Open LetBinding)
data LetBinding = LetBinding { LetBinding -> Origin
letOrigin :: Origin
, LetBinding -> Term
letTerm :: Term
, LetBinding -> Dom Type
letType :: Dom Type
}
deriving (Int -> LetBinding -> ShowS
[LetBinding] -> ShowS
LetBinding -> String
(Int -> LetBinding -> ShowS)
-> (LetBinding -> String)
-> ([LetBinding] -> ShowS)
-> Show LetBinding
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> LetBinding -> ShowS
showsPrec :: Int -> LetBinding -> ShowS
$cshow :: LetBinding -> String
show :: LetBinding -> String
$cshowList :: [LetBinding] -> ShowS
showList :: [LetBinding] -> ShowS
Show, (forall x. LetBinding -> Rep LetBinding x)
-> (forall x. Rep LetBinding x -> LetBinding) -> Generic LetBinding
forall x. Rep LetBinding x -> LetBinding
forall x. LetBinding -> Rep LetBinding x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. LetBinding -> Rep LetBinding x
from :: forall x. LetBinding -> Rep LetBinding x
$cto :: forall x. Rep LetBinding x -> LetBinding
to :: forall x. Rep LetBinding x -> LetBinding
Generic)
onLetBindingType :: (Dom Type -> Dom Type) -> LetBinding -> LetBinding
onLetBindingType :: (Dom Type -> Dom Type) -> LetBinding -> LetBinding
onLetBindingType Dom Type -> Dom Type
f LetBinding
b = LetBinding
b { letType = f $ letType b }
data AbstractMode
= AbstractMode
| ConcreteMode
| IgnoreAbstractMode
deriving (Int -> AbstractMode -> ShowS
[AbstractMode] -> ShowS
AbstractMode -> String
(Int -> AbstractMode -> ShowS)
-> (AbstractMode -> String)
-> ([AbstractMode] -> ShowS)
-> Show AbstractMode
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> AbstractMode -> ShowS
showsPrec :: Int -> AbstractMode -> ShowS
$cshow :: AbstractMode -> String
show :: AbstractMode -> String
$cshowList :: [AbstractMode] -> ShowS
showList :: [AbstractMode] -> ShowS
Show, AbstractMode -> AbstractMode -> Bool
(AbstractMode -> AbstractMode -> Bool)
-> (AbstractMode -> AbstractMode -> Bool) -> Eq AbstractMode
forall a. (a -> a -> Bool) -> (a -> a -> Bool) -> Eq a
$c== :: AbstractMode -> AbstractMode -> Bool
== :: AbstractMode -> AbstractMode -> Bool
$c/= :: AbstractMode -> AbstractMode -> Bool
/= :: AbstractMode -> AbstractMode -> Bool
Eq, (forall x. AbstractMode -> Rep AbstractMode x)
-> (forall x. Rep AbstractMode x -> AbstractMode)
-> Generic AbstractMode
forall x. Rep AbstractMode x -> AbstractMode
forall x. AbstractMode -> Rep AbstractMode x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. AbstractMode -> Rep AbstractMode x
from :: forall x. AbstractMode -> Rep AbstractMode x
$cto :: forall x. Rep AbstractMode x -> AbstractMode
to :: forall x. Rep AbstractMode x -> AbstractMode
Generic)
aDefToMode :: IsAbstract -> AbstractMode
aDefToMode :: IsAbstract -> AbstractMode
aDefToMode IsAbstract
AbstractDef = AbstractMode
AbstractMode
aDefToMode IsAbstract
ConcreteDef = AbstractMode
ConcreteMode
aModeToDef :: AbstractMode -> Maybe IsAbstract
aModeToDef :: AbstractMode -> Maybe IsAbstract
aModeToDef AbstractMode
AbstractMode = IsAbstract -> Maybe IsAbstract
forall a. a -> Maybe a
Just IsAbstract
AbstractDef
aModeToDef AbstractMode
ConcreteMode = IsAbstract -> Maybe IsAbstract
forall a. a -> Maybe a
Just IsAbstract
ConcreteDef
aModeToDef AbstractMode
_ = Maybe IsAbstract
forall a. Maybe a
Nothing
data OpaqueBlock = OpaqueBlock
{ OpaqueBlock -> OpaqueId
opaqueId :: {-# UNPACK #-} !OpaqueId
, OpaqueBlock -> HashSet QName
opaqueUnfolding :: HashSet QName
, OpaqueBlock -> HashSet QName
opaqueDecls :: HashSet QName
, OpaqueBlock -> Maybe OpaqueId
opaqueParent :: Maybe OpaqueId
, OpaqueBlock -> Range
opaqueRange :: Range
} deriving (Int -> OpaqueBlock -> ShowS
[OpaqueBlock] -> ShowS
OpaqueBlock -> String
(Int -> OpaqueBlock -> ShowS)
-> (OpaqueBlock -> String)
-> ([OpaqueBlock] -> ShowS)
-> Show OpaqueBlock
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> OpaqueBlock -> ShowS
showsPrec :: Int -> OpaqueBlock -> ShowS
$cshow :: OpaqueBlock -> String
show :: OpaqueBlock -> String
$cshowList :: [OpaqueBlock] -> ShowS
showList :: [OpaqueBlock] -> ShowS
Show, (forall x. OpaqueBlock -> Rep OpaqueBlock x)
-> (forall x. Rep OpaqueBlock x -> OpaqueBlock)
-> Generic OpaqueBlock
forall x. Rep OpaqueBlock x -> OpaqueBlock
forall x. OpaqueBlock -> Rep OpaqueBlock x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. OpaqueBlock -> Rep OpaqueBlock x
from :: forall x. OpaqueBlock -> Rep OpaqueBlock x
$cto :: forall x. Rep OpaqueBlock x -> OpaqueBlock
to :: forall x. Rep OpaqueBlock x -> OpaqueBlock
Generic)
instance Pretty OpaqueBlock where
pretty :: OpaqueBlock -> Doc
pretty (OpaqueBlock OpaqueId
_ HashSet QName
uf HashSet QName
ds Maybe OpaqueId
p Range
_) = [Doc] -> Doc
forall (t :: * -> *). Foldable t => t Doc -> Doc
vcat
([Doc] -> Doc) -> [Doc] -> Doc
forall a b. (a -> b) -> a -> b
$ [ Doc
"opaque (extends " Doc -> Doc -> Doc
forall a. Semigroup a => a -> a -> a
<> Maybe OpaqueId -> Doc
forall a. Pretty a => a -> Doc
pretty Maybe OpaqueId
p Doc -> Doc -> Doc
forall a. Semigroup a => a -> a -> a
<> Doc
") {"
, Int -> Doc -> Doc
forall a. Int -> Doc a -> Doc a
nest Int
2 Doc
"unfolds"
]
[Doc] -> [Doc] -> [Doc]
forall a. [a] -> [a] -> [a]
++ [ Int -> Doc -> Doc
forall a. Int -> Doc a -> Doc a
nest Int
4 (QName -> Doc
forall a. Pretty a => a -> Doc
pretty QName
n Doc -> Doc -> Doc
forall a. Semigroup a => a -> a -> a
<> Doc
",") | QName
n <- [QName] -> [QName]
forall a. Ord a => [a] -> [a]
List.sort ([QName] -> [QName]) -> [QName] -> [QName]
forall a b. (a -> b) -> a -> b
$ HashSet QName -> [QName]
forall a. HashSet a -> [a]
HashSet.toList HashSet QName
uf ]
[Doc] -> [Doc] -> [Doc]
forall a. [a] -> [a] -> [a]
++ [ Int -> Doc -> Doc
forall a. Int -> Doc a -> Doc a
nest Int
2 Doc
"declares" ]
[Doc] -> [Doc] -> [Doc]
forall a. [a] -> [a] -> [a]
++ [ Int -> Doc -> Doc
forall a. Int -> Doc a -> Doc a
nest Int
4 (QName -> Doc
forall a. Pretty a => a -> Doc
pretty QName
n Doc -> Doc -> Doc
forall a. Doc a -> Doc a -> Doc a
<+> Doc
": _") | QName
n <- [QName] -> [QName]
forall a. Ord a => [a] -> [a]
List.sort ([QName] -> [QName]) -> [QName] -> [QName]
forall a b. (a -> b) -> a -> b
$ HashSet QName -> [QName]
forall a. HashSet a -> [a]
HashSet.toList HashSet QName
ds ]
[Doc] -> [Doc] -> [Doc]
forall a. [a] -> [a] -> [a]
++ [ Doc
"}" ]
instance Eq OpaqueBlock where
OpaqueBlock
xs == :: OpaqueBlock -> OpaqueBlock -> Bool
== OpaqueBlock
ys = OpaqueBlock -> OpaqueId
opaqueId OpaqueBlock
xs OpaqueId -> OpaqueId -> Bool
forall a. Eq a => a -> a -> Bool
== OpaqueBlock -> OpaqueId
opaqueId OpaqueBlock
ys
instance Hashable OpaqueBlock where
hashWithSalt :: Int -> OpaqueBlock -> Int
hashWithSalt Int
s = Int -> OpaqueId -> Int
forall a. Hashable a => Int -> a -> Int
hashWithSalt Int
s (OpaqueId -> Int)
-> (OpaqueBlock -> OpaqueId) -> OpaqueBlock -> Int
forall b c a. (b -> c) -> (a -> b) -> a -> c
. OpaqueBlock -> OpaqueId
opaqueId
data ExpandHidden
= ExpandLast
| DontExpandLast
| ReallyDontExpandLast
deriving (ExpandHidden -> ExpandHidden -> Bool
(ExpandHidden -> ExpandHidden -> Bool)
-> (ExpandHidden -> ExpandHidden -> Bool) -> Eq ExpandHidden
forall a. (a -> a -> Bool) -> (a -> a -> Bool) -> Eq a
$c== :: ExpandHidden -> ExpandHidden -> Bool
== :: ExpandHidden -> ExpandHidden -> Bool
$c/= :: ExpandHidden -> ExpandHidden -> Bool
/= :: ExpandHidden -> ExpandHidden -> Bool
Eq, (forall x. ExpandHidden -> Rep ExpandHidden x)
-> (forall x. Rep ExpandHidden x -> ExpandHidden)
-> Generic ExpandHidden
forall x. Rep ExpandHidden x -> ExpandHidden
forall x. ExpandHidden -> Rep ExpandHidden x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. ExpandHidden -> Rep ExpandHidden x
from :: forall x. ExpandHidden -> Rep ExpandHidden x
$cto :: forall x. Rep ExpandHidden x -> ExpandHidden
to :: forall x. Rep ExpandHidden x -> ExpandHidden
Generic)
isExpandLast :: ExpandHidden -> Bool
isExpandLast :: ExpandHidden -> Bool
isExpandLast ExpandHidden
ExpandLast = Bool
True
isExpandLast ExpandHidden
DontExpandLast = Bool
False
isExpandLast ExpandHidden
ReallyDontExpandLast = Bool
False
isDontExpandLast :: ExpandHidden -> Bool
isDontExpandLast :: ExpandHidden -> Bool
isDontExpandLast = Bool -> Bool
not (Bool -> Bool) -> (ExpandHidden -> Bool) -> ExpandHidden -> Bool
forall b c a. (b -> c) -> (a -> b) -> a -> c
. ExpandHidden -> Bool
isExpandLast
toExpandLast :: Bool -> ExpandHidden
toExpandLast :: Bool -> ExpandHidden
toExpandLast Bool
True = ExpandHidden
ExpandLast
toExpandLast Bool
False = ExpandHidden
DontExpandLast
data CandidateKind
= LocalCandidate
| GlobalCandidate QName
deriving (Int -> CandidateKind -> ShowS
[CandidateKind] -> ShowS
CandidateKind -> String
(Int -> CandidateKind -> ShowS)
-> (CandidateKind -> String)
-> ([CandidateKind] -> ShowS)
-> Show CandidateKind
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> CandidateKind -> ShowS
showsPrec :: Int -> CandidateKind -> ShowS
$cshow :: CandidateKind -> String
show :: CandidateKind -> String
$cshowList :: [CandidateKind] -> ShowS
showList :: [CandidateKind] -> ShowS
Show, (forall x. CandidateKind -> Rep CandidateKind x)
-> (forall x. Rep CandidateKind x -> CandidateKind)
-> Generic CandidateKind
forall x. Rep CandidateKind x -> CandidateKind
forall x. CandidateKind -> Rep CandidateKind x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. CandidateKind -> Rep CandidateKind x
from :: forall x. CandidateKind -> Rep CandidateKind x
$cto :: forall x. Rep CandidateKind x -> CandidateKind
to :: forall x. Rep CandidateKind x -> CandidateKind
Generic)
data Candidate = Candidate
{ Candidate -> CandidateKind
candidateKind :: CandidateKind
, Candidate -> Term
candidateTerm :: Term
, Candidate -> Type
candidateType :: Type
, Candidate -> OverlapMode
candidateOverlap :: OverlapMode
}
deriving (Int -> Candidate -> ShowS
[Candidate] -> ShowS
Candidate -> String
(Int -> Candidate -> ShowS)
-> (Candidate -> String)
-> ([Candidate] -> ShowS)
-> Show Candidate
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> Candidate -> ShowS
showsPrec :: Int -> Candidate -> ShowS
$cshow :: Candidate -> String
show :: Candidate -> String
$cshowList :: [Candidate] -> ShowS
showList :: [Candidate] -> ShowS
Show, (forall x. Candidate -> Rep Candidate x)
-> (forall x. Rep Candidate x -> Candidate) -> Generic Candidate
forall x. Rep Candidate x -> Candidate
forall x. Candidate -> Rep Candidate x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. Candidate -> Rep Candidate x
from :: forall x. Candidate -> Rep Candidate x
$cto :: forall x. Rep Candidate x -> Candidate
to :: forall x. Rep Candidate x -> Candidate
Generic)
instance Free Candidate where
freeVars' :: forall a c. IsVarSet a c => Candidate -> FreeM a c
freeVars' (Candidate CandidateKind
_ Term
t Type
u OverlapMode
_) = (Term, Type) -> FreeM a c
forall t a c. (Free t, IsVarSet a c) => t -> FreeM a c
forall a c. IsVarSet a c => (Term, Type) -> FreeM a c
freeVars' (Term
t, Type
u)
instance HasOverlapMode Candidate where
lensOverlapMode :: Lens' Candidate OverlapMode
lensOverlapMode OverlapMode -> f OverlapMode
f Candidate
x = OverlapMode -> f OverlapMode
f (Candidate -> OverlapMode
candidateOverlap Candidate
x) f OverlapMode -> (OverlapMode -> Candidate) -> f Candidate
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \OverlapMode
m -> Candidate
x{ candidateOverlap = m }
data CheckedArg = CheckedArg
{ CheckedArg -> Elim
caElim :: Elim
, CheckedArg -> Maybe Range
caRange :: Maybe Range
, CheckedArg -> Maybe (Abs Constraint)
caConstraint :: Maybe (Abs Constraint)
}
deriving Int -> CheckedArg -> ShowS
[CheckedArg] -> ShowS
CheckedArg -> String
(Int -> CheckedArg -> ShowS)
-> (CheckedArg -> String)
-> ([CheckedArg] -> ShowS)
-> Show CheckedArg
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> CheckedArg -> ShowS
showsPrec :: Int -> CheckedArg -> ShowS
$cshow :: CheckedArg -> String
show :: CheckedArg -> String
$cshowList :: [CheckedArg] -> ShowS
showList :: [CheckedArg] -> ShowS
Show
data ArgsCheckState a = ACState
{ forall a. ArgsCheckState a -> [CheckedArg]
acCheckedArgs :: [CheckedArg]
, forall a. ArgsCheckState a -> Type
acType :: Type
, forall a. ArgsCheckState a -> a
acData :: a
}
deriving Int -> ArgsCheckState a -> ShowS
[ArgsCheckState a] -> ShowS
ArgsCheckState a -> String
(Int -> ArgsCheckState a -> ShowS)
-> (ArgsCheckState a -> String)
-> ([ArgsCheckState a] -> ShowS)
-> Show (ArgsCheckState a)
forall a. Show a => Int -> ArgsCheckState a -> ShowS
forall a. Show a => [ArgsCheckState a] -> ShowS
forall a. Show a => ArgsCheckState a -> String
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: forall a. Show a => Int -> ArgsCheckState a -> ShowS
showsPrec :: Int -> ArgsCheckState a -> ShowS
$cshow :: forall a. Show a => ArgsCheckState a -> String
show :: ArgsCheckState a -> String
$cshowList :: forall a. Show a => [ArgsCheckState a] -> ShowS
showList :: [ArgsCheckState a] -> ShowS
Show
data Warning
= NicifierIssue DeclarationWarning
| TerminationIssue (List1 TerminationError)
| UnreachableClauses QName (List1 Range)
| CoverageIssue QName (List1 (Telescope, [NamedArg DeBruijnPattern]))
| CoverageNoExactSplit QName (List1 Clause)
| InlineNoExactSplit QName Clause
| NotStrictlyPositive QName (Seq OccursWhere)
| ConstructorDoesNotFitInData QName Sort Sort TCErr
| CoinductiveEtaRecord QName
| UnsolvedMetaVariables (Set1 Range)
| UnsolvedInteractionMetas (Set1 Range)
| UnsolvedConstraints (List1 ProblemConstraint)
| InteractionMetaBoundaries (Set1 Range)
| CantGeneralizeOverSorts (Set1 MetaId)
| AbsurdPatternRequiresAbsentRHS
| OldBuiltin BuiltinId BuiltinId
| BuiltinDeclaresIdentifier BuiltinId
| DuplicateRecordDirective C.RecordDirective
| EmptyRewritePragma
| EmptyWhere
| FixingRelevance String Relevance Relevance
| IllformedAsClause String
| InvalidCharacterLiteral Char
| ClashesViaRenaming NameOrModule (Set1 C.Name)
| UselessPatternDeclarationForRecord String
| UselessPragma Range Doc
| UselessPublic
| UselessHiding (List1 C.ImportedName)
| UselessInline QName
| WrongInstanceDeclaration
| InstanceWithExplicitArg QName
| InstanceNoOutputTypeName Doc
| InstanceArgWithExplicitArg Doc
| InversionDepthReached QName
| NoGuardednessFlag QName
| SafeFlagPostulate C.Name
| SafeFlagPragma (Set String)
| SafeFlagWithoutKFlagPrimEraseEquality
| WithoutKFlagPrimEraseEquality
| ConflictingPragmaOptions String String
| OptionWarning OptionWarning
| ParseWarning ParseWarning
| LibraryWarning LibWarning
| DeprecationWarning String String String
| UserWarning Text
| DuplicateUsing (List1 C.ImportedName)
| FixityInRenamingModule (List1 Range)
| ModuleDoesntExport C.QName [C.Name] [C.Name] (List1 C.ImportedName)
| InfectiveImport Doc
| CoInfectiveImport Doc
| ConfluenceCheckingIncompleteBecauseOfMeta QName
| ConfluenceForCubicalNotSupported
| NotARewriteRule C.QName IsAmbiguous
| IllegalRewriteRule QName IllegalRewriteRuleReason
| RewriteNonConfluent Term Term Term Doc
| RewriteMaybeNonConfluent Term Term [Doc]
| RewriteAmbiguousRules Term Term Term
| RewriteMissingRule Term Term Term
| PragmaCompileErased BackendName QName
| PragmaCompileList
| PragmaCompileMaybe
| PragmaCompileUnparsable String
| PragmaCompileWrong QName String
| PragmaCompileWrongName C.QName IsAmbiguous
| PragmaExpectsDefinedSymbol String C.QName
| PragmaExpectsUnambiguousConstructorOrFunction String C.QName IsAmbiguous
| PragmaExpectsUnambiguousProjectionOrFunction String C.QName IsAmbiguous
| NoMain TopLevelModuleName
| NotInScopeW C.QName
| UnsupportedIndexedMatch Doc
| AsPatternShadowsConstructorOrPatternSynonym ConstructorOrPatternSynonym
| PatternShadowsConstructor C.Name A.QName
| PlentyInHardCompileTimeMode QωOrigin
| RecordFieldWarning RecordFieldWarning
| MissingTypeSignatureForOpaque QName IsOpaque
| NotAffectedByOpaque
| UnfoldingWrongName C.QName
| UnfoldTransparentName QName
| UselessOpaque
| HiddenNotInArgumentPosition C.Expr
| InstanceNotInArgumentPosition C.Expr
| MacroInLetBindings
| AbstractInLetBindings
| InvalidDisplayForm QName String
| UnusedVariablesInDisplayForm (List1 A.Name)
| TooManyArgumentsToSort QName (List1 (NamedArg A.Expr))
| WithClauseProjectionFixityMismatch
{ Warning -> NamedArg Pattern
withClausePattern :: NamedArg A.Pattern
, Warning -> ProjOrigin
withClauseProjectionOrigin :: ProjOrigin
, Warning -> NamedArg DeBruijnPattern
parentPattern :: NamedArg DeBruijnPattern
, Warning -> ProjOrigin
parentProjectionOrigin :: ProjOrigin
}
| FaceConstraintCannotBeHidden ArgInfo
| FaceConstraintCannotBeNamed NamedName
| DuplicateInterfaceFiles AbsolutePath AbsolutePath
| CustomBackendWarning String Doc
deriving (Int -> Warning -> ShowS
[Warning] -> ShowS
Warning -> String
(Int -> Warning -> ShowS)
-> (Warning -> String) -> ([Warning] -> ShowS) -> Show Warning
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> Warning -> ShowS
showsPrec :: Int -> Warning -> ShowS
$cshow :: Warning -> String
show :: Warning -> String
$cshowList :: [Warning] -> ShowS
showList :: [Warning] -> ShowS
Show, (forall x. Warning -> Rep Warning x)
-> (forall x. Rep Warning x -> Warning) -> Generic Warning
forall x. Rep Warning x -> Warning
forall x. Warning -> Rep Warning x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. Warning -> Rep Warning x
from :: forall x. Warning -> Rep Warning x
$cto :: forall x. Rep Warning x -> Warning
to :: forall x. Rep Warning x -> Warning
Generic)
recordFieldWarningToError :: RecordFieldWarning -> TypeError
recordFieldWarningToError :: RecordFieldWarning -> TypeError
recordFieldWarningToError = \case
W.DuplicateFields List1 (Name, Range)
xrs -> List1 Name -> TypeError
DuplicateFields (List1 Name -> TypeError) -> List1 Name -> TypeError
forall a b. (a -> b) -> a -> b
$ ((Name, Range) -> Name) -> List1 (Name, Range) -> List1 Name
forall a b. (a -> b) -> NonEmpty a -> NonEmpty b
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
fmap (Name, Range) -> Name
forall a b. (a, b) -> a
fst List1 (Name, Range)
xrs
W.TooManyFields QName
q [Name]
ys List1 (Name, Range)
xrs -> QName -> [Name] -> List1 Name -> TypeError
TooManyFields QName
q [Name]
ys (List1 Name -> TypeError) -> List1 Name -> TypeError
forall a b. (a -> b) -> a -> b
$ ((Name, Range) -> Name) -> List1 (Name, Range) -> List1 Name
forall a b. (a -> b) -> NonEmpty a -> NonEmpty b
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
fmap (Name, Range) -> Name
forall a b. (a, b) -> a
fst List1 (Name, Range)
xrs
warningName :: Warning -> WarningName
warningName :: Warning -> WarningName
warningName = \case
NicifierIssue DeclarationWarning
dw -> DeclarationWarning -> WarningName
declarationWarningName DeclarationWarning
dw
OptionWarning OptionWarning
ow -> OptionWarning -> WarningName
optionWarningName OptionWarning
ow
ParseWarning ParseWarning
pw -> ParseWarning -> WarningName
parseWarningName ParseWarning
pw
LibraryWarning LibWarning
lw -> LibWarning -> WarningName
libraryWarningName LibWarning
lw
AsPatternShadowsConstructorOrPatternSynonym{} -> WarningName
AsPatternShadowsConstructorOrPatternSynonym_
PatternShadowsConstructor{} -> WarningName
PatternShadowsConstructor_
AbsurdPatternRequiresAbsentRHS{} -> WarningName
AbsurdPatternRequiresAbsentRHS_
CantGeneralizeOverSorts{} -> WarningName
CantGeneralizeOverSorts_
CoverageIssue{} -> WarningName
CoverageIssue_
CoverageNoExactSplit{} -> WarningName
CoverageNoExactSplit_
InlineNoExactSplit{} -> WarningName
InlineNoExactSplit_
DeprecationWarning{} -> WarningName
DeprecationWarning_
DuplicateRecordDirective{} -> WarningName
DuplicateRecordDirective_
Warning
EmptyRewritePragma -> WarningName
EmptyRewritePragma_
Warning
EmptyWhere -> WarningName
EmptyWhere_
FixingRelevance{} -> WarningName
FixingRelevance_
IllformedAsClause{} -> WarningName
IllformedAsClause_
WrongInstanceDeclaration{} -> WarningName
WrongInstanceDeclaration_
InstanceWithExplicitArg{} -> WarningName
InstanceWithExplicitArg_
InstanceNoOutputTypeName{} -> WarningName
InstanceNoOutputTypeName_
InstanceArgWithExplicitArg{} -> WarningName
InstanceArgWithExplicitArg_
DuplicateUsing{} -> WarningName
DuplicateUsing_
FixityInRenamingModule{} -> WarningName
FixityInRenamingModule_
InvalidCharacterLiteral{} -> WarningName
InvalidCharacterLiteral_
UselessPragma{} -> WarningName
UselessPragma_
InversionDepthReached{} -> WarningName
InversionDepthReached_
InteractionMetaBoundaries{} -> InteractionMetaBoundaries_{}
ModuleDoesntExport{} -> WarningName
ModuleDoesntExport_
NoGuardednessFlag{} -> WarningName
NoGuardednessFlag_
NotInScopeW{} -> WarningName
NotInScope_
NotStrictlyPositive{} -> WarningName
NotStrictlyPositive_
ConstructorDoesNotFitInData{}-> WarningName
ConstructorDoesNotFitInData_
CoinductiveEtaRecord{} -> WarningName
CoinductiveEtaRecord_
UnsupportedIndexedMatch{} -> WarningName
UnsupportedIndexedMatch_
OldBuiltin{} -> WarningName
OldBuiltin_
BuiltinDeclaresIdentifier{} -> WarningName
BuiltinDeclaresIdentifier_
SafeFlagPostulate{} -> WarningName
SafeFlagPostulate_
SafeFlagPragma{} -> WarningName
SafeFlagPragma_
Warning
SafeFlagWithoutKFlagPrimEraseEquality -> WarningName
SafeFlagWithoutKFlagPrimEraseEquality_
ConflictingPragmaOptions{} -> WarningName
ConflictingPragmaOptions_
Warning
WithoutKFlagPrimEraseEquality -> WarningName
WithoutKFlagPrimEraseEquality_
TerminationIssue{} -> WarningName
TerminationIssue_
UnreachableClauses{} -> WarningName
UnreachableClauses_
UnsolvedInteractionMetas{} -> WarningName
UnsolvedInteractionMetas_
UnsolvedConstraints{} -> WarningName
UnsolvedConstraints_
UnsolvedMetaVariables{} -> WarningName
UnsolvedMetaVariables_
UselessHiding{} -> WarningName
UselessHiding_
UselessInline{} -> WarningName
UselessInline_
UselessPublic{} -> WarningName
UselessPublic_
UselessPatternDeclarationForRecord{} -> WarningName
UselessPatternDeclarationForRecord_
ClashesViaRenaming{} -> WarningName
ClashesViaRenaming_
UserWarning{} -> WarningName
UserWarning_
InfectiveImport{} -> WarningName
InfectiveImport_
CoInfectiveImport{} -> WarningName
CoInfectiveImport_
ConfluenceCheckingIncompleteBecauseOfMeta{} -> WarningName
ConfluenceCheckingIncompleteBecauseOfMeta_
ConfluenceForCubicalNotSupported{} -> WarningName
ConfluenceForCubicalNotSupported_
IllegalRewriteRule QName
_ IllegalRewriteRuleReason
reason -> IllegalRewriteRuleReason -> WarningName
illegalRewriteWarningName IllegalRewriteRuleReason
reason
NotARewriteRule{} -> WarningName
NotARewriteRule_
RewriteNonConfluent{} -> WarningName
RewriteNonConfluent_
RewriteMaybeNonConfluent{} -> WarningName
RewriteMaybeNonConfluent_
RewriteAmbiguousRules{} -> WarningName
RewriteAmbiguousRules_
RewriteMissingRule{} -> WarningName
RewriteMissingRule_
PragmaCompileErased{} -> WarningName
PragmaCompileErased_
PragmaCompileList{} -> WarningName
PragmaCompileList_
PragmaCompileMaybe{} -> WarningName
PragmaCompileMaybe_
PragmaCompileUnparsable{} -> WarningName
PragmaCompileUnparsable_
PragmaCompileWrong{} -> WarningName
PragmaCompileWrong_
PragmaCompileWrongName{} -> WarningName
PragmaCompileWrongName_
PragmaExpectsDefinedSymbol{} -> WarningName
PragmaExpectsDefinedSymbol_
PragmaExpectsUnambiguousConstructorOrFunction{} ->
WarningName
PragmaExpectsUnambiguousConstructorOrFunction_
PragmaExpectsUnambiguousProjectionOrFunction{} ->
WarningName
PragmaExpectsUnambiguousProjectionOrFunction_
NoMain{} -> WarningName
NoMain_
PlentyInHardCompileTimeMode{}
-> WarningName
PlentyInHardCompileTimeMode_
RecordFieldWarning RecordFieldWarning
w -> case RecordFieldWarning
w of
W.DuplicateFields{} -> WarningName
DuplicateFields_
W.TooManyFields{} -> WarningName
TooManyFields_
MissingTypeSignatureForOpaque{} -> WarningName
MissingTypeSignatureForOpaque_
NotAffectedByOpaque{} -> WarningName
NotAffectedByOpaque_
UselessOpaque{} -> WarningName
UselessOpaque_
UnfoldingWrongName{} -> WarningName
UnfoldingWrongName_
UnfoldTransparentName{} -> WarningName
UnfoldTransparentName_
HiddenNotInArgumentPosition{} -> WarningName
HiddenNotInArgumentPosition_
InstanceNotInArgumentPosition{} -> WarningName
InstanceNotInArgumentPosition_
MacroInLetBindings{} -> WarningName
MacroInLetBindings_
AbstractInLetBindings{} -> WarningName
AbstractInLetBindings_
InvalidDisplayForm{} -> WarningName
InvalidDisplayForm_
UnusedVariablesInDisplayForm{} -> WarningName
UnusedVariablesInDisplayForm_
TooManyArgumentsToSort{} -> WarningName
TooManyArgumentsToSort_
WithClauseProjectionFixityMismatch{} -> WarningName
WithClauseProjectionFixityMismatch_
FaceConstraintCannotBeHidden{} -> WarningName
FaceConstraintCannotBeHidden_
FaceConstraintCannotBeNamed{} -> WarningName
FaceConstraintCannotBeNamed_
DuplicateInterfaceFiles{} -> WarningName
DuplicateInterfaceFiles_
CustomBackendWarning{} -> WarningName
CustomBackendWarning_
illegalRewriteWarningName :: IllegalRewriteRuleReason -> WarningName
illegalRewriteWarningName :: IllegalRewriteRuleReason -> WarningName
illegalRewriteWarningName = \case
LHSNotDefinitionOrConstructor{} -> WarningName
RewriteLHSNotDefinitionOrConstructor_
VariablesNotBoundByLHS{} -> WarningName
RewriteVariablesNotBoundByLHS_
VariablesBoundMoreThanOnce{} -> WarningName
RewriteVariablesBoundMoreThanOnce_
LHSReduces{} -> WarningName
RewriteLHSReduces_
HeadSymbolIsProjectionLikeFunction{} -> WarningName
RewriteHeadSymbolIsProjectionLikeFunction_
HeadSymbolIsTypeConstructor{} -> WarningName
RewriteHeadSymbolIsTypeConstructor_
HeadSymbolContainsMetas{} -> WarningName
RewriteHeadSymbolContainsMetas_
ConstructorParametersNotGeneral{} -> WarningName
RewriteConstructorParametersNotGeneral_
ContainsUnsolvedMetaVariables{} -> WarningName
RewriteContainsUnsolvedMetaVariables_
BlockedOnProblems{} -> WarningName
RewriteBlockedOnProblems_
RequiresDefinitions{} -> WarningName
RewriteRequiresDefinitions_
IllegalRewriteRuleReason
DoesNotTargetRewriteRelation -> WarningName
RewriteDoesNotTargetRewriteRelation_
IllegalRewriteRuleReason
BeforeFunctionDefinition -> WarningName
RewriteBeforeFunctionDefinition_
BeforeMutualFunctionDefinition{} -> WarningName
RewriteBeforeMutualFunctionDefinition_
IllegalRewriteRuleReason
DuplicateRewriteRule -> WarningName
DuplicateRewriteRule_
isSourceCodeWarning :: WarningName -> Bool
isSourceCodeWarning :: WarningName -> Bool
isSourceCodeWarning = \case
WarningName
DuplicateInterfaceFiles_ -> Bool
False
WarningName
WarningProblem_ -> Bool
False
WarningName
_ -> Bool
True
data TCWarning
= TCWarning
{ TCWarning -> CallStack
tcWarningLocation :: CallStack
, TCWarning -> Range
tcWarningRange :: Range
, TCWarning -> Warning
tcWarning :: Warning
, TCWarning -> Doc
tcWarningDoc :: Doc
, TCWarning -> String
tcWarningString :: String
, TCWarning -> Bool
tcWarningCached :: Bool
}
deriving (Int -> TCWarning -> ShowS
[TCWarning] -> ShowS
TCWarning -> String
(Int -> TCWarning -> ShowS)
-> (TCWarning -> String)
-> ([TCWarning] -> ShowS)
-> Show TCWarning
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> TCWarning -> ShowS
showsPrec :: Int -> TCWarning -> ShowS
$cshow :: TCWarning -> String
show :: TCWarning -> String
$cshowList :: [TCWarning] -> ShowS
showList :: [TCWarning] -> ShowS
Show, (forall x. TCWarning -> Rep TCWarning x)
-> (forall x. Rep TCWarning x -> TCWarning) -> Generic TCWarning
forall x. Rep TCWarning x -> TCWarning
forall x. TCWarning -> Rep TCWarning x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. TCWarning -> Rep TCWarning x
from :: forall x. TCWarning -> Rep TCWarning x
$cto :: forall x. Rep TCWarning x -> TCWarning
to :: forall x. Rep TCWarning x -> TCWarning
Generic)
tcWarningOrigin :: TCWarning -> SrcFile
tcWarningOrigin :: TCWarning -> Maybe RangeFile
tcWarningOrigin = Range -> Maybe RangeFile
rangeFile (Range -> Maybe RangeFile)
-> (TCWarning -> Range) -> TCWarning -> Maybe RangeFile
forall b c a. (b -> c) -> (a -> b) -> a -> c
. TCWarning -> Range
tcWarningRange
instance HasRange TCWarning where
getRange :: TCWarning -> Range
getRange = TCWarning -> Range
tcWarningRange
instance Eq TCWarning where
== :: TCWarning -> TCWarning -> Bool
(==) = (Range, String) -> (Range, String) -> Bool
forall a. Eq a => a -> a -> Bool
(==) ((Range, String) -> (Range, String) -> Bool)
-> (TCWarning -> (Range, String)) -> TCWarning -> TCWarning -> Bool
forall b c a. (b -> b -> c) -> (a -> b) -> a -> a -> c
`on` TCWarning -> Range
tcWarningRange (TCWarning -> Range)
-> (TCWarning -> String) -> TCWarning -> (Range, String)
forall b c c'. (b -> c) -> (b -> c') -> b -> (c, c')
forall (a :: * -> * -> *) b c c'.
Arrow a =>
a b c -> a b c' -> a b (c, c')
&&& TCWarning -> String
tcWarningString
instance Ord TCWarning where
compare :: TCWarning -> TCWarning -> Ordering
compare = (Range, String) -> (Range, String) -> Ordering
forall a. Ord a => a -> a -> Ordering
compare ((Range, String) -> (Range, String) -> Ordering)
-> (TCWarning -> (Range, String))
-> TCWarning
-> TCWarning
-> Ordering
forall b c a. (b -> b -> c) -> (a -> b) -> a -> a -> c
`on` TCWarning -> Range
tcWarningRange (TCWarning -> Range)
-> (TCWarning -> String) -> TCWarning -> (Range, String)
forall b c c'. (b -> c) -> (b -> c') -> b -> (c, c')
forall (a :: * -> * -> *) b c c'.
Arrow a =>
a b c -> a b c' -> a b (c, c')
&&& TCWarning -> String
tcWarningString
data CallInfo = CallInfo
{ CallInfo -> QName
callInfoTarget :: QName
, CallInfo -> Closure Term
callInfoCall :: Closure Term
} deriving (Int -> CallInfo -> ShowS
[CallInfo] -> ShowS
CallInfo -> String
(Int -> CallInfo -> ShowS)
-> (CallInfo -> String) -> ([CallInfo] -> ShowS) -> Show CallInfo
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> CallInfo -> ShowS
showsPrec :: Int -> CallInfo -> ShowS
$cshow :: CallInfo -> String
show :: CallInfo -> String
$cshowList :: [CallInfo] -> ShowS
showList :: [CallInfo] -> ShowS
Show, (forall x. CallInfo -> Rep CallInfo x)
-> (forall x. Rep CallInfo x -> CallInfo) -> Generic CallInfo
forall x. Rep CallInfo x -> CallInfo
forall x. CallInfo -> Rep CallInfo x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. CallInfo -> Rep CallInfo x
from :: forall x. CallInfo -> Rep CallInfo x
$cto :: forall x. Rep CallInfo x -> CallInfo
to :: forall x. Rep CallInfo x -> CallInfo
Generic)
instance HasRange CallInfo where
getRange :: CallInfo -> Range
getRange = QName -> Range
forall a. HasRange a => a -> Range
getRange (QName -> Range) -> (CallInfo -> QName) -> CallInfo -> Range
forall b c a. (b -> c) -> (a -> b) -> a -> c
. CallInfo -> QName
callInfoTarget
instance Pretty CallInfo where pretty :: CallInfo -> Doc
pretty = QName -> Doc
forall a. Pretty a => a -> Doc
pretty (QName -> Doc) -> (CallInfo -> QName) -> CallInfo -> Doc
forall b c a. (b -> c) -> (a -> b) -> a -> c
. CallInfo -> QName
callInfoTarget
data TerminationError = TerminationError
{ TerminationError -> [QName]
termErrFunctions :: [QName]
, TerminationError -> [CallInfo]
termErrCalls :: [CallInfo]
} deriving (Int -> TerminationError -> ShowS
[TerminationError] -> ShowS
TerminationError -> String
(Int -> TerminationError -> ShowS)
-> (TerminationError -> String)
-> ([TerminationError] -> ShowS)
-> Show TerminationError
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> TerminationError -> ShowS
showsPrec :: Int -> TerminationError -> ShowS
$cshow :: TerminationError -> String
show :: TerminationError -> String
$cshowList :: [TerminationError] -> ShowS
showList :: [TerminationError] -> ShowS
Show, (forall x. TerminationError -> Rep TerminationError x)
-> (forall x. Rep TerminationError x -> TerminationError)
-> Generic TerminationError
forall x. Rep TerminationError x -> TerminationError
forall x. TerminationError -> Rep TerminationError x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. TerminationError -> Rep TerminationError x
from :: forall x. TerminationError -> Rep TerminationError x
$cto :: forall x. Rep TerminationError x -> TerminationError
to :: forall x. Rep TerminationError x -> TerminationError
Generic)
data SplitError
= NotADatatype (Closure Type)
| BlockedType Blocker (Closure Type)
| ErasedDatatype ErasedDatatypeReason (Closure Type)
| CoinductiveDatatype (Closure Type)
| UnificationStuck
{ SplitError -> Maybe Blocker
cantSplitBlocker :: Maybe Blocker
, SplitError -> QName
cantSplitConName :: QName
, SplitError -> Telescope
cantSplitTel :: Telescope
, SplitError -> Args
cantSplitConIdx :: Args
, SplitError -> Args
cantSplitGivenIdx :: Args
, SplitError -> [UnificationFailure]
cantSplitFailures :: [UnificationFailure]
}
| CosplitCatchall
| CosplitNoTarget
| CosplitNoRecordType (Closure Type)
| CannotCreateMissingClause QName (Telescope,[NamedArg DeBruijnPattern]) Doc (Closure (Abs Type))
| GenericSplitError String
deriving (Int -> SplitError -> ShowS
[SplitError] -> ShowS
SplitError -> String
(Int -> SplitError -> ShowS)
-> (SplitError -> String)
-> ([SplitError] -> ShowS)
-> Show SplitError
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> SplitError -> ShowS
showsPrec :: Int -> SplitError -> ShowS
$cshow :: SplitError -> String
show :: SplitError -> String
$cshowList :: [SplitError] -> ShowS
showList :: [SplitError] -> ShowS
Show, (forall x. SplitError -> Rep SplitError x)
-> (forall x. Rep SplitError x -> SplitError) -> Generic SplitError
forall x. Rep SplitError x -> SplitError
forall x. SplitError -> Rep SplitError x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. SplitError -> Rep SplitError x
from :: forall x. SplitError -> Rep SplitError x
$cto :: forall x. Rep SplitError x -> SplitError
to :: forall x. Rep SplitError x -> SplitError
Generic)
data NegativeUnification
= UnifyConflict Telescope Term Term
| UnifyCycle Telescope Int Term
deriving (Int -> NegativeUnification -> ShowS
[NegativeUnification] -> ShowS
NegativeUnification -> String
(Int -> NegativeUnification -> ShowS)
-> (NegativeUnification -> String)
-> ([NegativeUnification] -> ShowS)
-> Show NegativeUnification
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> NegativeUnification -> ShowS
showsPrec :: Int -> NegativeUnification -> ShowS
$cshow :: NegativeUnification -> String
show :: NegativeUnification -> String
$cshowList :: [NegativeUnification] -> ShowS
showList :: [NegativeUnification] -> ShowS
Show, (forall x. NegativeUnification -> Rep NegativeUnification x)
-> (forall x. Rep NegativeUnification x -> NegativeUnification)
-> Generic NegativeUnification
forall x. Rep NegativeUnification x -> NegativeUnification
forall x. NegativeUnification -> Rep NegativeUnification x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. NegativeUnification -> Rep NegativeUnification x
from :: forall x. NegativeUnification -> Rep NegativeUnification x
$cto :: forall x. Rep NegativeUnification x -> NegativeUnification
to :: forall x. Rep NegativeUnification x -> NegativeUnification
Generic)
data UnificationFailure
= UnifyIndicesNotVars Telescope Type Term Term Args
| UnifyRecursiveEq Telescope Type Int Term
| UnifyReflexiveEq Telescope Type Term
| UnifyUnusableModality Telescope Type Int Term Modality
deriving (Int -> UnificationFailure -> ShowS
[UnificationFailure] -> ShowS
UnificationFailure -> String
(Int -> UnificationFailure -> ShowS)
-> (UnificationFailure -> String)
-> ([UnificationFailure] -> ShowS)
-> Show UnificationFailure
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> UnificationFailure -> ShowS
showsPrec :: Int -> UnificationFailure -> ShowS
$cshow :: UnificationFailure -> String
show :: UnificationFailure -> String
$cshowList :: [UnificationFailure] -> ShowS
showList :: [UnificationFailure] -> ShowS
Show, (forall x. UnificationFailure -> Rep UnificationFailure x)
-> (forall x. Rep UnificationFailure x -> UnificationFailure)
-> Generic UnificationFailure
forall x. Rep UnificationFailure x -> UnificationFailure
forall x. UnificationFailure -> Rep UnificationFailure x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. UnificationFailure -> Rep UnificationFailure x
from :: forall x. UnificationFailure -> Rep UnificationFailure x
$cto :: forall x. Rep UnificationFailure x -> UnificationFailure
to :: forall x. Rep UnificationFailure x -> UnificationFailure
Generic)
data UnquoteError
= BlockedOnMeta TCState Blocker
| CannotDeclareHiddenFunction QName
| CommitAfterDef
| ConInsteadOfDef QName String String
| DefineDataNotData QName
| DefInsteadOfCon QName String String
| MissingDeclaration QName
| MissingDefinition QName
| NakedUnquote
| NonCanonical String I.Term
| PatLamWithoutClauses I.Term
| StaleMeta TopLevelModuleName MetaId
| TooManyParameters Nat A.Expr
| UnboundName QName
deriving (Int -> UnquoteError -> ShowS
[UnquoteError] -> ShowS
UnquoteError -> String
(Int -> UnquoteError -> ShowS)
-> (UnquoteError -> String)
-> ([UnquoteError] -> ShowS)
-> Show UnquoteError
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> UnquoteError -> ShowS
showsPrec :: Int -> UnquoteError -> ShowS
$cshow :: UnquoteError -> String
show :: UnquoteError -> String
$cshowList :: [UnquoteError] -> ShowS
showList :: [UnquoteError] -> ShowS
Show, (forall x. UnquoteError -> Rep UnquoteError x)
-> (forall x. Rep UnquoteError x -> UnquoteError)
-> Generic UnquoteError
forall x. Rep UnquoteError x -> UnquoteError
forall x. UnquoteError -> Rep UnquoteError x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. UnquoteError -> Rep UnquoteError x
from :: forall x. UnquoteError -> Rep UnquoteError x
$cto :: forall x. Rep UnquoteError x -> UnquoteError
to :: forall x. Rep UnquoteError x -> UnquoteError
Generic)
data ExecError
= ExeNotTrusted ExeName ExeMap
| ExeNotFound ExeName FilePath
| ExeNotExecutable ExeName FilePath
deriving (Int -> ExecError -> ShowS
[ExecError] -> ShowS
ExecError -> String
(Int -> ExecError -> ShowS)
-> (ExecError -> String)
-> ([ExecError] -> ShowS)
-> Show ExecError
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> ExecError -> ShowS
showsPrec :: Int -> ExecError -> ShowS
$cshow :: ExecError -> String
show :: ExecError -> String
$cshowList :: [ExecError] -> ShowS
showList :: [ExecError] -> ShowS
Show, (forall x. ExecError -> Rep ExecError x)
-> (forall x. Rep ExecError x -> ExecError) -> Generic ExecError
forall x. Rep ExecError x -> ExecError
forall x. ExecError -> Rep ExecError x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. ExecError -> Rep ExecError x
from :: forall x. ExecError -> Rep ExecError x
$cto :: forall x. Rep ExecError x -> ExecError
to :: forall x. Rep ExecError x -> ExecError
Generic)
data TypeError
= InternalError String
| NotImplemented String
| NotSupported String
| CompilationError String
| SyntaxError String
| OptionError OptionError
| NicifierError DeclarationException'
| DoNotationError String
| IdiomBracketError String
| NoKnownRecordWithSuchFields [C.Name]
| ShouldEndInApplicationOfTheDatatype Type
| ConstructorPatternInWrongDatatype QName QName
| CantResolveOverloadedConstructorsTargetingSameDatatype QName (List1 QName)
| ConstructorDoesNotTargetGivenType QName Type
| InvalidDottedExpression
| LiteralTooBig
| NegativeLiteralInPattern
| WrongHidingInLHS
| WrongHidingInLambda Type
| WrongHidingInApplication Type
| WrongHidingInProjection QName
| IllegalHidingInPostfixProjection (NamedArg C.Expr)
| WrongNamedArgument (NamedArg A.Expr) (List1 NamedName)
| WrongAnnotationInLambda
| WrongIrrelevanceInLambda
| WrongQuantityInLambda
| WrongCohesionInLambda
| QuantityMismatch Quantity Quantity
| HidingMismatch Hiding Hiding
| RelevanceMismatch Relevance Relevance
| ForcedConstructorNotInstantiated A.Pattern
| IllformedProjectionPatternAbstract A.Pattern
| IllformedProjectionPatternConcrete C.Pattern
| CannotEliminateWithPattern (Maybe Blocker) (NamedArg A.Pattern) Type
| CannotEliminateWithProjection (Arg Type) Bool QName
| WrongNumberOfConstructorArguments QName Nat Nat
| ShouldBeEmpty Type [DeBruijnPattern]
| ShouldBeASort Type
| ShouldBePi Type
| ShouldBePath Type
| ShouldBeRecordType Type
| ShouldBeRecordPattern DeBruijnPattern
| InvalidTypeSort Sort
| SplitOnCoinductive
| SplitOnIrrelevant (Dom Type)
| SplitOnUnusableCohesion (Dom Type)
| SplitOnNonVariable Term Type
| SplitOnNonEtaRecord QName
| SplitOnAbstract QName
| SplitOnUnchecked QName
| SplitOnPartial (Dom Type)
| SplitInProp DataOrRecordE
| DefinitionIsIrrelevant QName
| DefinitionIsErased QName
| ProjectionIsIrrelevant QName
| VariableIsIrrelevant Name
| VariableIsErased Name
| VariableIsOfUnusableCohesion Name Cohesion
| LambdaIsErased
| RecordIsErased
| InvalidModalTelescopeUse Term Modality Modality Definition
| UnequalLevel Comparison Level Level
| UnequalTerms Comparison Term Term CompareAs
| UnequalRelevance Comparison Term Term
| UnequalQuantity Comparison Term Term
| UnequalCohesion Comparison Term Term
| UnequalFiniteness Comparison Term Term
| UnequalHiding Term Term
| UnequalSorts Sort Sort
| NotLeqSort Sort Sort
| MetaCannotDependOn MetaId Term Nat
| MetaIrrelevantSolution MetaId Term
| MetaErasedSolution MetaId Term
| GenericError String
| GenericDocError Doc
| SortOfSplitVarError (Maybe Blocker) Doc
| WrongSharpArity A.QName
| BuiltinMustBeConstructor BuiltinId A.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 C.Name (List1 A.ModuleName)
| BuiltinInParameterisedModule BuiltinId
| IllegalDeclarationInDataDefinition (List1 C.Declaration)
| IllegalLetInTelescope C.TypedBinding
| IllegalPatternInTelescope C.Binder
| AbsentRHSRequiresAbsurdPattern
| TooManyFields QName [C.Name] (List1 C.Name)
| DuplicateFields (List1 C.Name)
| DuplicateConstructors (List1 C.Name)
| DuplicateOverlapPragma QName OverlapMode OverlapMode
| WithOnFreeVariable A.Expr Term
| UnexpectedWithPatterns (List1 A.Pattern)
| WithClausePatternMismatch A.Pattern (NamedArg DeBruijnPattern)
| IllTypedPatternAfterWithAbstraction A.Pattern
| TooFewPatternsInWithClause
| TooManyPatternsInWithClause
| FieldOutsideRecord
| ModuleArityMismatch A.ModuleName Telescope (Either (List1 (NamedArg A.Expr)) Args)
| GeneralizeCyclicDependency
| ReferencesFutureVariables Term (List1 Int) (Arg Term) Int
| DoesNotMentionTicks Term Type (Arg Term)
| MismatchedProjectionsError QName QName
| AttributeKindNotEnabled String String String
| InvalidProjectionParameter (NamedArg A.Expr)
| TacticAttributeNotAllowed
| CannotRewriteByNonEquation Type
| MacroResultTypeMismatch Type
| NamedWhereModuleInRefinedContext [Term] [String]
| CubicalPrimitiveNotFullyApplied QName
| ComatchingDisabledForRecord QName
| IncorrectTypeForRewriteRelation Term IncorrectTypeForRewriteRelationReason
| CannotGenerateHCompClause Type
| CannotGenerateTransportClause QName (Closure (Abs Type))
| ExpectedIntervalLiteral A.Expr
| FaceConstraintDisjunction
| FaceConstraintUnsatisfiable
| PatternInPathLambda
| PatternInSystem
| UnexpectedParameter A.LamBinding
| NoParameterOfName ArgName
| UnexpectedModalityAnnotationInParameter A.LamBinding
| ExpectedBindingForParameter (Dom Type) (Abs Type)
| UnexpectedTypeSignatureForParameter (List1 (NamedArg A.Binder))
| SortDoesNotAdmitDataDefinitions QName Sort
| SortCannotDependOnItsIndex QName Type
| UnusableAtModality WhyCheckModality Modality Term
| InvalidFieldModality Cohesion
| SplitError SplitError
| ImpossibleConstructor QName NegativeUnification
| TooManyPolarities QName Int
| RecursiveRecordNeedsInductivity QName
| CannotSolveSizeConstraints (List1 (ProblemConstraint, HypSizeConstraint)) Doc
| ContradictorySizeConstraint (ProblemConstraint, HypSizeConstraint)
| EmptyTypeOfSizes Term
| FunctionTypeInSizeUniv Term
| PostulatedSizeInModule
| LibraryError LibErrors
| LibTooFarDown TopLevelModuleName AgdaLibFile
| SolvedButOpenHoles
| CyclicModuleDependency (List2 TopLevelModuleName)
| FileNotFound TopLevelModuleName [AbsolutePath]
| OverlappingProjects AbsolutePath TopLevelModuleName TopLevelModuleName
| AmbiguousTopLevelModuleName TopLevelModuleName (List2 AbsolutePath)
| ModuleNameUnexpected TopLevelModuleName TopLevelModuleName
| ModuleNameDoesntMatchFileName TopLevelModuleName [AbsolutePath]
| ModuleDefinedInOtherFile TopLevelModuleName AbsolutePath AbsolutePath
| InvalidFileName AbsolutePath InvalidFileNameReason
| ModuleNameHashCollision RawTopLevelModuleName (Maybe RawTopLevelModuleName)
| AbstractConstructorNotInScope A.QName
| CopatternHeadNotProjection C.QName
| NotAllowedInDotPatterns NotAllowedInDotPatterns
| NotInScope C.QName
| NoSuchModule C.QName
| AmbiguousName C.QName AmbiguousNameReason
| AmbiguousModule C.QName (List1 A.ModuleName)
| AmbiguousField C.Name (List2 A.ModuleName)
| AmbiguousConstructor QName (List2 QName)
| ClashingDefinition C.QName A.QName (Maybe NiceDeclaration)
| ClashingModule A.ModuleName A.ModuleName
| DefinitionInDifferentModule A.QName
| DuplicateImports C.QName (List1 C.ImportedName)
| InvalidPattern C.Pattern
| InvalidPun ConstructorOrPatternSynonym C.QName
| RepeatedNamesInImportDirective (List1 (List2 C.ImportedName))
| RepeatedVariablesInPattern (List1 C.Name)
| GeneralizeNotSupportedHere A.QName
| GeneralizedVarInLetOpenedModule A.QName
| MultipleFixityDecls (List1 (C.Name, Pair Fixity'))
| MultiplePolarityPragmas (List1 C.Name)
| ConstructorNameOfNonRecord ResolvedName
| CannotQuote CannotQuote
| CannotQuoteTerm CannotQuoteTerm
| DeclarationsAfterTopLevelModule
| IllegalDeclarationBeforeTopLevelModule
| MissingTypeSignature MissingTypeSignatureInfo
| NotAnExpression C.Expr
| NotAValidLetBinding (Maybe NotAValidLetBinding)
| NotAValidLetExpression NotAValidLetExpression
| NotValidBeforeField NiceDeclaration
| OpenEverythingInRecordWhere
| PrivateRecordField
| QualifiedLocalModule
| AsPatternInPatternSynonym
| DotPatternInPatternSynonym
| BadArgumentsToPatternSynonym A.AmbiguousQName
| TooFewArgumentsToPatternSynonym A.AmbiguousQName
| CannotResolveAmbiguousPatternSynonym (List1 (A.QName, A.PatternSynDefn))
| IllegalInstanceVariableInPatternSynonym C.Name
| PatternSynonymArgumentShadows ConstructorOrPatternSynonym C.Name (List1 AbstractName)
| UnusedVariableInPatternSynonym C.Name
| UnboundVariablesInPatternSynonym (List1 A.Name)
| NoParseForApplication (List2 C.Expr)
| AmbiguousParseForApplication (List2 C.Expr) (List1 C.Expr)
| NoParseForLHS LHSOrPatSyn [C.Pattern] C.Pattern
| AmbiguousParseForLHS LHSOrPatSyn C.Pattern (List2 C.Pattern)
| AmbiguousProjection QName (List1 QName)
| AmbiguousOverloadedProjection (List1 QName) Doc
| OperatorInformation [NotationSection] TypeError
| InstanceNoCandidate Type [(Term, TCErr)]
| ExecError ExecError
| UnquoteFailed UnquoteError
| DeBruijnIndexOutOfScope Nat Telescope [Name]
| NeedOptionAllowExec
| NeedOptionCopatterns
| NeedOptionCubical Cubical String
| NeedOptionPatternMatching
| NeedOptionProp
| NeedOptionRewriting
| NeedOptionSizedTypes String
| NeedOptionTwoLevel
| NeedOptionUniversePolymorphism
| NonFatalErrors (Set1 TCWarning)
| InstanceSearchDepthExhausted Term Type Int
| TriedToCopyConstrainedPrim QName
| InvalidInstanceHeadType Type WhyInvalidInstanceType
| InteractionError InteractionError
| BackendDoesNotSupportOnlyScopeChecking BackendName
| CubicalCompilationNotSupported Cubical
| CustomBackendError BackendName Doc
| GHCBackendError GHCBackendError
| JSBackendError JSBackendError
| UnknownBackend BackendName (Set BackendName)
deriving (Int -> TypeError -> ShowS
[TypeError] -> ShowS
TypeError -> String
(Int -> TypeError -> ShowS)
-> (TypeError -> String)
-> ([TypeError] -> ShowS)
-> Show TypeError
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> TypeError -> ShowS
showsPrec :: Int -> TypeError -> ShowS
$cshow :: TypeError -> String
show :: TypeError -> String
$cshowList :: [TypeError] -> ShowS
showList :: [TypeError] -> ShowS
Show, (forall x. TypeError -> Rep TypeError x)
-> (forall x. Rep TypeError x -> TypeError) -> Generic TypeError
forall x. Rep TypeError x -> TypeError
forall x. TypeError -> Rep TypeError x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. TypeError -> Rep TypeError x
from :: forall x. TypeError -> Rep TypeError x
$cto :: forall x. Rep TypeError x -> TypeError
to :: forall x. Rep TypeError x -> TypeError
Generic)
data WhyInvalidInstanceType
= ImproperInstHead
| ImproperInstTele
deriving (Int -> WhyInvalidInstanceType -> ShowS
[WhyInvalidInstanceType] -> ShowS
WhyInvalidInstanceType -> String
(Int -> WhyInvalidInstanceType -> ShowS)
-> (WhyInvalidInstanceType -> String)
-> ([WhyInvalidInstanceType] -> ShowS)
-> Show WhyInvalidInstanceType
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> WhyInvalidInstanceType -> ShowS
showsPrec :: Int -> WhyInvalidInstanceType -> ShowS
$cshow :: WhyInvalidInstanceType -> String
show :: WhyInvalidInstanceType -> String
$cshowList :: [WhyInvalidInstanceType] -> ShowS
showList :: [WhyInvalidInstanceType] -> ShowS
Show, (forall x. WhyInvalidInstanceType -> Rep WhyInvalidInstanceType x)
-> (forall x.
Rep WhyInvalidInstanceType x -> WhyInvalidInstanceType)
-> Generic WhyInvalidInstanceType
forall x. Rep WhyInvalidInstanceType x -> WhyInvalidInstanceType
forall x. WhyInvalidInstanceType -> Rep WhyInvalidInstanceType x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. WhyInvalidInstanceType -> Rep WhyInvalidInstanceType x
from :: forall x. WhyInvalidInstanceType -> Rep WhyInvalidInstanceType x
$cto :: forall x. Rep WhyInvalidInstanceType x -> WhyInvalidInstanceType
to :: forall x. Rep WhyInvalidInstanceType x -> WhyInvalidInstanceType
Generic)
data InteractionError
= CannotRefine String
| CaseSplitError Doc
| ExpectedIdentifier C.Expr
| ExpectedApplication
| NoActionForInteractionPoint InteractionId
| NoSuchInteractionPoint InteractionId
| UnexpectedWhere
deriving (Int -> InteractionError -> ShowS
[InteractionError] -> ShowS
InteractionError -> String
(Int -> InteractionError -> ShowS)
-> (InteractionError -> String)
-> ([InteractionError] -> ShowS)
-> Show InteractionError
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> InteractionError -> ShowS
showsPrec :: Int -> InteractionError -> ShowS
$cshow :: InteractionError -> String
show :: InteractionError -> String
$cshowList :: [InteractionError] -> ShowS
showList :: [InteractionError] -> ShowS
Show, (forall x. InteractionError -> Rep InteractionError x)
-> (forall x. Rep InteractionError x -> InteractionError)
-> Generic InteractionError
forall x. Rep InteractionError x -> InteractionError
forall x. InteractionError -> Rep InteractionError x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. InteractionError -> Rep InteractionError x
from :: forall x. InteractionError -> Rep InteractionError x
$cto :: forall x. Rep InteractionError x -> InteractionError
to :: forall x. Rep InteractionError x -> InteractionError
Generic)
data GHCBackendError
= ConstructorCountMismatch QName [QName] [String]
| NotAHaskellType Term WhyNotAHaskellType
| WrongTypeOfMain QName Type
deriving (Int -> GHCBackendError -> ShowS
[GHCBackendError] -> ShowS
GHCBackendError -> String
(Int -> GHCBackendError -> ShowS)
-> (GHCBackendError -> String)
-> ([GHCBackendError] -> ShowS)
-> Show GHCBackendError
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> GHCBackendError -> ShowS
showsPrec :: Int -> GHCBackendError -> ShowS
$cshow :: GHCBackendError -> String
show :: GHCBackendError -> String
$cshowList :: [GHCBackendError] -> ShowS
showList :: [GHCBackendError] -> ShowS
Show, (forall x. GHCBackendError -> Rep GHCBackendError x)
-> (forall x. Rep GHCBackendError x -> GHCBackendError)
-> Generic GHCBackendError
forall x. Rep GHCBackendError x -> GHCBackendError
forall x. GHCBackendError -> Rep GHCBackendError x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. GHCBackendError -> Rep GHCBackendError x
from :: forall x. GHCBackendError -> Rep GHCBackendError x
$cto :: forall x. Rep GHCBackendError x -> GHCBackendError
to :: forall x. Rep GHCBackendError x -> GHCBackendError
Generic)
data JSBackendError
= BadCompilePragma
deriving (Int -> JSBackendError -> ShowS
[JSBackendError] -> ShowS
JSBackendError -> String
(Int -> JSBackendError -> ShowS)
-> (JSBackendError -> String)
-> ([JSBackendError] -> ShowS)
-> Show JSBackendError
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> JSBackendError -> ShowS
showsPrec :: Int -> JSBackendError -> ShowS
$cshow :: JSBackendError -> String
show :: JSBackendError -> String
$cshowList :: [JSBackendError] -> ShowS
showList :: [JSBackendError] -> ShowS
Show, (forall x. JSBackendError -> Rep JSBackendError x)
-> (forall x. Rep JSBackendError x -> JSBackendError)
-> Generic JSBackendError
forall x. Rep JSBackendError x -> JSBackendError
forall x. JSBackendError -> Rep JSBackendError x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. JSBackendError -> Rep JSBackendError x
from :: forall x. JSBackendError -> Rep JSBackendError x
$cto :: forall x. Rep JSBackendError x -> JSBackendError
to :: forall x. Rep JSBackendError x -> JSBackendError
Generic)
data MissingTypeSignatureInfo
= MissingDataSignature C.Name
| MissingRecordSignature C.Name
| MissingFunctionSignature C.LHS
deriving (Int -> MissingTypeSignatureInfo -> ShowS
[MissingTypeSignatureInfo] -> ShowS
MissingTypeSignatureInfo -> String
(Int -> MissingTypeSignatureInfo -> ShowS)
-> (MissingTypeSignatureInfo -> String)
-> ([MissingTypeSignatureInfo] -> ShowS)
-> Show MissingTypeSignatureInfo
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> MissingTypeSignatureInfo -> ShowS
showsPrec :: Int -> MissingTypeSignatureInfo -> ShowS
$cshow :: MissingTypeSignatureInfo -> String
show :: MissingTypeSignatureInfo -> String
$cshowList :: [MissingTypeSignatureInfo] -> ShowS
showList :: [MissingTypeSignatureInfo] -> ShowS
Show, (forall x.
MissingTypeSignatureInfo -> Rep MissingTypeSignatureInfo x)
-> (forall x.
Rep MissingTypeSignatureInfo x -> MissingTypeSignatureInfo)
-> Generic MissingTypeSignatureInfo
forall x.
Rep MissingTypeSignatureInfo x -> MissingTypeSignatureInfo
forall x.
MissingTypeSignatureInfo -> Rep MissingTypeSignatureInfo x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x.
MissingTypeSignatureInfo -> Rep MissingTypeSignatureInfo x
from :: forall x.
MissingTypeSignatureInfo -> Rep MissingTypeSignatureInfo x
$cto :: forall x.
Rep MissingTypeSignatureInfo x -> MissingTypeSignatureInfo
to :: forall x.
Rep MissingTypeSignatureInfo x -> MissingTypeSignatureInfo
Generic)
data WhyNotAHaskellType
= NoPragmaFor QName
| WrongPragmaFor Range QName
| BadLambda Term
| BadMeta Term
| BadDontCare Term
| NotCompiled QName
deriving (Int -> WhyNotAHaskellType -> ShowS
[WhyNotAHaskellType] -> ShowS
WhyNotAHaskellType -> String
(Int -> WhyNotAHaskellType -> ShowS)
-> (WhyNotAHaskellType -> String)
-> ([WhyNotAHaskellType] -> ShowS)
-> Show WhyNotAHaskellType
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> WhyNotAHaskellType -> ShowS
showsPrec :: Int -> WhyNotAHaskellType -> ShowS
$cshow :: WhyNotAHaskellType -> String
show :: WhyNotAHaskellType -> String
$cshowList :: [WhyNotAHaskellType] -> ShowS
showList :: [WhyNotAHaskellType] -> ShowS
Show, (forall x. WhyNotAHaskellType -> Rep WhyNotAHaskellType x)
-> (forall x. Rep WhyNotAHaskellType x -> WhyNotAHaskellType)
-> Generic WhyNotAHaskellType
forall x. Rep WhyNotAHaskellType x -> WhyNotAHaskellType
forall x. WhyNotAHaskellType -> Rep WhyNotAHaskellType x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. WhyNotAHaskellType -> Rep WhyNotAHaskellType x
from :: forall x. WhyNotAHaskellType -> Rep WhyNotAHaskellType x
$cto :: forall x. Rep WhyNotAHaskellType x -> WhyNotAHaskellType
to :: forall x. Rep WhyNotAHaskellType x -> WhyNotAHaskellType
Generic)
data InvalidFileNameReason
= DoesNotCorrespondToValidModuleName
| RootNameModuleNotAQualifiedModuleName Text
deriving (Int -> InvalidFileNameReason -> ShowS
[InvalidFileNameReason] -> ShowS
InvalidFileNameReason -> String
(Int -> InvalidFileNameReason -> ShowS)
-> (InvalidFileNameReason -> String)
-> ([InvalidFileNameReason] -> ShowS)
-> Show InvalidFileNameReason
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> InvalidFileNameReason -> ShowS
showsPrec :: Int -> InvalidFileNameReason -> ShowS
$cshow :: InvalidFileNameReason -> String
show :: InvalidFileNameReason -> String
$cshowList :: [InvalidFileNameReason] -> ShowS
showList :: [InvalidFileNameReason] -> ShowS
Show, (forall x. InvalidFileNameReason -> Rep InvalidFileNameReason x)
-> (forall x. Rep InvalidFileNameReason x -> InvalidFileNameReason)
-> Generic InvalidFileNameReason
forall x. Rep InvalidFileNameReason x -> InvalidFileNameReason
forall x. InvalidFileNameReason -> Rep InvalidFileNameReason x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. InvalidFileNameReason -> Rep InvalidFileNameReason x
from :: forall x. InvalidFileNameReason -> Rep InvalidFileNameReason x
$cto :: forall x. Rep InvalidFileNameReason x -> InvalidFileNameReason
to :: forall x. Rep InvalidFileNameReason x -> InvalidFileNameReason
Generic)
type DataOrRecordE = DataOrRecord' InductionAndEta
data InductionAndEta = InductionAndEta
{ InductionAndEta -> Maybe Induction
recordInduction :: Maybe Induction
, InductionAndEta -> EtaEquality
recordEtaEquality :: EtaEquality
} deriving (Int -> InductionAndEta -> ShowS
[InductionAndEta] -> ShowS
InductionAndEta -> String
(Int -> InductionAndEta -> ShowS)
-> (InductionAndEta -> String)
-> ([InductionAndEta] -> ShowS)
-> Show InductionAndEta
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> InductionAndEta -> ShowS
showsPrec :: Int -> InductionAndEta -> ShowS
$cshow :: InductionAndEta -> String
show :: InductionAndEta -> String
$cshowList :: [InductionAndEta] -> ShowS
showList :: [InductionAndEta] -> ShowS
Show, (forall x. InductionAndEta -> Rep InductionAndEta x)
-> (forall x. Rep InductionAndEta x -> InductionAndEta)
-> Generic InductionAndEta
forall x. Rep InductionAndEta x -> InductionAndEta
forall x. InductionAndEta -> Rep InductionAndEta x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. InductionAndEta -> Rep InductionAndEta x
from :: forall x. InductionAndEta -> Rep InductionAndEta x
$cto :: forall x. Rep InductionAndEta x -> InductionAndEta
to :: forall x. Rep InductionAndEta x -> InductionAndEta
Generic)
data IllegalRewriteRuleReason
= LHSNotDefinitionOrConstructor
| VariablesNotBoundByLHS IntSet
| VariablesBoundMoreThanOnce IntSet
| LHSReduces Term Term
| HeadSymbolIsProjectionLikeFunction QName
| HeadSymbolIsTypeConstructor QName
| HeadSymbolContainsMetas QName
| ConstructorParametersNotGeneral ConHead Args
| ContainsUnsolvedMetaVariables (Set1 MetaId)
| BlockedOnProblems (Set1 ProblemId)
| RequiresDefinitions (Set1 QName)
| DoesNotTargetRewriteRelation
| BeforeFunctionDefinition
| BeforeMutualFunctionDefinition QName
| DuplicateRewriteRule
deriving (Int -> IllegalRewriteRuleReason -> ShowS
[IllegalRewriteRuleReason] -> ShowS
IllegalRewriteRuleReason -> String
(Int -> IllegalRewriteRuleReason -> ShowS)
-> (IllegalRewriteRuleReason -> String)
-> ([IllegalRewriteRuleReason] -> ShowS)
-> Show IllegalRewriteRuleReason
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> IllegalRewriteRuleReason -> ShowS
showsPrec :: Int -> IllegalRewriteRuleReason -> ShowS
$cshow :: IllegalRewriteRuleReason -> String
show :: IllegalRewriteRuleReason -> String
$cshowList :: [IllegalRewriteRuleReason] -> ShowS
showList :: [IllegalRewriteRuleReason] -> ShowS
Show, (forall x.
IllegalRewriteRuleReason -> Rep IllegalRewriteRuleReason x)
-> (forall x.
Rep IllegalRewriteRuleReason x -> IllegalRewriteRuleReason)
-> Generic IllegalRewriteRuleReason
forall x.
Rep IllegalRewriteRuleReason x -> IllegalRewriteRuleReason
forall x.
IllegalRewriteRuleReason -> Rep IllegalRewriteRuleReason x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x.
IllegalRewriteRuleReason -> Rep IllegalRewriteRuleReason x
from :: forall x.
IllegalRewriteRuleReason -> Rep IllegalRewriteRuleReason x
$cto :: forall x.
Rep IllegalRewriteRuleReason x -> IllegalRewriteRuleReason
to :: forall x.
Rep IllegalRewriteRuleReason x -> IllegalRewriteRuleReason
Generic)
data IsAmbiguous
= YesAmbiguous AmbiguousQName
| NotAmbiguous
deriving (Int -> IsAmbiguous -> ShowS
[IsAmbiguous] -> ShowS
IsAmbiguous -> String
(Int -> IsAmbiguous -> ShowS)
-> (IsAmbiguous -> String)
-> ([IsAmbiguous] -> ShowS)
-> Show IsAmbiguous
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> IsAmbiguous -> ShowS
showsPrec :: Int -> IsAmbiguous -> ShowS
$cshow :: IsAmbiguous -> String
show :: IsAmbiguous -> String
$cshowList :: [IsAmbiguous] -> ShowS
showList :: [IsAmbiguous] -> ShowS
Show, (forall x. IsAmbiguous -> Rep IsAmbiguous x)
-> (forall x. Rep IsAmbiguous x -> IsAmbiguous)
-> Generic IsAmbiguous
forall x. Rep IsAmbiguous x -> IsAmbiguous
forall x. IsAmbiguous -> Rep IsAmbiguous x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. IsAmbiguous -> Rep IsAmbiguous x
from :: forall x. IsAmbiguous -> Rep IsAmbiguous x
$cto :: forall x. Rep IsAmbiguous x -> IsAmbiguous
to :: forall x. Rep IsAmbiguous x -> IsAmbiguous
Generic)
data IncorrectTypeForRewriteRelationReason
= ShouldAcceptAtLeastTwoArguments
| FinalTwoArgumentsNotVisible
| TypeDoesNotEndInSort Type Telescope
deriving (Int -> IncorrectTypeForRewriteRelationReason -> ShowS
[IncorrectTypeForRewriteRelationReason] -> ShowS
IncorrectTypeForRewriteRelationReason -> String
(Int -> IncorrectTypeForRewriteRelationReason -> ShowS)
-> (IncorrectTypeForRewriteRelationReason -> String)
-> ([IncorrectTypeForRewriteRelationReason] -> ShowS)
-> Show IncorrectTypeForRewriteRelationReason
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> IncorrectTypeForRewriteRelationReason -> ShowS
showsPrec :: Int -> IncorrectTypeForRewriteRelationReason -> ShowS
$cshow :: IncorrectTypeForRewriteRelationReason -> String
show :: IncorrectTypeForRewriteRelationReason -> String
$cshowList :: [IncorrectTypeForRewriteRelationReason] -> ShowS
showList :: [IncorrectTypeForRewriteRelationReason] -> ShowS
Show, (forall x.
IncorrectTypeForRewriteRelationReason
-> Rep IncorrectTypeForRewriteRelationReason x)
-> (forall x.
Rep IncorrectTypeForRewriteRelationReason x
-> IncorrectTypeForRewriteRelationReason)
-> Generic IncorrectTypeForRewriteRelationReason
forall x.
Rep IncorrectTypeForRewriteRelationReason x
-> IncorrectTypeForRewriteRelationReason
forall x.
IncorrectTypeForRewriteRelationReason
-> Rep IncorrectTypeForRewriteRelationReason x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x.
IncorrectTypeForRewriteRelationReason
-> Rep IncorrectTypeForRewriteRelationReason x
from :: forall x.
IncorrectTypeForRewriteRelationReason
-> Rep IncorrectTypeForRewriteRelationReason x
$cto :: forall x.
Rep IncorrectTypeForRewriteRelationReason x
-> IncorrectTypeForRewriteRelationReason
to :: forall x.
Rep IncorrectTypeForRewriteRelationReason x
-> IncorrectTypeForRewriteRelationReason
Generic)
data CannotQuote
= CannotQuoteAmbiguous (List2 A.QName)
| CannotQuoteExpression A.Expr
| CannotQuoteHidden
| CannotQuoteNothing
| CannotQuotePattern (NamedArg C.Pattern)
deriving (Int -> CannotQuote -> ShowS
[CannotQuote] -> ShowS
CannotQuote -> String
(Int -> CannotQuote -> ShowS)
-> (CannotQuote -> String)
-> ([CannotQuote] -> ShowS)
-> Show CannotQuote
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> CannotQuote -> ShowS
showsPrec :: Int -> CannotQuote -> ShowS
$cshow :: CannotQuote -> String
show :: CannotQuote -> String
$cshowList :: [CannotQuote] -> ShowS
showList :: [CannotQuote] -> ShowS
Show, (forall x. CannotQuote -> Rep CannotQuote x)
-> (forall x. Rep CannotQuote x -> CannotQuote)
-> Generic CannotQuote
forall x. Rep CannotQuote x -> CannotQuote
forall x. CannotQuote -> Rep CannotQuote x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. CannotQuote -> Rep CannotQuote x
from :: forall x. CannotQuote -> Rep CannotQuote x
$cto :: forall x. Rep CannotQuote x -> CannotQuote
to :: forall x. Rep CannotQuote x -> CannotQuote
Generic)
data LHSOrPatSyn = IsLHS | IsPatSyn
deriving (LHSOrPatSyn -> LHSOrPatSyn -> Bool
(LHSOrPatSyn -> LHSOrPatSyn -> Bool)
-> (LHSOrPatSyn -> LHSOrPatSyn -> Bool) -> Eq LHSOrPatSyn
forall a. (a -> a -> Bool) -> (a -> a -> Bool) -> Eq a
$c== :: LHSOrPatSyn -> LHSOrPatSyn -> Bool
== :: LHSOrPatSyn -> LHSOrPatSyn -> Bool
$c/= :: LHSOrPatSyn -> LHSOrPatSyn -> Bool
/= :: LHSOrPatSyn -> LHSOrPatSyn -> Bool
Eq, Int -> LHSOrPatSyn -> ShowS
[LHSOrPatSyn] -> ShowS
LHSOrPatSyn -> String
(Int -> LHSOrPatSyn -> ShowS)
-> (LHSOrPatSyn -> String)
-> ([LHSOrPatSyn] -> ShowS)
-> Show LHSOrPatSyn
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> LHSOrPatSyn -> ShowS
showsPrec :: Int -> LHSOrPatSyn -> ShowS
$cshow :: LHSOrPatSyn -> String
show :: LHSOrPatSyn -> String
$cshowList :: [LHSOrPatSyn] -> ShowS
showList :: [LHSOrPatSyn] -> ShowS
Show, (forall x. LHSOrPatSyn -> Rep LHSOrPatSyn x)
-> (forall x. Rep LHSOrPatSyn x -> LHSOrPatSyn)
-> Generic LHSOrPatSyn
forall x. Rep LHSOrPatSyn x -> LHSOrPatSyn
forall x. LHSOrPatSyn -> Rep LHSOrPatSyn x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. LHSOrPatSyn -> Rep LHSOrPatSyn x
from :: forall x. LHSOrPatSyn -> Rep LHSOrPatSyn x
$cto :: forall x. Rep LHSOrPatSyn x -> LHSOrPatSyn
to :: forall x. Rep LHSOrPatSyn x -> LHSOrPatSyn
Generic, LHSOrPatSyn
LHSOrPatSyn -> LHSOrPatSyn -> Bounded LHSOrPatSyn
forall a. a -> a -> Bounded a
$cminBound :: LHSOrPatSyn
minBound :: LHSOrPatSyn
$cmaxBound :: LHSOrPatSyn
maxBound :: LHSOrPatSyn
Bounded, Int -> LHSOrPatSyn
LHSOrPatSyn -> Int
LHSOrPatSyn -> [LHSOrPatSyn]
LHSOrPatSyn -> LHSOrPatSyn
LHSOrPatSyn -> LHSOrPatSyn -> [LHSOrPatSyn]
LHSOrPatSyn -> LHSOrPatSyn -> LHSOrPatSyn -> [LHSOrPatSyn]
(LHSOrPatSyn -> LHSOrPatSyn)
-> (LHSOrPatSyn -> LHSOrPatSyn)
-> (Int -> LHSOrPatSyn)
-> (LHSOrPatSyn -> Int)
-> (LHSOrPatSyn -> [LHSOrPatSyn])
-> (LHSOrPatSyn -> LHSOrPatSyn -> [LHSOrPatSyn])
-> (LHSOrPatSyn -> LHSOrPatSyn -> [LHSOrPatSyn])
-> (LHSOrPatSyn -> LHSOrPatSyn -> LHSOrPatSyn -> [LHSOrPatSyn])
-> Enum LHSOrPatSyn
forall a.
(a -> a)
-> (a -> a)
-> (Int -> a)
-> (a -> Int)
-> (a -> [a])
-> (a -> a -> [a])
-> (a -> a -> [a])
-> (a -> a -> a -> [a])
-> Enum a
$csucc :: LHSOrPatSyn -> LHSOrPatSyn
succ :: LHSOrPatSyn -> LHSOrPatSyn
$cpred :: LHSOrPatSyn -> LHSOrPatSyn
pred :: LHSOrPatSyn -> LHSOrPatSyn
$ctoEnum :: Int -> LHSOrPatSyn
toEnum :: Int -> LHSOrPatSyn
$cfromEnum :: LHSOrPatSyn -> Int
fromEnum :: LHSOrPatSyn -> Int
$cenumFrom :: LHSOrPatSyn -> [LHSOrPatSyn]
enumFrom :: LHSOrPatSyn -> [LHSOrPatSyn]
$cenumFromThen :: LHSOrPatSyn -> LHSOrPatSyn -> [LHSOrPatSyn]
enumFromThen :: LHSOrPatSyn -> LHSOrPatSyn -> [LHSOrPatSyn]
$cenumFromTo :: LHSOrPatSyn -> LHSOrPatSyn -> [LHSOrPatSyn]
enumFromTo :: LHSOrPatSyn -> LHSOrPatSyn -> [LHSOrPatSyn]
$cenumFromThenTo :: LHSOrPatSyn -> LHSOrPatSyn -> LHSOrPatSyn -> [LHSOrPatSyn]
enumFromThenTo :: LHSOrPatSyn -> LHSOrPatSyn -> LHSOrPatSyn -> [LHSOrPatSyn]
Enum)
data TCErr
= TypeError
{ TCErr -> CallStack
tcErrLocation :: CallStack
, TCErr -> TCState
tcErrState :: TCState
, TCErr -> Closure TypeError
tcErrClosErr :: Closure TypeError
}
| ParserError ParseError
| GenericException String
| IOException (Maybe TCState) Range E.IOException
| PatternErr Blocker
instance Show TCErr where
show :: TCErr -> String
show = \case
TypeError CallStack
_ TCState
_ Closure TypeError
e -> Range -> String
forall a. Pretty a => a -> String
prettyShow (TCEnv -> Range
envRange (TCEnv -> Range) -> TCEnv -> Range
forall a b. (a -> b) -> a -> b
$ Closure TypeError -> TCEnv
forall a. Closure a -> TCEnv
clEnv Closure TypeError
e) String -> ShowS
forall a. [a] -> [a] -> [a]
++ String
": " String -> ShowS
forall a. [a] -> [a] -> [a]
++ TypeError -> String
forall a. Show a => a -> String
show (Closure TypeError -> TypeError
forall a. Closure a -> a
clValue Closure TypeError
e)
ParserError ParseError
e -> ParseError -> String
forall a. Pretty a => a -> String
prettyShow ParseError
e
GenericException String
msg -> String
msg
IOException Maybe TCState
_ Range
r IOException
e -> Range -> String
forall a. Pretty a => a -> String
prettyShow Range
r String -> ShowS
forall a. [a] -> [a] -> [a]
++ String
": " String -> ShowS
forall a. [a] -> [a] -> [a]
++ IOException -> String
forall e. Exception e => e -> String
showIOException IOException
e
PatternErr{} -> String
"Pattern violation (you shouldn't see this)"
instance HasRange TCErr where
getRange :: TCErr -> Range
getRange (TypeError CallStack
_ TCState
_ Closure TypeError
cl) = TCEnv -> Range
envRange (TCEnv -> Range) -> TCEnv -> Range
forall a b. (a -> b) -> a -> b
$ Closure TypeError -> TCEnv
forall a. Closure a -> TCEnv
clEnv Closure TypeError
cl
getRange (ParserError ParseError
e) = ParseError -> Range
forall a. HasRange a => a -> Range
getRange ParseError
e
getRange GenericException{} = Range
forall a. Range' a
noRange
getRange (IOException Maybe TCState
_ Range
r IOException
_) = Range
r
getRange PatternErr{} = Range
forall a. Range' a
noRange
instance E.Exception TCErr
data WarningsAndNonFatalErrors = WarningsAndNonFatalErrors
{ WarningsAndNonFatalErrors -> Set TCWarning
tcWarnings :: Set TCWarning
, WarningsAndNonFatalErrors -> Set TCWarning
nonFatalErrors :: Set TCWarning
}
instance Null WarningsAndNonFatalErrors where
null :: WarningsAndNonFatalErrors -> Bool
null (WarningsAndNonFatalErrors Set TCWarning
ws Set TCWarning
errs) = Set TCWarning -> Bool
forall a. Null a => a -> Bool
null Set TCWarning
ws Bool -> Bool -> Bool
&& Set TCWarning -> Bool
forall a. Null a => a -> Bool
null Set TCWarning
errs
empty :: WarningsAndNonFatalErrors
empty = Set TCWarning -> Set TCWarning -> WarningsAndNonFatalErrors
WarningsAndNonFatalErrors Set TCWarning
forall a. Null a => a
empty Set TCWarning
forall a. Null a => a
empty
instance MonadIO m => HasOptions (TCMT m) where
pragmaOptions :: TCMT m PragmaOptions
pragmaOptions = Lens' TCState PragmaOptions -> TCMT m PragmaOptions
forall (m :: * -> *) a. ReadTCState m => Lens' TCState a -> m a
useTC (PragmaOptions -> f PragmaOptions) -> TCState -> f TCState
Lens' TCState PragmaOptions
stPragmaOptions
{-# INLINE pragmaOptions #-}
commandLineOptions :: TCMT m CommandLineOptions
commandLineOptions = do
p <- Lens' TCState PragmaOptions -> TCMT m PragmaOptions
forall (m :: * -> *) a. ReadTCState m => Lens' TCState a -> m a
useTC (PragmaOptions -> f PragmaOptions) -> TCState -> f TCState
Lens' TCState PragmaOptions
stPragmaOptions
cl <- stPersistentOptions . stPersistentState <$> getTC
return $ cl { optPragmaOptions = p }
{-# SPECIALIZE commandLineOptions :: TCM CommandLineOptions #-}
sizedTypesOption :: HasOptions m => m Bool
sizedTypesOption :: forall (m :: * -> *). HasOptions m => m Bool
sizedTypesOption = PragmaOptions -> Bool
optSizedTypes (PragmaOptions -> Bool) -> m PragmaOptions -> m Bool
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> m PragmaOptions
forall (m :: * -> *). HasOptions m => m PragmaOptions
pragmaOptions
{-# INLINE sizedTypesOption #-}
guardednessOption :: HasOptions m => m Bool
guardednessOption :: forall (m :: * -> *). HasOptions m => m Bool
guardednessOption = PragmaOptions -> Bool
optGuardedness (PragmaOptions -> Bool) -> m PragmaOptions -> m Bool
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> m PragmaOptions
forall (m :: * -> *). HasOptions m => m PragmaOptions
pragmaOptions
{-# INLINE guardednessOption #-}
withoutKOption :: HasOptions m => m Bool
withoutKOption :: forall (m :: * -> *). HasOptions m => m Bool
withoutKOption = PragmaOptions -> Bool
optWithoutK (PragmaOptions -> Bool) -> m PragmaOptions -> m Bool
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> m PragmaOptions
forall (m :: * -> *). HasOptions m => m PragmaOptions
pragmaOptions
{-# INLINE withoutKOption #-}
cubicalOption :: HasOptions m => m (Maybe Cubical)
cubicalOption :: forall (m :: * -> *). HasOptions m => m (Maybe Cubical)
cubicalOption = PragmaOptions -> Maybe Cubical
optCubical (PragmaOptions -> Maybe Cubical)
-> m PragmaOptions -> m (Maybe Cubical)
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> m PragmaOptions
forall (m :: * -> *). HasOptions m => m PragmaOptions
pragmaOptions
{-# INLINE cubicalOption #-}
cubicalCompatibleOption :: HasOptions m => m Bool
cubicalCompatibleOption :: forall (m :: * -> *). HasOptions m => m Bool
cubicalCompatibleOption = PragmaOptions -> Bool
optCubicalCompatible (PragmaOptions -> Bool) -> m PragmaOptions -> m Bool
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> m PragmaOptions
forall (m :: * -> *). HasOptions m => m PragmaOptions
pragmaOptions
{-# INLINE cubicalCompatibleOption #-}
enableCaching :: HasOptions m => m Bool
enableCaching :: forall (m :: * -> *). HasOptions m => m Bool
enableCaching = PragmaOptions -> Bool
optCaching (PragmaOptions -> Bool) -> m PragmaOptions -> m Bool
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> m PragmaOptions
forall (m :: * -> *). HasOptions m => m PragmaOptions
pragmaOptions
{-# INLINE enableCaching #-}
data ReduceEnv = ReduceEnv
{ ReduceEnv -> TCEnv
redEnv :: TCEnv
, ReduceEnv -> TCState
redSt :: TCState
, ReduceEnv -> Maybe (MetaId -> ReduceM Bool)
redPred :: Maybe (MetaId -> ReduceM Bool)
}
mapRedEnv :: (TCEnv -> TCEnv) -> ReduceEnv -> ReduceEnv
mapRedEnv :: (TCEnv -> TCEnv) -> ReduceEnv -> ReduceEnv
mapRedEnv TCEnv -> TCEnv
f ReduceEnv
s = ReduceEnv
s { redEnv = f (redEnv s) }
{-# INLINE mapRedEnv #-}
mapRedSt :: (TCState -> TCState) -> ReduceEnv -> ReduceEnv
mapRedSt :: (TCState -> TCState) -> ReduceEnv -> ReduceEnv
mapRedSt TCState -> TCState
f ReduceEnv
s = ReduceEnv
s { redSt = f (redSt s) }
{-# INLINE mapRedSt #-}
mapRedEnvSt :: (TCEnv -> TCEnv) -> (TCState -> TCState) -> ReduceEnv
-> ReduceEnv
mapRedEnvSt :: (TCEnv -> TCEnv) -> (TCState -> TCState) -> ReduceEnv -> ReduceEnv
mapRedEnvSt TCEnv -> TCEnv
f TCState -> TCState
g (ReduceEnv TCEnv
e TCState
s Maybe (MetaId -> ReduceM Bool)
p) = TCEnv -> TCState -> Maybe (MetaId -> ReduceM Bool) -> ReduceEnv
ReduceEnv (TCEnv -> TCEnv
f TCEnv
e) (TCState -> TCState
g TCState
s) Maybe (MetaId -> ReduceM Bool)
p
{-# INLINE mapRedEnvSt #-}
reduceEnv :: Lens' ReduceEnv TCEnv
reduceEnv :: Lens' ReduceEnv TCEnv
reduceEnv TCEnv -> f TCEnv
f ReduceEnv
s = TCEnv -> f TCEnv
f (ReduceEnv -> TCEnv
redEnv ReduceEnv
s) f TCEnv -> (TCEnv -> ReduceEnv) -> f ReduceEnv
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ TCEnv
e -> ReduceEnv
s { redEnv = e }
{-# INLINE reduceEnv #-}
reduceSt :: Lens' ReduceEnv TCState
reduceSt :: Lens' ReduceEnv TCState
reduceSt TCState -> f TCState
f ReduceEnv
s = TCState -> f TCState
f (ReduceEnv -> TCState
redSt ReduceEnv
s) f TCState -> (TCState -> ReduceEnv) -> f ReduceEnv
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ TCState
e -> ReduceEnv
s { redSt = e }
{-# INLINE reduceSt #-}
newtype ReduceM a = ReduceM { forall a. ReduceM a -> ReduceEnv -> a
unReduceM :: ReduceEnv -> a }
unKleisli :: (a -> ReduceM b) -> ReduceM (a -> b)
unKleisli :: forall a b. (a -> ReduceM b) -> ReduceM (a -> b)
unKleisli a -> ReduceM b
f = (ReduceEnv -> a -> b) -> ReduceM (a -> b)
forall a. (ReduceEnv -> a) -> ReduceM a
ReduceM ((ReduceEnv -> a -> b) -> ReduceM (a -> b))
-> (ReduceEnv -> a -> b) -> ReduceM (a -> b)
forall a b. (a -> b) -> a -> b
$ \ ReduceEnv
env a
x -> ReduceM b -> ReduceEnv -> b
forall a. ReduceM a -> ReduceEnv -> a
unReduceM (a -> ReduceM b
f a
x) ReduceEnv
env
onReduceEnv :: (ReduceEnv -> ReduceEnv) -> ReduceM a -> ReduceM a
onReduceEnv :: forall a. (ReduceEnv -> ReduceEnv) -> ReduceM a -> ReduceM a
onReduceEnv ReduceEnv -> ReduceEnv
f (ReduceM ReduceEnv -> a
m) = (ReduceEnv -> a) -> ReduceM a
forall a. (ReduceEnv -> a) -> ReduceM a
ReduceM (ReduceEnv -> a
m (ReduceEnv -> a) -> (ReduceEnv -> ReduceEnv) -> ReduceEnv -> a
forall b c a. (b -> c) -> (a -> b) -> a -> c
. ReduceEnv -> ReduceEnv
f)
{-# INLINE onReduceEnv #-}
fmapReduce :: (a -> b) -> ReduceM a -> ReduceM b
fmapReduce :: forall a b. (a -> b) -> ReduceM a -> ReduceM b
fmapReduce a -> b
f (ReduceM ReduceEnv -> a
m) = (ReduceEnv -> b) -> ReduceM b
forall a. (ReduceEnv -> a) -> ReduceM a
ReduceM ((ReduceEnv -> b) -> ReduceM b) -> (ReduceEnv -> b) -> ReduceM b
forall a b. (a -> b) -> a -> b
$ \ ReduceEnv
e -> a -> b
f (a -> b) -> a -> b
forall a b. (a -> b) -> a -> b
$! ReduceEnv -> a
m ReduceEnv
e
{-# INLINE fmapReduce #-}
apReduce :: ReduceM (a -> b) -> ReduceM a -> ReduceM b
apReduce :: forall a b. ReduceM (a -> b) -> ReduceM a -> ReduceM b
apReduce (ReduceM ReduceEnv -> a -> b
f) (ReduceM ReduceEnv -> a
x) = (ReduceEnv -> b) -> ReduceM b
forall a. (ReduceEnv -> a) -> ReduceM a
ReduceM ((ReduceEnv -> b) -> ReduceM b) -> (ReduceEnv -> b) -> ReduceM b
forall a b. (a -> b) -> a -> b
$ \ ReduceEnv
e ->
let g :: a -> b
g = ReduceEnv -> a -> b
f ReduceEnv
e
a :: a
a = ReduceEnv -> a
x ReduceEnv
e
in a -> b
g (a -> b) -> b -> b
forall a b. a -> b -> b
`pseq` a
a a -> b -> b
forall a b. a -> b -> b
`pseq` a -> b
g a
a
{-# INLINE apReduce #-}
thenReduce :: ReduceM a -> ReduceM b -> ReduceM b
thenReduce :: forall a b. ReduceM a -> ReduceM b -> ReduceM b
thenReduce (ReduceM ReduceEnv -> a
x) (ReduceM ReduceEnv -> b
y) = (ReduceEnv -> b) -> ReduceM b
forall a. (ReduceEnv -> a) -> ReduceM a
ReduceM ((ReduceEnv -> b) -> ReduceM b) -> (ReduceEnv -> b) -> ReduceM b
forall a b. (a -> b) -> a -> b
$ \ ReduceEnv
e -> ReduceEnv -> a
x ReduceEnv
e a -> b -> b
forall a b. a -> b -> b
`pseq` ReduceEnv -> b
y ReduceEnv
e
{-# INLINE thenReduce #-}
beforeReduce :: ReduceM a -> ReduceM b -> ReduceM a
beforeReduce :: forall a b. ReduceM a -> ReduceM b -> ReduceM a
beforeReduce (ReduceM ReduceEnv -> a
x) (ReduceM ReduceEnv -> b
y) = (ReduceEnv -> a) -> ReduceM a
forall a. (ReduceEnv -> a) -> ReduceM a
ReduceM ((ReduceEnv -> a) -> ReduceM a) -> (ReduceEnv -> a) -> ReduceM a
forall a b. (a -> b) -> a -> b
$ \ ReduceEnv
e ->
let a :: a
a = ReduceEnv -> a
x ReduceEnv
e
in a
a a -> a -> a
forall a b. a -> b -> b
`pseq` ReduceEnv -> b
y ReduceEnv
e b -> a -> a
forall a b. a -> b -> b
`pseq` a
a
{-# INLINE beforeReduce #-}
bindReduce :: ReduceM a -> (a -> ReduceM b) -> ReduceM b
bindReduce :: forall a b. ReduceM a -> (a -> ReduceM b) -> ReduceM b
bindReduce (ReduceM ReduceEnv -> a
m) a -> ReduceM b
f = (ReduceEnv -> b) -> ReduceM b
forall a. (ReduceEnv -> a) -> ReduceM a
ReduceM ((ReduceEnv -> b) -> ReduceM b) -> (ReduceEnv -> b) -> ReduceM b
forall a b. (a -> b) -> a -> b
$ \ ReduceEnv
e -> ReduceM b -> ReduceEnv -> b
forall a. ReduceM a -> ReduceEnv -> a
unReduceM (a -> ReduceM b
f (a -> ReduceM b) -> a -> ReduceM b
forall a b. (a -> b) -> a -> b
$! ReduceEnv -> a
m ReduceEnv
e) ReduceEnv
e
{-# INLINE bindReduce #-}
instance Functor ReduceM where
fmap :: forall a b. (a -> b) -> ReduceM a -> ReduceM b
fmap = (a -> b) -> ReduceM a -> ReduceM b
forall a b. (a -> b) -> ReduceM a -> ReduceM b
fmapReduce
instance Applicative ReduceM where
pure :: forall a. a -> ReduceM a
pure a
x = (ReduceEnv -> a) -> ReduceM a
forall a. (ReduceEnv -> a) -> ReduceM a
ReduceM (a -> ReduceEnv -> a
forall a b. a -> b -> a
const a
x)
<*> :: forall a b. ReduceM (a -> b) -> ReduceM a -> ReduceM b
(<*>) = ReduceM (a -> b) -> ReduceM a -> ReduceM b
forall a b. ReduceM (a -> b) -> ReduceM a -> ReduceM b
apReduce
*> :: forall a b. ReduceM a -> ReduceM b -> ReduceM b
(*>) = ReduceM a -> ReduceM b -> ReduceM b
forall a b. ReduceM a -> ReduceM b -> ReduceM b
thenReduce
<* :: forall a b. ReduceM a -> ReduceM b -> ReduceM a
(<*) = ReduceM a -> ReduceM b -> ReduceM a
forall a b. ReduceM a -> ReduceM b -> ReduceM a
beforeReduce
instance Monad ReduceM where
return :: forall a. a -> ReduceM a
return = a -> ReduceM a
forall a. a -> ReduceM a
forall (f :: * -> *) a. Applicative f => a -> f a
pure
>>= :: forall a b. ReduceM a -> (a -> ReduceM b) -> ReduceM b
(>>=) = ReduceM a -> (a -> ReduceM b) -> ReduceM b
forall a b. ReduceM a -> (a -> ReduceM b) -> ReduceM b
bindReduce
>> :: forall a b. ReduceM a -> ReduceM b -> ReduceM b
(>>) = ReduceM a -> ReduceM b -> ReduceM b
forall a b. ReduceM a -> ReduceM b -> ReduceM b
forall (f :: * -> *) a b. Applicative f => f a -> f b -> f b
(*>)
instance MonadFail ReduceM where
fail :: forall a. String -> ReduceM a
fail = String -> ReduceM a
forall a. HasCallStack => String -> a
error
instance ReadTCState ReduceM where
getTCState :: ReduceM TCState
getTCState = (ReduceEnv -> TCState) -> ReduceM TCState
forall a. (ReduceEnv -> a) -> ReduceM a
ReduceM ReduceEnv -> TCState
redSt
locallyTCState :: forall a b. Lens' TCState a -> (a -> a) -> ReduceM b -> ReduceM b
locallyTCState Lens' TCState a
l a -> a
f = (ReduceEnv -> ReduceEnv) -> ReduceM b -> ReduceM b
forall a. (ReduceEnv -> ReduceEnv) -> ReduceM a -> ReduceM a
onReduceEnv ((ReduceEnv -> ReduceEnv) -> ReduceM b -> ReduceM b)
-> (ReduceEnv -> ReduceEnv) -> ReduceM b -> ReduceM b
forall a b. (a -> b) -> a -> b
$ (TCState -> TCState) -> ReduceEnv -> ReduceEnv
mapRedSt ((TCState -> TCState) -> ReduceEnv -> ReduceEnv)
-> (TCState -> TCState) -> ReduceEnv -> ReduceEnv
forall a b. (a -> b) -> a -> b
$ Lens' TCState a -> LensMap TCState a
forall o i. Lens' o i -> LensMap o i
over (a -> f a) -> TCState -> f TCState
Lens' TCState a
l a -> a
f
runReduceM :: ReduceM a -> TCM a
runReduceM :: forall a. ReduceM a -> TCM a
runReduceM ReduceM a
m = (IORef TCState -> TCEnv -> IO a) -> TCMT IO a
forall (m :: * -> *) a. (IORef TCState -> TCEnv -> m a) -> TCMT m a
TCM ((IORef TCState -> TCEnv -> IO a) -> TCMT IO a)
-> (IORef TCState -> TCEnv -> IO a) -> TCMT IO a
forall a b. (a -> b) -> a -> b
$ \ IORef TCState
r TCEnv
e -> do
s <- IORef TCState -> IO TCState
forall a. IORef a -> IO a
readIORef IORef TCState
r
E.evaluate $ unReduceM m $ ReduceEnv e s Nothing
runReduceF :: (a -> ReduceM b) -> TCM (a -> b)
runReduceF :: forall a b. (a -> ReduceM b) -> TCM (a -> b)
runReduceF a -> ReduceM b
f = do
e <- TCMT IO TCEnv
forall (m :: * -> *). MonadTCEnv m => m TCEnv
askTC
s <- getTC
return $ \a
x -> ReduceM b -> ReduceEnv -> b
forall a. ReduceM a -> ReduceEnv -> a
unReduceM (a -> ReduceM b
f a
x) (TCEnv -> TCState -> Maybe (MetaId -> ReduceM Bool) -> ReduceEnv
ReduceEnv TCEnv
e TCState
s Maybe (MetaId -> ReduceM Bool)
forall a. Maybe a
Nothing)
instance MonadTCEnv ReduceM where
askTC :: ReduceM TCEnv
askTC = (ReduceEnv -> TCEnv) -> ReduceM TCEnv
forall a. (ReduceEnv -> a) -> ReduceM a
ReduceM ReduceEnv -> TCEnv
redEnv
localTC :: forall a. (TCEnv -> TCEnv) -> ReduceM a -> ReduceM a
localTC = (ReduceEnv -> ReduceEnv) -> ReduceM a -> ReduceM a
forall a. (ReduceEnv -> ReduceEnv) -> ReduceM a -> ReduceM a
onReduceEnv ((ReduceEnv -> ReduceEnv) -> ReduceM a -> ReduceM a)
-> ((TCEnv -> TCEnv) -> ReduceEnv -> ReduceEnv)
-> (TCEnv -> TCEnv)
-> ReduceM a
-> ReduceM a
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (TCEnv -> TCEnv) -> ReduceEnv -> ReduceEnv
mapRedEnv
useR :: (ReadTCState m) => Lens' TCState a -> m a
useR :: forall (m :: * -> *) a. ReadTCState m => Lens' TCState a -> m a
useR Lens' TCState a
l = do
!x <- m TCState
forall (m :: * -> *). ReadTCState m => m TCState
getTCState m TCState -> (TCState -> a) -> m a
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> (TCState -> Lens' TCState a -> a
forall o i. o -> Lens' o i -> i
^. (a -> f a) -> TCState -> f TCState
Lens' TCState a
l)
return x
{-# INLINE useR #-}
askR :: ReduceM ReduceEnv
askR :: ReduceM ReduceEnv
askR = (ReduceEnv -> ReduceEnv) -> ReduceM ReduceEnv
forall a. (ReduceEnv -> a) -> ReduceM a
ReduceM ReduceEnv -> ReduceEnv
forall r (m :: * -> *). MonadReader r m => m r
ask
{-# INLINE askR #-}
localR :: (ReduceEnv -> ReduceEnv) -> ReduceM a -> ReduceM a
localR :: forall a. (ReduceEnv -> ReduceEnv) -> ReduceM a -> ReduceM a
localR ReduceEnv -> ReduceEnv
f = (ReduceEnv -> a) -> ReduceM a
forall a. (ReduceEnv -> a) -> ReduceM a
ReduceM ((ReduceEnv -> a) -> ReduceM a)
-> (ReduceM a -> ReduceEnv -> a) -> ReduceM a -> ReduceM a
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (ReduceEnv -> ReduceEnv) -> (ReduceEnv -> a) -> ReduceEnv -> a
forall a.
(ReduceEnv -> ReduceEnv) -> (ReduceEnv -> a) -> ReduceEnv -> a
forall r (m :: * -> *) a. MonadReader r m => (r -> r) -> m a -> m a
local ReduceEnv -> ReduceEnv
f ((ReduceEnv -> a) -> ReduceEnv -> a)
-> (ReduceM a -> ReduceEnv -> a) -> ReduceM a -> ReduceEnv -> a
forall b c a. (b -> c) -> (a -> b) -> a -> c
. ReduceM a -> ReduceEnv -> a
forall a. ReduceM a -> ReduceEnv -> a
unReduceM
{-# INLINE localR #-}
instance HasOptions ReduceM where
pragmaOptions :: ReduceM PragmaOptions
pragmaOptions = Lens' TCState PragmaOptions -> ReduceM PragmaOptions
forall (m :: * -> *) a. ReadTCState m => Lens' TCState a -> m a
useR (PragmaOptions -> f PragmaOptions) -> TCState -> f TCState
Lens' TCState PragmaOptions
stPragmaOptions
commandLineOptions :: ReduceM CommandLineOptions
commandLineOptions = do
p <- Lens' TCState PragmaOptions -> ReduceM PragmaOptions
forall (m :: * -> *) a. ReadTCState m => Lens' TCState a -> m a
useR (PragmaOptions -> f PragmaOptions) -> TCState -> f TCState
Lens' TCState PragmaOptions
stPragmaOptions
cl <- stPersistentOptions . stPersistentState <$> getTCState
return $ cl{ optPragmaOptions = p }
class ( Applicative m
, MonadTCEnv m
, ReadTCState m
, HasOptions m
) => MonadReduce m where
liftReduce :: ReduceM a -> m a
default liftReduce :: (MonadTrans t, MonadReduce n, t n ~ m) => ReduceM a -> m a
liftReduce = n a -> m a
n a -> t n a
forall (m :: * -> *) a. Monad m => m a -> t m a
forall (t :: (* -> *) -> * -> *) (m :: * -> *) a.
(MonadTrans t, Monad m) =>
m a -> t m a
lift (n a -> m a) -> (ReduceM a -> n a) -> ReduceM a -> m a
forall b c a. (b -> c) -> (a -> b) -> a -> c
. ReduceM a -> n a
forall a. ReduceM a -> n a
forall (m :: * -> *) a. MonadReduce m => ReduceM a -> m a
liftReduce
instance MonadReduce ReduceM where
liftReduce :: forall a. ReduceM a -> ReduceM a
liftReduce = ReduceM a -> ReduceM a
forall a. a -> a
id
instance MonadReduce m => MonadReduce (ChangeT m)
instance MonadReduce m => MonadReduce (ExceptT err m)
instance MonadReduce m => MonadReduce (IdentityT m)
instance MonadReduce m => MonadReduce (ListT m)
instance MonadReduce m => MonadReduce (MaybeT m)
instance MonadReduce m => MonadReduce (ReaderT r m)
instance MonadReduce m => MonadReduce (StateT w m)
instance (Monoid w, MonadReduce m) => MonadReduce (WriterT w m)
instance MonadReduce m => MonadReduce (BlockT m)
class Monad m => MonadTCEnv m where
askTC :: m TCEnv
localTC :: (TCEnv -> TCEnv) -> m a -> m a
default askTC :: (MonadTrans t, MonadTCEnv n, t n ~ m) => m TCEnv
askTC = n TCEnv -> t n TCEnv
forall (m :: * -> *) a. Monad m => m a -> t m a
forall (t :: (* -> *) -> * -> *) (m :: * -> *) a.
(MonadTrans t, Monad m) =>
m a -> t m a
lift n TCEnv
forall (m :: * -> *). MonadTCEnv m => m TCEnv
askTC
default localTC
:: (MonadTransControl t, MonadTCEnv n, t n ~ m)
=> (TCEnv -> TCEnv) -> m a -> m a
localTC = (n (StT t a) -> n (StT t a)) -> m a -> m a
(n (StT t a) -> n (StT t a)) -> t n a -> t n a
forall (t :: (* -> *) -> * -> *) (m :: * -> *) a b.
(MonadTransControl t, Monad (t m), Monad m) =>
(m (StT t a) -> m (StT t b)) -> t m a -> t m b
liftThrough ((n (StT t a) -> n (StT t a)) -> m a -> m a)
-> ((TCEnv -> TCEnv) -> n (StT t a) -> n (StT t a))
-> (TCEnv -> TCEnv)
-> m a
-> m a
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (TCEnv -> TCEnv) -> n (StT t a) -> n (StT t a)
forall a. (TCEnv -> TCEnv) -> n a -> n a
forall (m :: * -> *) a.
MonadTCEnv m =>
(TCEnv -> TCEnv) -> m a -> m a
localTC
instance MonadTCEnv m => MonadTCEnv (ChangeT m)
instance MonadTCEnv m => MonadTCEnv (ExceptT err m)
instance MonadTCEnv m => MonadTCEnv (IdentityT m)
instance MonadTCEnv m => MonadTCEnv (MaybeT m)
instance MonadTCEnv m => MonadTCEnv (ReaderT r m)
instance MonadTCEnv m => MonadTCEnv (StateT s m)
instance (Monoid w, MonadTCEnv m) => MonadTCEnv (WriterT w m)
instance MonadTCEnv m => MonadTCEnv (ListT m) where
localTC :: forall a. (TCEnv -> TCEnv) -> ListT m a -> ListT m a
localTC = (m (Maybe (a, ListT m a)) -> m (Maybe (a, ListT m a)))
-> ListT m a -> ListT m a
forall (m :: * -> *) a (n :: * -> *) b.
(m (Maybe (a, ListT m a)) -> n (Maybe (b, ListT n b)))
-> ListT m a -> ListT n b
mapListT ((m (Maybe (a, ListT m a)) -> m (Maybe (a, ListT m a)))
-> ListT m a -> ListT m a)
-> ((TCEnv -> TCEnv)
-> m (Maybe (a, ListT m a)) -> m (Maybe (a, ListT m a)))
-> (TCEnv -> TCEnv)
-> ListT m a
-> ListT m a
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (TCEnv -> TCEnv)
-> m (Maybe (a, ListT m a)) -> m (Maybe (a, ListT m a))
forall a. (TCEnv -> TCEnv) -> m a -> m a
forall (m :: * -> *) a.
MonadTCEnv m =>
(TCEnv -> TCEnv) -> m a -> m a
localTC
{-# INLINE asksTC #-}
asksTC :: MonadTCEnv m => (TCEnv -> a) -> m a
asksTC :: forall (m :: * -> *) a. MonadTCEnv m => (TCEnv -> a) -> m a
asksTC TCEnv -> a
f = TCEnv -> a
f (TCEnv -> a) -> m TCEnv -> m a
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> m TCEnv
forall (m :: * -> *). MonadTCEnv m => m TCEnv
askTC
{-# INLINE viewTC #-}
viewTC :: MonadTCEnv m => Lens' TCEnv a -> m a
viewTC :: forall (m :: * -> *) a. MonadTCEnv m => Lens' TCEnv a -> m a
viewTC Lens' TCEnv a
l = (TCEnv -> a) -> m a
forall (m :: * -> *) a. MonadTCEnv m => (TCEnv -> a) -> m a
asksTC (TCEnv -> Lens' TCEnv a -> a
forall o i. o -> Lens' o i -> i
^. (a -> f a) -> TCEnv -> f TCEnv
Lens' TCEnv a
l)
{-# INLINE locallyTC #-}
locallyTC :: MonadTCEnv m => Lens' TCEnv a -> (a -> a) -> m b -> m b
locallyTC :: forall (m :: * -> *) a b.
MonadTCEnv m =>
Lens' TCEnv a -> (a -> a) -> m b -> m b
locallyTC Lens' TCEnv a
l = (TCEnv -> TCEnv) -> m b -> m b
forall a. (TCEnv -> TCEnv) -> m a -> m a
forall (m :: * -> *) a.
MonadTCEnv m =>
(TCEnv -> TCEnv) -> m a -> m a
localTC ((TCEnv -> TCEnv) -> m b -> m b)
-> ((a -> a) -> TCEnv -> TCEnv) -> (a -> a) -> m b -> m b
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Lens' TCEnv a -> (a -> a) -> TCEnv -> TCEnv
forall o i. Lens' o i -> LensMap o i
over (a -> f a) -> TCEnv -> f TCEnv
Lens' TCEnv a
l
class Monad m => MonadTCState m where
getTC :: m TCState
putTC :: TCState -> m ()
modifyTC :: (TCState -> TCState) -> m ()
default getTC :: (MonadTrans t, MonadTCState n, t n ~ m) => m TCState
getTC = n TCState -> t n TCState
forall (m :: * -> *) a. Monad m => m a -> t m a
forall (t :: (* -> *) -> * -> *) (m :: * -> *) a.
(MonadTrans t, Monad m) =>
m a -> t m a
lift n TCState
forall (m :: * -> *). MonadTCState m => m TCState
getTC
default putTC :: (MonadTrans t, MonadTCState n, t n ~ m) => TCState -> m ()
putTC = n () -> m ()
n () -> t n ()
forall (m :: * -> *) a. Monad m => m a -> t m a
forall (t :: (* -> *) -> * -> *) (m :: * -> *) a.
(MonadTrans t, Monad m) =>
m a -> t m a
lift (n () -> m ()) -> (TCState -> n ()) -> TCState -> m ()
forall b c a. (b -> c) -> (a -> b) -> a -> c
. TCState -> n ()
forall (m :: * -> *). MonadTCState m => TCState -> m ()
putTC
default modifyTC :: (MonadTrans t, MonadTCState n, t n ~ m) => (TCState -> TCState) -> m ()
modifyTC = n () -> m ()
n () -> t n ()
forall (m :: * -> *) a. Monad m => m a -> t m a
forall (t :: (* -> *) -> * -> *) (m :: * -> *) a.
(MonadTrans t, Monad m) =>
m a -> t m a
lift (n () -> m ())
-> ((TCState -> TCState) -> n ()) -> (TCState -> TCState) -> m ()
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (TCState -> TCState) -> n ()
forall (m :: * -> *).
MonadTCState m =>
(TCState -> TCState) -> m ()
modifyTC
instance MonadTCState m => MonadTCState (MaybeT m)
instance MonadTCState m => MonadTCState (ListT m)
instance MonadTCState m => MonadTCState (ExceptT err m)
instance MonadTCState m => MonadTCState (ReaderT r m)
instance MonadTCState m => MonadTCState (StateT s m)
instance MonadTCState m => MonadTCState (ChangeT m)
instance MonadTCState m => MonadTCState (IdentityT m)
instance (Monoid w, MonadTCState m) => MonadTCState (WriterT w m)
{-# INLINE getsTC #-}
getsTC :: ReadTCState m => (TCState -> a) -> m a
getsTC :: forall (m :: * -> *) a. ReadTCState m => (TCState -> a) -> m a
getsTC TCState -> a
f = TCState -> a
f (TCState -> a) -> m TCState -> m a
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> m TCState
forall (m :: * -> *). ReadTCState m => m TCState
getTCState
{-# INLINE modifyTC' #-}
modifyTC' :: MonadTCState m => (TCState -> TCState) -> m ()
modifyTC' :: forall (m :: * -> *).
MonadTCState m =>
(TCState -> TCState) -> m ()
modifyTC' TCState -> TCState
f = do
s' <- m TCState
forall (m :: * -> *). MonadTCState m => m TCState
getTC
putTC $! f s'
{-# INLINE useTC #-}
useTC :: ReadTCState m => Lens' TCState a -> m a
useTC :: forall (m :: * -> *) a. ReadTCState m => Lens' TCState a -> m a
useTC Lens' TCState a
l = do
!x <- (TCState -> a) -> m a
forall (m :: * -> *) a. ReadTCState m => (TCState -> a) -> m a
getsTC (TCState -> Lens' TCState a -> a
forall o i. o -> Lens' o i -> i
^. (a -> f a) -> TCState -> f TCState
Lens' TCState a
l)
return x
infix 4 `setTCLens`
{-# INLINE setTCLens #-}
setTCLens :: MonadTCState m => Lens' TCState a -> a -> m ()
setTCLens :: forall (m :: * -> *) a.
MonadTCState m =>
Lens' TCState a -> a -> m ()
setTCLens Lens' TCState a
l = (TCState -> TCState) -> m ()
forall (m :: * -> *).
MonadTCState m =>
(TCState -> TCState) -> m ()
modifyTC ((TCState -> TCState) -> m ())
-> (a -> TCState -> TCState) -> a -> m ()
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Lens' TCState a -> a -> TCState -> TCState
forall o i. Lens' o i -> LensSet o i
set (a -> f a) -> TCState -> f TCState
Lens' TCState a
l
{-# INLINE setTCLens' #-}
setTCLens' :: MonadTCState m => Lens' TCState a -> a -> m ()
setTCLens' :: forall (m :: * -> *) a.
MonadTCState m =>
Lens' TCState a -> a -> m ()
setTCLens' Lens' TCState a
l = (TCState -> TCState) -> m ()
forall (m :: * -> *).
MonadTCState m =>
(TCState -> TCState) -> m ()
modifyTC' ((TCState -> TCState) -> m ())
-> (a -> TCState -> TCState) -> a -> m ()
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Lens' TCState a -> a -> TCState -> TCState
forall o i. Lens' o i -> LensSet o i
set (a -> f a) -> TCState -> f TCState
Lens' TCState a
l
{-# INLINE modifyTCLens #-}
modifyTCLens :: MonadTCState m => Lens' TCState a -> (a -> a) -> m ()
modifyTCLens :: forall (m :: * -> *) a.
MonadTCState m =>
Lens' TCState a -> (a -> a) -> m ()
modifyTCLens Lens' TCState a
l = (TCState -> TCState) -> m ()
forall (m :: * -> *).
MonadTCState m =>
(TCState -> TCState) -> m ()
modifyTC ((TCState -> TCState) -> m ())
-> ((a -> a) -> TCState -> TCState) -> (a -> a) -> m ()
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Lens' TCState a -> (a -> a) -> TCState -> TCState
forall o i. Lens' o i -> LensMap o i
over (a -> f a) -> TCState -> f TCState
Lens' TCState a
l
{-# INLINE modifyTCLens' #-}
modifyTCLens' :: MonadTCState m => Lens' TCState a -> (a -> a) -> m ()
modifyTCLens' :: forall (m :: * -> *) a.
MonadTCState m =>
Lens' TCState a -> (a -> a) -> m ()
modifyTCLens' Lens' TCState a
l = (TCState -> TCState) -> m ()
forall (m :: * -> *).
MonadTCState m =>
(TCState -> TCState) -> m ()
modifyTC' ((TCState -> TCState) -> m ())
-> ((a -> a) -> TCState -> TCState) -> (a -> a) -> m ()
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Lens' TCState a -> (a -> a) -> TCState -> TCState
forall o i. Lens' o i -> LensMap o i
over (a -> f a) -> TCState -> f TCState
Lens' TCState a
l
{-# INLINE modifyTCLensM #-}
modifyTCLensM :: (MonadTCState m, ReadTCState m) => Lens' TCState a -> (a -> m a) -> m ()
modifyTCLensM :: forall (m :: * -> *) a.
(MonadTCState m, ReadTCState m) =>
Lens' TCState a -> (a -> m a) -> m ()
modifyTCLensM Lens' TCState a
l a -> m a
f = Lens' TCState a -> m a
forall (m :: * -> *) a. ReadTCState m => Lens' TCState a -> m a
useTC (a -> f a) -> TCState -> f TCState
Lens' TCState a
l m a -> (a -> m a) -> m a
forall a b. m a -> (a -> m b) -> m b
forall (m :: * -> *) a b. Monad m => m a -> (a -> m b) -> m b
>>= a -> m a
f m a -> (a -> m ()) -> m ()
forall a b. m a -> (a -> m b) -> m b
forall (m :: * -> *) a b. Monad m => m a -> (a -> m b) -> m b
>>= Lens' TCState a -> a -> m ()
forall (m :: * -> *) a.
MonadTCState m =>
Lens' TCState a -> a -> m ()
setTCLens (a -> f a) -> TCState -> f TCState
Lens' TCState a
l
{-# INLINE stateTCLens #-}
stateTCLens :: (MonadTCState m, ReadTCState m) => Lens' TCState a -> (a -> (r , a)) -> m r
stateTCLens :: forall (m :: * -> *) a r.
(MonadTCState m, ReadTCState m) =>
Lens' TCState a -> (a -> (r, a)) -> m r
stateTCLens Lens' TCState a
l a -> (r, a)
f = Lens' TCState a -> (a -> m (r, a)) -> m r
forall (m :: * -> *) a r.
(MonadTCState m, ReadTCState m) =>
Lens' TCState a -> (a -> m (r, a)) -> m r
stateTCLensM (a -> f a) -> TCState -> f TCState
Lens' TCState a
l ((a -> m (r, a)) -> m r) -> (a -> m (r, a)) -> m r
forall a b. (a -> b) -> a -> b
$ (r, a) -> m (r, a)
forall a. a -> m a
forall (m :: * -> *) a. Monad m => a -> m a
return ((r, a) -> m (r, a)) -> (a -> (r, a)) -> a -> m (r, a)
forall b c a. (b -> c) -> (a -> b) -> a -> c
. a -> (r, a)
f
{-# INLINE stateTCLensM #-}
stateTCLensM :: (MonadTCState m, ReadTCState m) => Lens' TCState a -> (a -> m (r , a)) -> m r
stateTCLensM :: forall (m :: * -> *) a r.
(MonadTCState m, ReadTCState m) =>
Lens' TCState a -> (a -> m (r, a)) -> m r
stateTCLensM Lens' TCState a
l a -> m (r, a)
f = do
a <- Lens' TCState a -> m a
forall (m :: * -> *) a. ReadTCState m => Lens' TCState a -> m a
useTC (a -> f a) -> TCState -> f TCState
Lens' TCState a
l
(result , a') <- f a
result <$ setTCLens l a'
class Monad m => MonadBlock m where
patternViolation :: Blocker -> m a
default patternViolation :: (MonadTrans t, MonadBlock n, m ~ t n) => Blocker -> m a
patternViolation = n a -> m a
n a -> t n a
forall (m :: * -> *) a. Monad m => m a -> t m a
forall (t :: (* -> *) -> * -> *) (m :: * -> *) a.
(MonadTrans t, Monad m) =>
m a -> t m a
lift (n a -> m a) -> (Blocker -> n a) -> Blocker -> m a
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Blocker -> n a
forall a. Blocker -> n a
forall (m :: * -> *) a. MonadBlock m => Blocker -> m a
patternViolation
catchPatternErr :: (Blocker -> m a) -> m a -> m a
newtype BlockT m a = BlockT { forall (m :: * -> *) a. BlockT m a -> ExceptT Blocker m a
unBlockT :: ExceptT Blocker m a }
deriving ( (forall a b. (a -> b) -> BlockT m a -> BlockT m b)
-> (forall a b. a -> BlockT m b -> BlockT m a)
-> Functor (BlockT m)
forall a b. a -> BlockT m b -> BlockT m a
forall a b. (a -> b) -> BlockT m a -> BlockT m b
forall (m :: * -> *) a b.
Functor m =>
a -> BlockT m b -> BlockT m a
forall (m :: * -> *) a b.
Functor m =>
(a -> b) -> BlockT m a -> BlockT m b
forall (f :: * -> *).
(forall a b. (a -> b) -> f a -> f b)
-> (forall a b. a -> f b -> f a) -> Functor f
$cfmap :: forall (m :: * -> *) a b.
Functor m =>
(a -> b) -> BlockT m a -> BlockT m b
fmap :: forall a b. (a -> b) -> BlockT m a -> BlockT m b
$c<$ :: forall (m :: * -> *) a b.
Functor m =>
a -> BlockT m b -> BlockT m a
<$ :: forall a b. a -> BlockT m b -> BlockT m a
Functor, Functor (BlockT m)
Functor (BlockT m) =>
(forall a. a -> BlockT m a)
-> (forall a b. BlockT m (a -> b) -> BlockT m a -> BlockT m b)
-> (forall a b c.
(a -> b -> c) -> BlockT m a -> BlockT m b -> BlockT m c)
-> (forall a b. BlockT m a -> BlockT m b -> BlockT m b)
-> (forall a b. BlockT m a -> BlockT m b -> BlockT m a)
-> Applicative (BlockT m)
forall a. a -> BlockT m a
forall a b. BlockT m a -> BlockT m b -> BlockT m a
forall a b. BlockT m a -> BlockT m b -> BlockT m b
forall a b. BlockT m (a -> b) -> BlockT m a -> BlockT m b
forall a b c.
(a -> b -> c) -> BlockT m a -> BlockT m b -> BlockT m c
forall (m :: * -> *). Monad m => Functor (BlockT m)
forall (m :: * -> *) a. Monad m => a -> BlockT m a
forall (m :: * -> *) a b.
Monad m =>
BlockT m a -> BlockT m b -> BlockT m a
forall (m :: * -> *) a b.
Monad m =>
BlockT m a -> BlockT m b -> BlockT m b
forall (m :: * -> *) a b.
Monad m =>
BlockT m (a -> b) -> BlockT m a -> BlockT m b
forall (m :: * -> *) a b c.
Monad m =>
(a -> b -> c) -> BlockT m a -> BlockT m b -> BlockT m c
forall (f :: * -> *).
Functor f =>
(forall a. a -> f a)
-> (forall a b. f (a -> b) -> f a -> f b)
-> (forall a b c. (a -> b -> c) -> f a -> f b -> f c)
-> (forall a b. f a -> f b -> f b)
-> (forall a b. f a -> f b -> f a)
-> Applicative f
$cpure :: forall (m :: * -> *) a. Monad m => a -> BlockT m a
pure :: forall a. a -> BlockT m a
$c<*> :: forall (m :: * -> *) a b.
Monad m =>
BlockT m (a -> b) -> BlockT m a -> BlockT m b
<*> :: forall a b. BlockT m (a -> b) -> BlockT m a -> BlockT m b
$cliftA2 :: forall (m :: * -> *) a b c.
Monad m =>
(a -> b -> c) -> BlockT m a -> BlockT m b -> BlockT m c
liftA2 :: forall a b c.
(a -> b -> c) -> BlockT m a -> BlockT m b -> BlockT m c
$c*> :: forall (m :: * -> *) a b.
Monad m =>
BlockT m a -> BlockT m b -> BlockT m b
*> :: forall a b. BlockT m a -> BlockT m b -> BlockT m b
$c<* :: forall (m :: * -> *) a b.
Monad m =>
BlockT m a -> BlockT m b -> BlockT m a
<* :: forall a b. BlockT m a -> BlockT m b -> BlockT m a
Applicative, Applicative (BlockT m)
Applicative (BlockT m) =>
(forall a b. BlockT m a -> (a -> BlockT m b) -> BlockT m b)
-> (forall a b. BlockT m a -> BlockT m b -> BlockT m b)
-> (forall a. a -> BlockT m a)
-> Monad (BlockT m)
forall a. a -> BlockT m a
forall a b. BlockT m a -> BlockT m b -> BlockT m b
forall a b. BlockT m a -> (a -> BlockT m b) -> BlockT m b
forall (m :: * -> *). Monad m => Applicative (BlockT m)
forall (m :: * -> *) a. Monad m => a -> BlockT m a
forall (m :: * -> *) a b.
Monad m =>
BlockT m a -> BlockT m b -> BlockT m b
forall (m :: * -> *) a b.
Monad m =>
BlockT m a -> (a -> BlockT m b) -> BlockT m b
forall (m :: * -> *).
Applicative m =>
(forall a b. m a -> (a -> m b) -> m b)
-> (forall a b. m a -> m b -> m b)
-> (forall a. a -> m a)
-> Monad m
$c>>= :: forall (m :: * -> *) a b.
Monad m =>
BlockT m a -> (a -> BlockT m b) -> BlockT m b
>>= :: forall a b. BlockT m a -> (a -> BlockT m b) -> BlockT m b
$c>> :: forall (m :: * -> *) a b.
Monad m =>
BlockT m a -> BlockT m b -> BlockT m b
>> :: forall a b. BlockT m a -> BlockT m b -> BlockT m b
$creturn :: forall (m :: * -> *) a. Monad m => a -> BlockT m a
return :: forall a. a -> BlockT m a
Monad, (forall (m :: * -> *). Monad m => Monad (BlockT m)) =>
(forall (m :: * -> *) a. Monad m => m a -> BlockT m a)
-> MonadTrans BlockT
forall (m :: * -> *). Monad m => Monad (BlockT m)
forall (m :: * -> *) a. Monad m => m a -> BlockT m a
forall (t :: (* -> *) -> * -> *).
(forall (m :: * -> *). Monad m => Monad (t m)) =>
(forall (m :: * -> *) a. Monad m => m a -> t m a) -> MonadTrans t
$clift :: forall (m :: * -> *) a. Monad m => m a -> BlockT m a
lift :: forall (m :: * -> *) a. Monad m => m a -> BlockT m a
MonadTrans
, Monad (BlockT m)
Monad (BlockT m) =>
(forall a. IO a -> BlockT m a) -> MonadIO (BlockT m)
forall a. IO a -> BlockT m a
forall (m :: * -> *).
Monad m =>
(forall a. IO a -> m a) -> MonadIO m
forall (m :: * -> *). MonadIO m => Monad (BlockT m)
forall (m :: * -> *) a. MonadIO m => IO a -> BlockT m a
$cliftIO :: forall (m :: * -> *) a. MonadIO m => IO a -> BlockT m a
liftIO :: forall a. IO a -> BlockT m a
MonadIO, Monad (BlockT m)
Monad (BlockT m) =>
(forall a. String -> BlockT m a) -> MonadFail (BlockT m)
forall a. String -> BlockT m a
forall (m :: * -> *).
Monad m =>
(forall a. String -> m a) -> MonadFail m
forall (m :: * -> *). MonadFail m => Monad (BlockT m)
forall (m :: * -> *) a. MonadFail m => String -> BlockT m a
$cfail :: forall (m :: * -> *) a. MonadFail m => String -> BlockT m a
fail :: forall a. String -> BlockT m a
MonadFail
, Monad (BlockT m)
BlockT m TCState
Monad (BlockT m) =>
BlockT m TCState
-> (forall a b.
Lens' TCState a -> (a -> a) -> BlockT m b -> BlockT m b)
-> (forall a. (TCState -> TCState) -> BlockT m a -> BlockT m a)
-> ReadTCState (BlockT m)
forall a. (TCState -> TCState) -> BlockT m a -> BlockT m a
forall a b. Lens' TCState a -> (a -> a) -> BlockT m b -> BlockT m b
forall (m :: * -> *).
Monad m =>
m TCState
-> (forall a b. Lens' TCState a -> (a -> a) -> m b -> m b)
-> (forall a. (TCState -> TCState) -> m a -> m a)
-> ReadTCState m
forall (m :: * -> *). ReadTCState m => Monad (BlockT m)
forall (m :: * -> *). ReadTCState m => BlockT m TCState
forall (m :: * -> *) a.
ReadTCState m =>
(TCState -> TCState) -> BlockT m a -> BlockT m a
forall (m :: * -> *) a b.
ReadTCState m =>
Lens' TCState a -> (a -> a) -> BlockT m b -> BlockT m b
$cgetTCState :: forall (m :: * -> *). ReadTCState m => BlockT m TCState
getTCState :: BlockT m TCState
$clocallyTCState :: forall (m :: * -> *) a b.
ReadTCState m =>
Lens' TCState a -> (a -> a) -> BlockT m b -> BlockT m b
locallyTCState :: forall a b. Lens' TCState a -> (a -> a) -> BlockT m b -> BlockT m b
$cwithTCState :: forall (m :: * -> *) a.
ReadTCState m =>
(TCState -> TCState) -> BlockT m a -> BlockT m a
withTCState :: forall a. (TCState -> TCState) -> BlockT m a -> BlockT m a
ReadTCState, Monad (BlockT m)
Functor (BlockT m)
Applicative (BlockT m)
BlockT m PragmaOptions
BlockT m CommandLineOptions
(Functor (BlockT m), Applicative (BlockT m), Monad (BlockT m)) =>
BlockT m PragmaOptions
-> BlockT m CommandLineOptions -> HasOptions (BlockT m)
forall (m :: * -> *).
(Functor m, Applicative m, Monad m) =>
m PragmaOptions -> m CommandLineOptions -> HasOptions m
forall (m :: * -> *). HasOptions m => Monad (BlockT m)
forall (m :: * -> *). HasOptions m => Functor (BlockT m)
forall (m :: * -> *). HasOptions m => Applicative (BlockT m)
forall (m :: * -> *). HasOptions m => BlockT m PragmaOptions
forall (m :: * -> *). HasOptions m => BlockT m CommandLineOptions
$cpragmaOptions :: forall (m :: * -> *). HasOptions m => BlockT m PragmaOptions
pragmaOptions :: BlockT m PragmaOptions
$ccommandLineOptions :: forall (m :: * -> *). HasOptions m => BlockT m CommandLineOptions
commandLineOptions :: BlockT m CommandLineOptions
HasOptions
, Monad (BlockT m)
BlockT m TCEnv
Monad (BlockT m) =>
BlockT m TCEnv
-> (forall a. (TCEnv -> TCEnv) -> BlockT m a -> BlockT m a)
-> MonadTCEnv (BlockT m)
forall a. (TCEnv -> TCEnv) -> BlockT m a -> BlockT m a
forall (m :: * -> *).
Monad m =>
m TCEnv
-> (forall a. (TCEnv -> TCEnv) -> m a -> m a) -> MonadTCEnv m
forall (m :: * -> *). MonadTCEnv m => Monad (BlockT m)
forall (m :: * -> *). MonadTCEnv m => BlockT m TCEnv
forall (m :: * -> *) a.
MonadTCEnv m =>
(TCEnv -> TCEnv) -> BlockT m a -> BlockT m a
$caskTC :: forall (m :: * -> *). MonadTCEnv m => BlockT m TCEnv
askTC :: BlockT m TCEnv
$clocalTC :: forall (m :: * -> *) a.
MonadTCEnv m =>
(TCEnv -> TCEnv) -> BlockT m a -> BlockT m a
localTC :: forall a. (TCEnv -> TCEnv) -> BlockT m a -> BlockT m a
MonadTCEnv, Monad (BlockT m)
BlockT m TCState
Monad (BlockT m) =>
BlockT m TCState
-> (TCState -> BlockT m ())
-> ((TCState -> TCState) -> BlockT m ())
-> MonadTCState (BlockT m)
TCState -> BlockT m ()
(TCState -> TCState) -> BlockT m ()
forall (m :: * -> *).
Monad m =>
m TCState
-> (TCState -> m ())
-> ((TCState -> TCState) -> m ())
-> MonadTCState m
forall (m :: * -> *). MonadTCState m => Monad (BlockT m)
forall (m :: * -> *). MonadTCState m => BlockT m TCState
forall (m :: * -> *). MonadTCState m => TCState -> BlockT m ()
forall (m :: * -> *).
MonadTCState m =>
(TCState -> TCState) -> BlockT m ()
$cgetTC :: forall (m :: * -> *). MonadTCState m => BlockT m TCState
getTC :: BlockT m TCState
$cputTC :: forall (m :: * -> *). MonadTCState m => TCState -> BlockT m ()
putTC :: TCState -> BlockT m ()
$cmodifyTC :: forall (m :: * -> *).
MonadTCState m =>
(TCState -> TCState) -> BlockT m ()
modifyTC :: (TCState -> TCState) -> BlockT m ()
MonadTCState, Applicative (BlockT m)
MonadIO (BlockT m)
HasOptions (BlockT m)
MonadTCState (BlockT m)
MonadTCEnv (BlockT m)
(Applicative (BlockT m), MonadIO (BlockT m), MonadTCEnv (BlockT m),
MonadTCState (BlockT m), HasOptions (BlockT m)) =>
(forall a. TCM a -> BlockT m a) -> MonadTCM (BlockT m)
forall a. TCM a -> BlockT m a
forall (tcm :: * -> *).
(Applicative tcm, MonadIO tcm, MonadTCEnv tcm, MonadTCState tcm,
HasOptions tcm) =>
(forall a. TCM a -> tcm a) -> MonadTCM tcm
forall (m :: * -> *). MonadTCM m => Applicative (BlockT m)
forall (m :: * -> *). MonadTCM m => MonadIO (BlockT m)
forall (m :: * -> *). MonadTCM m => HasOptions (BlockT m)
forall (m :: * -> *). MonadTCM m => MonadTCState (BlockT m)
forall (m :: * -> *). MonadTCM m => MonadTCEnv (BlockT m)
forall (m :: * -> *) a. MonadTCM m => TCM a -> BlockT m a
$cliftTCM :: forall (m :: * -> *) a. MonadTCM m => TCM a -> BlockT m a
liftTCM :: forall a. TCM a -> BlockT m a
MonadTCM
)
instance Monad m => MonadBlock (BlockT m) where
patternViolation :: forall a. Blocker -> BlockT m a
patternViolation = ExceptT Blocker m a -> BlockT m a
forall (m :: * -> *) a. ExceptT Blocker m a -> BlockT m a
BlockT (ExceptT Blocker m a -> BlockT m a)
-> (Blocker -> ExceptT Blocker m a) -> Blocker -> BlockT m a
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Blocker -> ExceptT Blocker m a
forall a. Blocker -> ExceptT Blocker m a
forall e (m :: * -> *) a. MonadError e m => e -> m a
throwError
catchPatternErr :: forall a. (Blocker -> BlockT m a) -> BlockT m a -> BlockT m a
catchPatternErr Blocker -> BlockT m a
h BlockT m a
f = ExceptT Blocker m a -> BlockT m a
forall (m :: * -> *) a. ExceptT Blocker m a -> BlockT m a
BlockT (ExceptT Blocker m a -> BlockT m a)
-> ExceptT Blocker m a -> BlockT m a
forall a b. (a -> b) -> a -> b
$ ExceptT Blocker m a
-> (Blocker -> ExceptT Blocker m a) -> ExceptT Blocker m a
forall a.
ExceptT Blocker m a
-> (Blocker -> ExceptT Blocker m a) -> ExceptT Blocker m a
forall e (m :: * -> *) a.
MonadError e m =>
m a -> (e -> m a) -> m a
catchError (BlockT m a -> ExceptT Blocker m a
forall (m :: * -> *) a. BlockT m a -> ExceptT Blocker m a
unBlockT BlockT m a
f) (BlockT m a -> ExceptT Blocker m a
forall (m :: * -> *) a. BlockT m a -> ExceptT Blocker m a
unBlockT (BlockT m a -> ExceptT Blocker m a)
-> (Blocker -> BlockT m a) -> Blocker -> ExceptT Blocker m a
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Blocker -> BlockT m a
h)
instance Monad m => MonadBlock (ExceptT TCErr m) where
patternViolation :: forall a. Blocker -> ExceptT TCErr m a
patternViolation = TCErr -> ExceptT TCErr m a
forall a. TCErr -> ExceptT TCErr m a
forall e (m :: * -> *) a. MonadError e m => e -> m a
throwError (TCErr -> ExceptT TCErr m a)
-> (Blocker -> TCErr) -> Blocker -> ExceptT TCErr m a
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Blocker -> TCErr
PatternErr
catchPatternErr :: forall a.
(Blocker -> ExceptT TCErr m a)
-> ExceptT TCErr m a -> ExceptT TCErr m a
catchPatternErr Blocker -> ExceptT TCErr m a
h ExceptT TCErr m a
f = ExceptT TCErr m a
-> (TCErr -> ExceptT TCErr m a) -> ExceptT TCErr m a
forall a.
ExceptT TCErr m a
-> (TCErr -> ExceptT TCErr m a) -> ExceptT TCErr m a
forall e (m :: * -> *) a.
MonadError e m =>
m a -> (e -> m a) -> m a
catchError ExceptT TCErr m a
f ((TCErr -> ExceptT TCErr m a) -> ExceptT TCErr m a)
-> (TCErr -> ExceptT TCErr m a) -> ExceptT TCErr m a
forall a b. (a -> b) -> a -> b
$ \case
PatternErr Blocker
b -> Blocker -> ExceptT TCErr m a
h Blocker
b
TCErr
err -> TCErr -> ExceptT TCErr m a
forall a. TCErr -> ExceptT TCErr m a
forall e (m :: * -> *) a. MonadError e m => e -> m a
throwError TCErr
err
runBlocked :: Monad m => BlockT m a -> m (Either Blocker a)
runBlocked :: forall (m :: * -> *) a.
Monad m =>
BlockT m a -> m (Either Blocker a)
runBlocked = ExceptT Blocker m a -> m (Either Blocker a)
forall e (m :: * -> *) a. ExceptT e m a -> m (Either e a)
runExceptT (ExceptT Blocker m a -> m (Either Blocker a))
-> (BlockT m a -> ExceptT Blocker m a)
-> BlockT m a
-> m (Either Blocker a)
forall b c a. (b -> c) -> (a -> b) -> a -> c
. BlockT m a -> ExceptT Blocker m a
forall (m :: * -> *) a. BlockT m a -> ExceptT Blocker m a
unBlockT
{-# INLINE runBlocked #-}
instance MonadBlock m => MonadBlock (MaybeT m) where
catchPatternErr :: forall a. (Blocker -> MaybeT m a) -> MaybeT m a -> MaybeT m a
catchPatternErr Blocker -> MaybeT m a
h MaybeT m a
m = m (Maybe a) -> MaybeT m a
forall (m :: * -> *) a. m (Maybe a) -> MaybeT m a
MaybeT (m (Maybe a) -> MaybeT m a) -> m (Maybe a) -> MaybeT m a
forall a b. (a -> b) -> a -> b
$ (Blocker -> m (Maybe a)) -> m (Maybe a) -> m (Maybe a)
forall a. (Blocker -> m a) -> m a -> m a
forall (m :: * -> *) a.
MonadBlock m =>
(Blocker -> m a) -> m a -> m a
catchPatternErr (MaybeT m a -> m (Maybe a)
forall (m :: * -> *) a. MaybeT m a -> m (Maybe a)
runMaybeT (MaybeT m a -> m (Maybe a))
-> (Blocker -> MaybeT m a) -> Blocker -> m (Maybe a)
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Blocker -> MaybeT m a
h) (m (Maybe a) -> m (Maybe a)) -> m (Maybe a) -> m (Maybe a)
forall a b. (a -> b) -> a -> b
$ MaybeT m a -> m (Maybe a)
forall (m :: * -> *) a. MaybeT m a -> m (Maybe a)
runMaybeT MaybeT m a
m
instance MonadBlock m => MonadBlock (ReaderT e m) where
catchPatternErr :: forall a.
(Blocker -> ReaderT e m a) -> ReaderT e m a -> ReaderT e m a
catchPatternErr Blocker -> ReaderT e m a
h ReaderT e m a
m = (e -> m a) -> ReaderT e m a
forall r (m :: * -> *) a. (r -> m a) -> ReaderT r m a
ReaderT ((e -> m a) -> ReaderT e m a) -> (e -> m a) -> ReaderT e m a
forall a b. (a -> b) -> a -> b
$ \ e
e ->
let run :: ReaderT e m a -> m a
run = (ReaderT e m a -> e -> m a) -> e -> ReaderT e m a -> m a
forall a b c. (a -> b -> c) -> b -> a -> c
flip ReaderT e m a -> e -> m a
forall r (m :: * -> *) a. ReaderT r m a -> r -> m a
runReaderT e
e in (Blocker -> m a) -> m a -> m a
forall a. (Blocker -> m a) -> m a -> m a
forall (m :: * -> *) a.
MonadBlock m =>
(Blocker -> m a) -> m a -> m a
catchPatternErr (ReaderT e m a -> m a
run (ReaderT e m a -> m a)
-> (Blocker -> ReaderT e m a) -> Blocker -> m a
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Blocker -> ReaderT e m a
h) (ReaderT e m a -> m a
run ReaderT e m a
m)
newtype TCMT m a = TCM { forall (m :: * -> *) a. TCMT m a -> IORef TCState -> TCEnv -> m a
unTCM :: IORef TCState -> TCEnv -> m a }
type TCM = TCMT IO
{-# SPECIALIZE INLINE mapTCMT :: (forall a. IO a -> IO a) -> TCM a -> TCM a #-}
mapTCMT :: (forall a. m a -> n a) -> TCMT m a -> TCMT n a
mapTCMT :: forall (m :: * -> *) (n :: * -> *) a.
(forall a. m a -> n a) -> TCMT m a -> TCMT n a
mapTCMT forall a. m a -> n a
f (TCM IORef TCState -> TCEnv -> m a
m) = (IORef TCState -> TCEnv -> n a) -> TCMT n a
forall (m :: * -> *) a. (IORef TCState -> TCEnv -> m a) -> TCMT m a
TCM ((IORef TCState -> TCEnv -> n a) -> TCMT n a)
-> (IORef TCState -> TCEnv -> n a) -> TCMT n a
forall a b. (a -> b) -> a -> b
$ \ IORef TCState
s TCEnv
e -> m a -> n a
forall a. m a -> n a
f (IORef TCState -> TCEnv -> m a
m IORef TCState
s TCEnv
e)
pureTCM :: MonadIO m => (TCState -> TCEnv -> a) -> TCMT m a
pureTCM :: forall (m :: * -> *) a.
MonadIO m =>
(TCState -> TCEnv -> a) -> TCMT m a
pureTCM TCState -> TCEnv -> a
f = (IORef TCState -> TCEnv -> m a) -> TCMT m a
forall (m :: * -> *) a. (IORef TCState -> TCEnv -> m a) -> TCMT m a
TCM ((IORef TCState -> TCEnv -> m a) -> TCMT m a)
-> (IORef TCState -> TCEnv -> m a) -> TCMT m a
forall a b. (a -> b) -> a -> b
$ \ IORef TCState
r TCEnv
e -> do
s <- IO TCState -> m TCState
forall a. IO a -> m a
forall (m :: * -> *) a. MonadIO m => IO a -> m a
liftIO (IO TCState -> m TCState) -> IO TCState -> m TCState
forall a b. (a -> b) -> a -> b
$ IORef TCState -> IO TCState
forall a. IORef a -> IO a
readIORef IORef TCState
r
return (f s e)
{-# INLINE pureTCM #-}
returnTCMT :: Applicative m => a -> TCMT m a
returnTCMT :: forall (m :: * -> *) a. Applicative m => a -> TCMT m a
returnTCMT = \a
x -> (IORef TCState -> TCEnv -> m a) -> TCMT m a
forall (m :: * -> *) a. (IORef TCState -> TCEnv -> m a) -> TCMT m a
TCM ((IORef TCState -> TCEnv -> m a) -> TCMT m a)
-> (IORef TCState -> TCEnv -> m a) -> TCMT m a
forall a b. (a -> b) -> a -> b
$ \IORef TCState
_ TCEnv
_ -> a -> m a
forall a. a -> m a
forall (f :: * -> *) a. Applicative f => a -> f a
pure a
x
{-# INLINE returnTCMT #-}
bindTCMT :: Monad m => TCMT m a -> (a -> TCMT m b) -> TCMT m b
bindTCMT :: forall (m :: * -> *) a b.
Monad m =>
TCMT m a -> (a -> TCMT m b) -> TCMT m b
bindTCMT = \(TCM IORef TCState -> TCEnv -> m a
m) a -> TCMT m b
k -> (IORef TCState -> TCEnv -> m b) -> TCMT m b
forall (m :: * -> *) a. (IORef TCState -> TCEnv -> m a) -> TCMT m a
TCM ((IORef TCState -> TCEnv -> m b) -> TCMT m b)
-> (IORef TCState -> TCEnv -> m b) -> TCMT m b
forall a b. (a -> b) -> a -> b
$ \IORef TCState
r TCEnv
e -> IORef TCState -> TCEnv -> m a
m IORef TCState
r TCEnv
e m a -> (a -> m b) -> m b
forall a b. m a -> (a -> m b) -> m b
forall (m :: * -> *) a b. Monad m => m a -> (a -> m b) -> m b
>>= \a
x -> TCMT m b -> IORef TCState -> TCEnv -> m b
forall (m :: * -> *) a. TCMT m a -> IORef TCState -> TCEnv -> m a
unTCM (a -> TCMT m b
k a
x) IORef TCState
r TCEnv
e
{-# INLINE bindTCMT #-}
thenTCMT :: Applicative m => TCMT m a -> TCMT m b -> TCMT m b
thenTCMT :: forall (m :: * -> *) a b.
Applicative m =>
TCMT m a -> TCMT m b -> TCMT m b
thenTCMT = \(TCM IORef TCState -> TCEnv -> m a
m1) (TCM IORef TCState -> TCEnv -> m b
m2) -> (IORef TCState -> TCEnv -> m b) -> TCMT m b
forall (m :: * -> *) a. (IORef TCState -> TCEnv -> m a) -> TCMT m a
TCM ((IORef TCState -> TCEnv -> m b) -> TCMT m b)
-> (IORef TCState -> TCEnv -> m b) -> TCMT m b
forall a b. (a -> b) -> a -> b
$ \IORef TCState
r TCEnv
e -> IORef TCState -> TCEnv -> m a
m1 IORef TCState
r TCEnv
e m a -> m b -> m b
forall a b. m a -> m b -> m b
forall (f :: * -> *) a b. Applicative f => f a -> f b -> f b
*> IORef TCState -> TCEnv -> m b
m2 IORef TCState
r TCEnv
e
{-# INLINE thenTCMT #-}
instance Functor m => Functor (TCMT m) where
fmap :: forall a b. (a -> b) -> TCMT m a -> TCMT m b
fmap = (a -> b) -> TCMT m a -> TCMT m b
forall (m :: * -> *) a b.
Functor m =>
(a -> b) -> TCMT m a -> TCMT m b
fmapTCMT; {-# INLINE fmap #-}
fmapTCMT :: Functor m => (a -> b) -> TCMT m a -> TCMT m b
fmapTCMT :: forall (m :: * -> *) a b.
Functor m =>
(a -> b) -> TCMT m a -> TCMT m b
fmapTCMT = \a -> b
f (TCM IORef TCState -> TCEnv -> m a
m) -> (IORef TCState -> TCEnv -> m b) -> TCMT m b
forall (m :: * -> *) a. (IORef TCState -> TCEnv -> m a) -> TCMT m a
TCM ((IORef TCState -> TCEnv -> m b) -> TCMT m b)
-> (IORef TCState -> TCEnv -> m b) -> TCMT m b
forall a b. (a -> b) -> a -> b
$ \IORef TCState
r TCEnv
e -> (a -> b) -> m a -> m b
forall a b. (a -> b) -> m a -> m b
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
fmap a -> b
f (IORef TCState -> TCEnv -> m a
m IORef TCState
r TCEnv
e)
{-# INLINE fmapTCMT #-}
instance Applicative m => Applicative (TCMT m) where
pure :: forall a. a -> TCMT m a
pure = a -> TCMT m a
forall (m :: * -> *) a. Applicative m => a -> TCMT m a
returnTCMT; {-# INLINE pure #-}
<*> :: forall a b. TCMT m (a -> b) -> TCMT m a -> TCMT m b
(<*>) = TCMT m (a -> b) -> TCMT m a -> TCMT m b
forall (m :: * -> *) a b.
Applicative m =>
TCMT m (a -> b) -> TCMT m a -> TCMT m b
apTCMT; {-# INLINE (<*>) #-}
apTCMT :: Applicative m => TCMT m (a -> b) -> TCMT m a -> TCMT m b
apTCMT :: forall (m :: * -> *) a b.
Applicative m =>
TCMT m (a -> b) -> TCMT m a -> TCMT m b
apTCMT = \(TCM IORef TCState -> TCEnv -> m (a -> b)
mf) (TCM IORef TCState -> TCEnv -> m a
m) -> (IORef TCState -> TCEnv -> m b) -> TCMT m b
forall (m :: * -> *) a. (IORef TCState -> TCEnv -> m a) -> TCMT m a
TCM ((IORef TCState -> TCEnv -> m b) -> TCMT m b)
-> (IORef TCState -> TCEnv -> m b) -> TCMT m b
forall a b. (a -> b) -> a -> b
$ \IORef TCState
r TCEnv
e -> IORef TCState -> TCEnv -> m (a -> b)
mf IORef TCState
r TCEnv
e m (a -> b) -> m a -> m b
forall a b. m (a -> b) -> m a -> m b
forall (f :: * -> *) a b. Applicative f => f (a -> b) -> f a -> f b
<*> IORef TCState -> TCEnv -> m a
m IORef TCState
r TCEnv
e
{-# INLINE apTCMT #-}
instance MonadTrans TCMT where
lift :: forall (m :: * -> *) a. Monad m => m a -> TCMT m a
lift m a
m = (IORef TCState -> TCEnv -> m a) -> TCMT m a
forall (m :: * -> *) a. (IORef TCState -> TCEnv -> m a) -> TCMT m a
TCM ((IORef TCState -> TCEnv -> m a) -> TCMT m a)
-> (IORef TCState -> TCEnv -> m a) -> TCMT m a
forall a b. (a -> b) -> a -> b
$ \IORef TCState
_ TCEnv
_ -> m a
m; {-# INLINE lift #-}
instance Monad m => Monad (TCMT m) where
return :: forall a. a -> TCMT m a
return = a -> TCMT m a
forall a. a -> TCMT m a
forall (f :: * -> *) a. Applicative f => a -> f a
pure; {-# INLINE return #-}
>>= :: forall a b. TCMT m a -> (a -> TCMT m b) -> TCMT m b
(>>=) = TCMT m a -> (a -> TCMT m b) -> TCMT m b
forall (m :: * -> *) a b.
Monad m =>
TCMT m a -> (a -> TCMT m b) -> TCMT m b
bindTCMT; {-# INLINE (>>=) #-}
>> :: forall a b. TCMT m a -> TCMT m b -> TCMT m b
(>>) = TCMT m a -> TCMT m b -> TCMT m b
forall a b. TCMT m a -> TCMT m b -> TCMT m b
forall (f :: * -> *) a b. Applicative f => f a -> f b -> f b
(*>); {-# INLINE (>>) #-}
instance (CatchIO m, MonadIO m) => MonadFail (TCMT m) where
fail :: forall a. String -> TCMT m a
fail = String -> TCMT m a
forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
String -> m a
internalError
instance MonadIO m => MonadIO (TCMT m) where
liftIO :: forall a. IO a -> TCMT m a
liftIO IO a
m = (IORef TCState -> TCEnv -> m a) -> TCMT m a
forall (m :: * -> *) a. (IORef TCState -> TCEnv -> m a) -> TCMT m a
TCM ((IORef TCState -> TCEnv -> m a) -> TCMT m a)
-> (IORef TCState -> TCEnv -> m a) -> TCMT m a
forall a b. (a -> b) -> a -> b
$ \ IORef TCState
s TCEnv
env -> do
IO a -> m a
forall a. IO a -> m a
forall (m :: * -> *) a. MonadIO m => IO a -> m a
liftIO (IO a -> m a) -> IO a -> m a
forall a b. (a -> b) -> a -> b
$ IORef TCState -> Range -> IO a -> IO a
forall {a}. IORef TCState -> Range -> IO a -> IO a
wrap IORef TCState
s (TCEnv -> Range
envRange TCEnv
env) (IO a -> IO a) -> IO a -> IO a
forall a b. (a -> b) -> a -> b
$ do
x <- IO a
m
x `seq` return x
where
wrap :: IORef TCState -> Range -> IO a -> IO a
wrap IORef TCState
s Range
r IO a
m = IO a -> (IOException -> IO a) -> IO a
forall e a. Exception e => IO a -> (e -> IO a) -> IO a
E.catch IO a
m ((IOException -> IO a) -> IO a) -> (IOException -> IO a) -> IO a
forall a b. (a -> b) -> a -> b
$ \ IOException
err -> do
s <- IORef TCState -> IO TCState
forall a. IORef a -> IO a
readIORef IORef TCState
s
E.throwIO $ IOException (Just s) r err
instance MonadIO m => MonadTCEnv (TCMT m) where
askTC :: TCMT m TCEnv
askTC = (IORef TCState -> TCEnv -> m TCEnv) -> TCMT m TCEnv
forall (m :: * -> *) a. (IORef TCState -> TCEnv -> m a) -> TCMT m a
TCM ((IORef TCState -> TCEnv -> m TCEnv) -> TCMT m TCEnv)
-> (IORef TCState -> TCEnv -> m TCEnv) -> TCMT m TCEnv
forall a b. (a -> b) -> a -> b
$ \ IORef TCState
_ TCEnv
e -> TCEnv -> m TCEnv
forall a. a -> m a
forall (m :: * -> *) a. Monad m => a -> m a
return TCEnv
e; {-# INLINE askTC #-}
localTC :: forall a. (TCEnv -> TCEnv) -> TCMT m a -> TCMT m a
localTC TCEnv -> TCEnv
f (TCM IORef TCState -> TCEnv -> m a
m) = (IORef TCState -> TCEnv -> m a) -> TCMT m a
forall (m :: * -> *) a. (IORef TCState -> TCEnv -> m a) -> TCMT m a
TCM ((IORef TCState -> TCEnv -> m a) -> TCMT m a)
-> (IORef TCState -> TCEnv -> m a) -> TCMT m a
forall a b. (a -> b) -> a -> b
$ \ IORef TCState
s TCEnv
e -> IORef TCState -> TCEnv -> m a
m IORef TCState
s (TCEnv -> TCEnv
f TCEnv
e); {-# INLINE localTC #-}
instance MonadIO m => MonadTCState (TCMT m) where
getTC :: TCMT m TCState
getTC = (IORef TCState -> TCEnv -> m TCState) -> TCMT m TCState
forall (m :: * -> *) a. (IORef TCState -> TCEnv -> m a) -> TCMT m a
TCM ((IORef TCState -> TCEnv -> m TCState) -> TCMT m TCState)
-> (IORef TCState -> TCEnv -> m TCState) -> TCMT m TCState
forall a b. (a -> b) -> a -> b
$ \ IORef TCState
r TCEnv
_e -> IO TCState -> m TCState
forall a. IO a -> m a
forall (m :: * -> *) a. MonadIO m => IO a -> m a
liftIO (IORef TCState -> IO TCState
forall a. IORef a -> IO a
readIORef IORef TCState
r); {-# INLINE getTC #-}
putTC :: TCState -> TCMT m ()
putTC TCState
s = (IORef TCState -> TCEnv -> m ()) -> TCMT m ()
forall (m :: * -> *) a. (IORef TCState -> TCEnv -> m a) -> TCMT m a
TCM ((IORef TCState -> TCEnv -> m ()) -> TCMT m ())
-> (IORef TCState -> TCEnv -> m ()) -> TCMT m ()
forall a b. (a -> b) -> a -> b
$ \ IORef TCState
r TCEnv
_e -> IO () -> m ()
forall a. IO a -> m a
forall (m :: * -> *) a. MonadIO m => IO a -> m a
liftIO (IORef TCState -> TCState -> IO ()
forall a. IORef a -> a -> IO ()
writeIORef IORef TCState
r TCState
s); {-# INLINE putTC #-}
modifyTC :: (TCState -> TCState) -> TCMT m ()
modifyTC TCState -> TCState
f = TCState -> TCMT m ()
forall (m :: * -> *). MonadTCState m => TCState -> m ()
putTC (TCState -> TCMT m ())
-> (TCState -> TCState) -> TCState -> TCMT m ()
forall b c a. (b -> c) -> (a -> b) -> a -> c
. TCState -> TCState
f (TCState -> TCMT m ()) -> TCMT m TCState -> TCMT m ()
forall (m :: * -> *) a b. Monad m => (a -> m b) -> m a -> m b
=<< TCMT m TCState
forall (m :: * -> *). MonadTCState m => m TCState
getTC; {-# INLINE modifyTC #-}
instance MonadIO m => ReadTCState (TCMT m) where
getTCState :: TCMT m TCState
getTCState = TCMT m TCState
forall (m :: * -> *). MonadTCState m => m TCState
getTC; {-# INLINE getTCState #-}
locallyTCState :: forall a b. Lens' TCState a -> (a -> a) -> TCMT m b -> TCMT m b
locallyTCState Lens' TCState a
l a -> a
f = TCMT m a -> (a -> TCMT m ()) -> TCMT m b -> TCMT m b
forall (m :: * -> *) a b.
Monad m =>
m a -> (a -> m ()) -> m b -> m b
bracket_ (Lens' TCState a -> TCMT m a
forall (m :: * -> *) a. ReadTCState m => Lens' TCState a -> m a
useTC (a -> f a) -> TCState -> f TCState
Lens' TCState a
l TCMT m a -> TCMT m () -> TCMT m a
forall a b. TCMT m a -> TCMT m b -> TCMT m a
forall (f :: * -> *) a b. Applicative f => f a -> f b -> f a
<* Lens' TCState a -> (a -> a) -> TCMT m ()
forall (m :: * -> *) a.
MonadTCState m =>
Lens' TCState a -> (a -> a) -> m ()
modifyTCLens (a -> f a) -> TCState -> f TCState
Lens' TCState a
l a -> a
f) (Lens' TCState a -> a -> TCMT m ()
forall (m :: * -> *) a.
MonadTCState m =>
Lens' TCState a -> a -> m ()
setTCLens (a -> f a) -> TCState -> f TCState
Lens' TCState a
l); {-# INLINE locallyTCState #-}
instance MonadBlock TCM where
patternViolation :: forall a. Blocker -> TCM a
patternViolation Blocker
b = TCErr -> TCMT IO a
forall a. TCErr -> TCMT IO a
forall e (m :: * -> *) a. MonadError e m => e -> m a
throwError (Blocker -> TCErr
PatternErr Blocker
b)
catchPatternErr :: forall a. (Blocker -> TCM a) -> TCM a -> TCM a
catchPatternErr Blocker -> TCM a
handle TCM a
v =
TCM a -> (TCErr -> TCM a) -> TCM a
forall a. TCM a -> (TCErr -> TCM a) -> TCM a
catchError_ TCM a
v ((TCErr -> TCM a) -> TCM a) -> (TCErr -> TCM a) -> TCM a
forall a b. (a -> b) -> a -> b
$ \TCErr
err ->
case TCErr
err of
PatternErr Blocker
u -> Blocker -> TCM a
handle Blocker
u
TCErr
_ -> TCErr -> TCM a
forall a. TCErr -> TCMT IO a
forall e (m :: * -> *) a. MonadError e m => e -> m a
throwError TCErr
err
instance (CatchIO m, MonadIO m) => MonadError TCErr (TCMT m) where
throwError :: forall a. TCErr -> TCMT m a
throwError = IO a -> TCMT m a
forall a. IO a -> TCMT m a
forall (m :: * -> *) a. MonadIO m => IO a -> m a
liftIO (IO a -> TCMT m a) -> (TCErr -> IO a) -> TCErr -> TCMT m a
forall b c a. (b -> c) -> (a -> b) -> a -> c
. TCErr -> IO a
forall e a. (HasCallStack, Exception e) => e -> IO a
E.throwIO
catchError :: forall a. TCMT m a -> (TCErr -> TCMT m a) -> TCMT m a
catchError TCMT m a
m TCErr -> TCMT m a
h = (IORef TCState -> TCEnv -> m a) -> TCMT m a
forall (m :: * -> *) a. (IORef TCState -> TCEnv -> m a) -> TCMT m a
TCM ((IORef TCState -> TCEnv -> m a) -> TCMT m a)
-> (IORef TCState -> TCEnv -> m a) -> TCMT m a
forall a b. (a -> b) -> a -> b
$ \ IORef TCState
r TCEnv
e -> do
oldState <- IO TCState -> m TCState
forall a. IO a -> m a
forall (m :: * -> *) a. MonadIO m => IO a -> m a
liftIO (IO TCState -> m TCState) -> IO TCState -> m TCState
forall a b. (a -> b) -> a -> b
$ IORef TCState -> IO TCState
forall a. IORef a -> IO a
readIORef IORef TCState
r
unTCM m r e `catchIO` \TCErr
err -> do
case TCErr
err of
PatternErr{} -> () -> m ()
forall a. a -> m a
forall (m :: * -> *) a. Monad m => a -> m a
return ()
TCErr
_ ->
IO () -> m ()
forall a. IO a -> m a
forall (m :: * -> *) a. MonadIO m => IO a -> m a
liftIO (IO () -> m ()) -> IO () -> m ()
forall a b. (a -> b) -> a -> b
$ do
newState <- IORef TCState -> IO TCState
forall a. IORef a -> IO a
readIORef IORef TCState
r
writeIORef r $ oldState { stPersistentState = stPersistentState newState }
TCMT m a -> IORef TCState -> TCEnv -> m a
forall (m :: * -> *) a. TCMT m a -> IORef TCState -> TCEnv -> m a
unTCM (TCErr -> TCMT m a
h TCErr
err) IORef TCState
r TCEnv
e
instance CatchImpossible TCM where
catchImpossibleJust :: forall b a.
(Impossible -> Maybe b) -> TCM a -> (b -> TCM a) -> TCM a
catchImpossibleJust Impossible -> Maybe b
f TCM a
m b -> TCM a
h = (IORef TCState -> TCEnv -> IO a) -> TCM a
forall (m :: * -> *) a. (IORef TCState -> TCEnv -> m a) -> TCMT m a
TCM ((IORef TCState -> TCEnv -> IO a) -> TCM a)
-> (IORef TCState -> TCEnv -> IO a) -> TCM a
forall a b. (a -> b) -> a -> b
$ \ IORef TCState
r TCEnv
e -> do
s <- IORef TCState -> IO TCState
forall a. IORef a -> IO a
readIORef IORef TCState
r
catchImpossibleJust f (unTCM m r e) $ \ b
err -> do
IORef TCState -> TCState -> IO ()
forall a. IORef a -> a -> IO ()
writeIORef IORef TCState
r TCState
s
TCM a -> IORef TCState -> TCEnv -> IO a
forall (m :: * -> *) a. TCMT m a -> IORef TCState -> TCEnv -> m a
unTCM (b -> TCM a
h b
err) IORef TCState
r TCEnv
e
instance MonadIO m => MonadReduce (TCMT m) where
liftReduce :: forall a. ReduceM a -> TCMT m a
liftReduce = TCM a -> TCMT m a
forall a. TCM a -> TCMT m a
forall (tcm :: * -> *) a. MonadTCM tcm => TCM a -> tcm a
liftTCM (TCM a -> TCMT m a)
-> (ReduceM a -> TCM a) -> ReduceM a -> TCMT m a
forall b c a. (b -> c) -> (a -> b) -> a -> c
. ReduceM a -> TCM a
forall a. ReduceM a -> TCM a
runReduceM; {-# INLINE liftReduce #-}
instance (IsString a, MonadIO m) => IsString (TCMT m a) where
fromString :: String -> TCMT m a
fromString String
s = a -> TCMT m a
forall a. a -> TCMT m a
forall (m :: * -> *) a. Monad m => a -> m a
return (String -> a
forall a. IsString a => String -> a
fromString String
s)
instance {-# OVERLAPPABLE #-} (MonadIO m, Semigroup a) => Semigroup (TCMT m a) where
<> :: TCMT m a -> TCMT m a -> TCMT m a
(<>) = (a -> a -> a) -> TCMT m a -> TCMT m a -> TCMT m a
forall a b c. (a -> b -> c) -> TCMT m a -> TCMT m b -> TCMT m c
forall (f :: * -> *) a b c.
Applicative f =>
(a -> b -> c) -> f a -> f b -> f c
liftA2 a -> a -> a
forall a. Semigroup a => a -> a -> a
(<>)
instance {-# OVERLAPPABLE #-} (MonadIO m, Semigroup a, Monoid a) => Monoid (TCMT m a) where
mempty :: TCMT m a
mempty = a -> TCMT m a
forall a. a -> TCMT m a
forall (f :: * -> *) a. Applicative f => a -> f a
pure a
forall a. Monoid a => a
mempty
mappend :: TCMT m a -> TCMT m a -> TCMT m a
mappend = TCMT m a -> TCMT m a -> TCMT m a
forall a. Semigroup a => a -> a -> a
(<>)
mconcat :: [TCMT m a] -> TCMT m a
mconcat = [a] -> a
forall a. Monoid a => [a] -> a
mconcat ([a] -> a) -> ([TCMT m a] -> TCMT m [a]) -> [TCMT m a] -> TCMT m a
forall (m :: * -> *) b c a.
Functor m =>
(b -> c) -> (a -> m b) -> a -> m c
<.> [TCMT m a] -> TCMT m [a]
forall (t :: * -> *) (m :: * -> *) a.
(Traversable t, Monad m) =>
t (m a) -> m (t a)
forall (m :: * -> *) a. Monad m => [m a] -> m [a]
sequence
instance {-# OVERLAPPABLE #-} (MonadIO m, Null a) => Null (TCMT m a) where
empty :: TCMT m a
empty = a -> TCMT m a
forall a. a -> TCMT m a
forall (m :: * -> *) a. Monad m => a -> m a
return a
forall a. Null a => a
empty
null :: TCMT m a -> Bool
null = TCMT m a -> Bool
forall a. HasCallStack => a
__IMPOSSIBLE__
catchError_ :: TCM a -> (TCErr -> TCM a) -> TCM a
catchError_ :: forall a. TCM a -> (TCErr -> TCM a) -> TCM a
catchError_ TCM a
m TCErr -> TCM a
h = (IORef TCState -> TCEnv -> IO a) -> TCM a
forall (m :: * -> *) a. (IORef TCState -> TCEnv -> m a) -> TCMT m a
TCM ((IORef TCState -> TCEnv -> IO a) -> TCM a)
-> (IORef TCState -> TCEnv -> IO a) -> TCM a
forall a b. (a -> b) -> a -> b
$ \IORef TCState
r TCEnv
e ->
TCM a -> IORef TCState -> TCEnv -> IO a
forall (m :: * -> *) a. TCMT m a -> IORef TCState -> TCEnv -> m a
unTCM TCM a
m IORef TCState
r TCEnv
e
IO a -> (TCErr -> IO a) -> IO a
forall e a. Exception e => IO a -> (e -> IO a) -> IO a
`E.catch` \TCErr
err -> TCM a -> IORef TCState -> TCEnv -> IO a
forall (m :: * -> *) a. TCMT m a -> IORef TCState -> TCEnv -> m a
unTCM (TCErr -> TCM a
h TCErr
err) IORef TCState
r TCEnv
e
finally_ :: TCM a -> TCM b -> TCM a
finally_ :: forall a b. TCM a -> TCM b -> TCM a
finally_ TCM a
m TCM b
f = do
x <- TCM a
m TCM a -> (TCErr -> TCM a) -> TCM a
forall a. TCM a -> (TCErr -> TCM a) -> TCM a
`catchError_` \ TCErr
err -> TCM b
f TCM b -> TCM a -> TCM a
forall a b. TCMT IO a -> TCMT IO b -> TCMT IO b
forall (m :: * -> *) a b. Monad m => m a -> m b -> m b
>> TCErr -> TCM a
forall a. TCErr -> TCMT IO a
forall e (m :: * -> *) a. MonadError e m => e -> m a
throwError TCErr
err
_ <- f
return x
class ( Applicative tcm, MonadIO tcm
, MonadTCEnv tcm
, MonadTCState tcm
, HasOptions tcm
) => MonadTCM tcm where
liftTCM :: TCM a -> tcm a
default liftTCM :: (MonadTCM m, MonadTrans t, tcm ~ t m) => TCM a -> tcm a
liftTCM = m a -> tcm a
m a -> t m a
forall (m :: * -> *) a. Monad m => m a -> t m a
forall (t :: (* -> *) -> * -> *) (m :: * -> *) a.
(MonadTrans t, Monad m) =>
m a -> t m a
lift (m a -> tcm a) -> (TCM a -> m a) -> TCM a -> tcm a
forall b c a. (b -> c) -> (a -> b) -> a -> c
. TCM a -> m a
forall a. TCM a -> m a
forall (tcm :: * -> *) a. MonadTCM tcm => TCM a -> tcm a
liftTCM
{-# INLINE liftTCM #-}
{-# RULES "liftTCM/id" liftTCM = id #-}
instance MonadIO m => MonadTCM (TCMT m) where
liftTCM :: forall a. TCM a -> TCMT m a
liftTCM = (forall a. IO a -> m a) -> TCMT IO a -> TCMT m a
forall (m :: * -> *) (n :: * -> *) a.
(forall a. m a -> n a) -> TCMT m a -> TCMT n a
mapTCMT IO a -> m a
forall a. IO a -> m a
forall (m :: * -> *) a. MonadIO m => IO a -> m a
liftIO
{-# INLINE liftTCM #-}
instance MonadTCM tcm => MonadTCM (ChangeT tcm)
instance MonadTCM tcm => MonadTCM (ExceptT err tcm)
instance MonadTCM tcm => MonadTCM (IdentityT tcm)
instance MonadTCM tcm => MonadTCM (ListT tcm)
instance MonadTCM tcm => MonadTCM (MaybeT tcm)
instance MonadTCM tcm => MonadTCM (ReaderT r tcm)
instance MonadTCM tcm => MonadTCM (StateT s tcm)
instance (Monoid w, MonadTCM tcm) => MonadTCM (WriterT w tcm)
instance MonadBench TCM where
type BenchPhase TCM = Phase
getBenchmark :: TCM (Benchmark (BenchPhase (TCMT IO)))
getBenchmark = IO (Benchmark (BenchPhase (TCMT IO)))
-> TCM (Benchmark (BenchPhase (TCMT IO)))
forall a. IO a -> TCMT IO a
forall (m :: * -> *) a. MonadIO m => IO a -> m a
liftIO (IO (Benchmark (BenchPhase (TCMT IO)))
-> TCM (Benchmark (BenchPhase (TCMT IO))))
-> IO (Benchmark (BenchPhase (TCMT IO)))
-> TCM (Benchmark (BenchPhase (TCMT IO)))
forall a b. (a -> b) -> a -> b
$ IO (Benchmark (BenchPhase IO))
IO (Benchmark (BenchPhase (TCMT IO)))
forall (m :: * -> *). MonadBench m => m (Benchmark (BenchPhase m))
getBenchmark
putBenchmark :: Benchmark (BenchPhase (TCMT IO)) -> TCMT IO ()
putBenchmark = IO () -> TCMT IO ()
forall a. IO a -> TCMT IO a
forall (m :: * -> *) a. MonadIO m => IO a -> m a
liftIO (IO () -> TCMT IO ())
-> (Benchmark -> IO ()) -> Benchmark -> TCMT IO ()
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Benchmark (BenchPhase IO) -> IO ()
Benchmark -> IO ()
forall (m :: * -> *).
MonadBench m =>
Benchmark (BenchPhase m) -> m ()
putBenchmark
finally :: forall a b. TCM a -> TCM b -> TCM a
finally = TCM b -> TCM c -> TCM b
forall a b. TCM a -> TCM b -> TCM a
finally_
instance Null (TCM Doc) where
empty :: TCM Doc
empty = Doc -> TCM Doc
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return Doc
forall a. Null a => a
empty
null :: TCM Doc -> Bool
null = TCM Doc -> Bool
forall a. HasCallStack => a
__IMPOSSIBLE__
internalError :: (HasCallStack, MonadTCError m) => String -> m a
internalError :: forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
String -> m a
internalError String
s = (CallStack -> m a) -> m a
forall b. HasCallStack => (CallStack -> b) -> b
withCallerCallStack ((CallStack -> m a) -> m a) -> (CallStack -> m a) -> m a
forall a b. (a -> b) -> a -> b
$ \ CallStack
loc ->
CallStack -> TypeError -> m a
forall (m :: * -> *) a.
MonadTCError m =>
CallStack -> TypeError -> m a
typeError' CallStack
loc (TypeError -> m a) -> TypeError -> m a
forall a b. (a -> b) -> a -> b
$ String -> TypeError
InternalError String
s
type MonadTCError m = (MonadTCEnv m, ReadTCState m, MonadError TCErr m)
locatedTypeError :: MonadTCError m => (a -> TypeError) -> (HasCallStack => a -> m b)
locatedTypeError :: forall (m :: * -> *) a b.
MonadTCError m =>
(a -> TypeError) -> HasCallStack => a -> m b
locatedTypeError a -> TypeError
f a
e = (CallStack -> m b) -> m b
forall b. HasCallStack => (CallStack -> b) -> b
withCallerCallStack ((CallStack -> TypeError -> m b) -> TypeError -> CallStack -> m b
forall a b c. (a -> b -> c) -> b -> a -> c
flip CallStack -> TypeError -> m b
forall (m :: * -> *) a.
MonadTCError m =>
CallStack -> TypeError -> m a
typeError' (a -> TypeError
f a
e))
genericError :: (HasCallStack, MonadTCError m) => String -> m a
genericError :: forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
String -> m a
genericError = (String -> TypeError) -> HasCallStack => String -> m a
forall (m :: * -> *) a b.
MonadTCError m =>
(a -> TypeError) -> HasCallStack => a -> m b
locatedTypeError String -> TypeError
GenericError
{-# SPECIALIZE genericDocError :: Doc -> TCM a #-}
genericDocError :: (HasCallStack, MonadTCError m) => Doc -> m a
genericDocError :: forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
Doc -> m a
genericDocError = (Doc -> TypeError) -> HasCallStack => Doc -> m a
forall (m :: * -> *) a b.
MonadTCError m =>
(a -> TypeError) -> HasCallStack => a -> m b
locatedTypeError Doc -> TypeError
GenericDocError
{-# SPECIALIZE typeError' :: CallStack -> TypeError -> TCM a #-}
typeError' :: MonadTCError m => CallStack -> TypeError -> m a
typeError' :: forall (m :: * -> *) a.
MonadTCError m =>
CallStack -> TypeError -> m a
typeError' CallStack
loc TypeError
err = TCErr -> m a
forall a. TCErr -> m a
forall e (m :: * -> *) a. MonadError e m => e -> m a
throwError (TCErr -> m a) -> m TCErr -> m a
forall (m :: * -> *) a b. Monad m => (a -> m b) -> m a -> m b
=<< CallStack -> TypeError -> m TCErr
forall (m :: * -> *).
(MonadTCEnv m, ReadTCState m) =>
CallStack -> TypeError -> m TCErr
typeError'_ CallStack
loc TypeError
err
{-# SPECIALIZE typeError :: HasCallStack => TypeError -> TCM a #-}
typeError :: (HasCallStack, MonadTCError m) => TypeError -> m a
typeError :: forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
TypeError -> m a
typeError TypeError
err = (CallStack -> m a) -> m a
forall b. HasCallStack => (CallStack -> b) -> b
withCallerCallStack ((CallStack -> m a) -> m a) -> (CallStack -> m a) -> m a
forall a b. (a -> b) -> a -> b
$ \CallStack
loc -> TCErr -> m a
forall a. TCErr -> m a
forall e (m :: * -> *) a. MonadError e m => e -> m a
throwError (TCErr -> m a) -> m TCErr -> m a
forall (m :: * -> *) a b. Monad m => (a -> m b) -> m a -> m b
=<< CallStack -> TypeError -> m TCErr
forall (m :: * -> *).
(MonadTCEnv m, ReadTCState m) =>
CallStack -> TypeError -> m TCErr
typeError'_ CallStack
loc TypeError
err
{-# SPECIALIZE typeError'_ :: CallStack -> TypeError -> TCM TCErr #-}
typeError'_ :: (MonadTCEnv m, ReadTCState m) => CallStack -> TypeError -> m TCErr
typeError'_ :: forall (m :: * -> *).
(MonadTCEnv m, ReadTCState m) =>
CallStack -> TypeError -> m TCErr
typeError'_ CallStack
loc TypeError
err = CallStack -> TCState -> Closure TypeError -> TCErr
TypeError CallStack
loc (TCState -> Closure TypeError -> TCErr)
-> m TCState -> m (Closure TypeError -> TCErr)
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> m TCState
forall (m :: * -> *). ReadTCState m => m TCState
getTCState m (Closure TypeError -> TCErr) -> m (Closure TypeError) -> m TCErr
forall a b. m (a -> b) -> m a -> m b
forall (f :: * -> *) a b. Applicative f => f (a -> b) -> f a -> f b
<*> TypeError -> m (Closure TypeError)
forall (m :: * -> *) a.
(MonadTCEnv m, ReadTCState m) =>
a -> m (Closure a)
buildClosure TypeError
err
{-# SPECIALIZE typeError_ :: HasCallStack => TypeError -> TCM TCErr #-}
typeError_ :: (HasCallStack, MonadTCEnv m, ReadTCState m) => TypeError -> m TCErr
typeError_ :: forall (m :: * -> *).
(HasCallStack, MonadTCEnv m, ReadTCState m) =>
TypeError -> m TCErr
typeError_ = (CallStack -> m TCErr) -> m TCErr
forall b. HasCallStack => (CallStack -> b) -> b
withCallerCallStack ((CallStack -> m TCErr) -> m TCErr)
-> (TypeError -> CallStack -> m TCErr) -> TypeError -> m TCErr
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (CallStack -> TypeError -> m TCErr)
-> TypeError -> CallStack -> m TCErr
forall a b c. (a -> b -> c) -> b -> a -> c
flip CallStack -> TypeError -> m TCErr
forall (m :: * -> *).
(MonadTCEnv m, ReadTCState m) =>
CallStack -> TypeError -> m TCErr
typeError'_
interactionError :: (HasCallStack, MonadTCError m) => InteractionError -> m a
interactionError :: forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
InteractionError -> m a
interactionError = (InteractionError -> TypeError)
-> HasCallStack => InteractionError -> m a
forall (m :: * -> *) a b.
MonadTCError m =>
(a -> TypeError) -> HasCallStack => a -> m b
locatedTypeError InteractionError -> TypeError
InteractionError
syntaxError :: (HasCallStack, MonadTCError m) => String -> m a
syntaxError :: forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
String -> m a
syntaxError = (String -> TypeError) -> HasCallStack => String -> m a
forall (m :: * -> *) a b.
MonadTCError m =>
(a -> TypeError) -> HasCallStack => a -> m b
locatedTypeError String -> TypeError
SyntaxError
unquoteError :: (HasCallStack, MonadTCError m) => UnquoteError -> m a
unquoteError :: forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
UnquoteError -> m a
unquoteError = (UnquoteError -> TypeError) -> HasCallStack => UnquoteError -> m a
forall (m :: * -> *) a b.
MonadTCError m =>
(a -> TypeError) -> HasCallStack => a -> m b
locatedTypeError UnquoteError -> TypeError
UnquoteFailed
execError :: (HasCallStack, MonadTCError m) => ExecError -> m a
execError :: forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
ExecError -> m a
execError = (ExecError -> TypeError) -> HasCallStack => ExecError -> m a
forall (m :: * -> *) a b.
MonadTCError m =>
(a -> TypeError) -> HasCallStack => a -> m b
locatedTypeError ExecError -> TypeError
ExecError
{-# SPECIALIZE runTCM :: TCEnv -> TCState -> TCM a -> IO (a, TCState) #-}
runTCM :: MonadIO m => TCEnv -> TCState -> TCMT m a -> m (a, TCState)
runTCM :: forall (m :: * -> *) a.
MonadIO m =>
TCEnv -> TCState -> TCMT m a -> m (a, TCState)
runTCM TCEnv
e TCState
s TCMT m a
m = do
r <- IO (IORef TCState) -> m (IORef TCState)
forall a. IO a -> m a
forall (m :: * -> *) a. MonadIO m => IO a -> m a
liftIO (IO (IORef TCState) -> m (IORef TCState))
-> IO (IORef TCState) -> m (IORef TCState)
forall a b. (a -> b) -> a -> b
$ TCState -> IO (IORef TCState)
forall a. a -> IO (IORef a)
newIORef TCState
s
a <- unTCM m r e
s <- liftIO $ readIORef r
return (a, s)
runTCMTop :: TCM a -> IO (Either TCErr a)
runTCMTop :: forall a. TCM a -> IO (Either TCErr a)
runTCMTop TCM a
m = (a -> Either TCErr a
forall a b. b -> Either a b
Right (a -> Either TCErr a) -> IO a -> IO (Either TCErr a)
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> TCM a -> IO a
forall (m :: * -> *) a. MonadIO m => TCMT m a -> m a
runTCMTop' TCM a
m) IO (Either TCErr a)
-> (TCErr -> IO (Either TCErr a)) -> IO (Either TCErr a)
forall e a. Exception e => IO a -> (e -> IO a) -> IO a
`E.catch` (Either TCErr a -> IO (Either TCErr a)
forall a. a -> IO a
forall (m :: * -> *) a. Monad m => a -> m a
return (Either TCErr a -> IO (Either TCErr a))
-> (TCErr -> Either TCErr a) -> TCErr -> IO (Either TCErr a)
forall b c a. (b -> c) -> (a -> b) -> a -> c
. TCErr -> Either TCErr a
forall a b. a -> Either a b
Left)
runTCMTop' :: MonadIO m => TCMT m a -> m a
runTCMTop' :: forall (m :: * -> *) a. MonadIO m => TCMT m a -> m a
runTCMTop' TCMT m a
m = do
r <- IO (IORef TCState) -> m (IORef TCState)
forall a. IO a -> m a
forall (m :: * -> *) a. MonadIO m => IO a -> m a
liftIO (IO (IORef TCState) -> m (IORef TCState))
-> IO (IORef TCState) -> m (IORef TCState)
forall a b. (a -> b) -> a -> b
$ TCState -> IO (IORef TCState)
forall a. a -> IO (IORef a)
newIORef TCState
initState
unTCM m r initEnv
runSafeTCM :: TCM a -> TCState -> IO (a, TCState)
runSafeTCM :: forall a. TCM a -> TCState -> IO (a, TCState)
runSafeTCM TCM a
m TCState
st =
TCEnv -> TCState -> TCM a -> IO (a, TCState)
forall (m :: * -> *) a.
MonadIO m =>
TCEnv -> TCState -> TCMT m a -> m (a, TCState)
runTCM TCEnv
initEnv TCState
st TCM a
m IO (a, TCState) -> (TCErr -> IO (a, TCState)) -> IO (a, TCState)
forall e a. Exception e => IO a -> (e -> IO a) -> IO a
`E.catch` \(TCErr
e :: TCErr) -> case TCErr
e of
IOException Maybe TCState
_ Range
_ IOException
err -> IOException -> IO (a, TCState)
forall e a. (HasCallStack, Exception e) => e -> IO a
E.throwIO IOException
err
TCErr
_ -> IO (a, TCState)
forall a. HasCallStack => a
__IMPOSSIBLE__
forkTCM :: TCM () -> TCM ()
forkTCM :: TCMT IO () -> TCMT IO ()
forkTCM TCMT IO ()
m = do
s <- TCMT IO TCState
forall (m :: * -> *). MonadTCState m => m TCState
getTC
e <- askTC
liftIO $ void $ forkIO $ void $ runTCM e s m
type InteractionOutputCallback = Response_boot TCErr TCWarning WarningsAndNonFatalErrors -> TCM ()
defaultInteractionOutputCallback :: InteractionOutputCallback
defaultInteractionOutputCallback :: InteractionOutputCallback
defaultInteractionOutputCallback = \case
Resp_HighlightingInfo {} -> TCMT IO ()
forall a. HasCallStack => a
__IMPOSSIBLE__
Resp_Status {} -> TCMT IO ()
forall a. HasCallStack => a
__IMPOSSIBLE__
Resp_JumpToError {} -> TCMT IO ()
forall a. HasCallStack => a
__IMPOSSIBLE__
Resp_InteractionPoints {} -> TCMT IO ()
forall a. HasCallStack => a
__IMPOSSIBLE__
Resp_GiveAction {} -> TCMT IO ()
forall a. HasCallStack => a
__IMPOSSIBLE__
Resp_MakeCase {} -> TCMT IO ()
forall a. HasCallStack => a
__IMPOSSIBLE__
Resp_SolveAll {} -> TCMT IO ()
forall a. HasCallStack => a
__IMPOSSIBLE__
Resp_Mimer {} -> TCMT IO ()
forall a. HasCallStack => a
__IMPOSSIBLE__
Resp_DisplayInfo {} -> TCMT IO ()
forall a. HasCallStack => a
__IMPOSSIBLE__
Resp_RunningInfo Int
_ String
s -> IO () -> TCMT IO ()
forall a. IO a -> TCMT IO a
forall (m :: * -> *) a. MonadIO m => IO a -> m a
liftIO (IO () -> TCMT IO ()) -> IO () -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ do
String -> IO ()
putStr String
s
Handle -> IO ()
hFlush Handle
stdout
Resp_ClearRunningInfo {} -> TCMT IO ()
forall a. HasCallStack => a
__IMPOSSIBLE__
Resp_ClearHighlighting {} -> TCMT IO ()
forall a. HasCallStack => a
__IMPOSSIBLE__
Resp_DoneAborting {} -> TCMT IO ()
forall a. HasCallStack => a
__IMPOSSIBLE__
Resp_DoneExiting {} -> TCMT IO ()
forall a. HasCallStack => a
__IMPOSSIBLE__
patternInTeleName :: String
patternInTeleName :: String
patternInTeleName = String
".patternInTele"
extendedLambdaName :: String
extendedLambdaName :: String
extendedLambdaName = String
".extendedlambda"
isExtendedLambdaName :: A.QName -> Bool
isExtendedLambdaName :: QName -> Bool
isExtendedLambdaName = (String
extendedLambdaName String -> String -> Bool
forall a. Eq a => [a] -> [a] -> Bool
`List.isPrefixOf`) (String -> Bool) -> (QName -> String) -> QName -> Bool
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Name -> String
forall a. Pretty a => a -> String
prettyShow (Name -> String) -> (QName -> Name) -> QName -> String
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Name -> Name
nameConcrete (Name -> Name) -> (QName -> Name) -> QName -> Name
forall b c a. (b -> c) -> (a -> b) -> a -> c
. QName -> Name
qnameName
absurdLambdaName :: String
absurdLambdaName :: String
absurdLambdaName = String
".absurdlambda"
isAbsurdLambdaName :: QName -> Bool
isAbsurdLambdaName :: QName -> Bool
isAbsurdLambdaName = (String
absurdLambdaName String -> String -> Bool
forall a. Eq a => a -> a -> Bool
==) (String -> Bool) -> (QName -> String) -> QName -> Bool
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Name -> String
forall a. Pretty a => a -> String
prettyShow (Name -> String) -> (QName -> Name) -> QName -> String
forall b c a. (b -> c) -> (a -> b) -> a -> c
. QName -> Name
qnameName
generalizedFieldName :: String
generalizedFieldName :: String
generalizedFieldName = String
".generalizedField-"
getGeneralizedFieldName :: A.QName -> Maybe String
getGeneralizedFieldName :: QName -> Maybe String
getGeneralizedFieldName = String -> String -> Maybe String
forall a. Eq a => [a] -> [a] -> Maybe [a]
List.stripPrefix String
generalizedFieldName (String -> Maybe String)
-> (QName -> String) -> QName -> Maybe String
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Name -> String
forall a. Pretty a => a -> String
prettyShow (Name -> String) -> (QName -> Name) -> QName -> String
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Name -> Name
nameConcrete (Name -> Name) -> (QName -> Name) -> QName -> Name
forall b c a. (b -> c) -> (a -> b) -> a -> c
. QName -> Name
qnameName
instance KillRange Signature where
killRange :: KillRangeT Signature
killRange (Sig Sections
secs Definitions
defs RewriteRuleMap
rews InstanceTable
inst) = (Sections
-> Definitions -> RewriteRuleMap -> InstanceTable -> Signature)
-> Sections
-> Definitions
-> RewriteRuleMap
-> InstanceTable
-> Signature
forall t (b :: Bool).
(KILLRANGE t b, IsBase t ~ b, All KillRange (Domains t)) =>
t -> t
killRangeN Sections
-> Definitions -> RewriteRuleMap -> InstanceTable -> Signature
Sig Sections
secs Definitions
defs RewriteRuleMap
rews InstanceTable
inst
instance KillRange InstanceTable where
killRange :: InstanceTable -> InstanceTable
killRange (InstanceTable DiscrimTree QName
tree Map QName Int
count) = (DiscrimTree QName -> Map QName Int -> InstanceTable)
-> DiscrimTree QName -> Map QName Int -> InstanceTable
forall t (b :: Bool).
(KILLRANGE t b, IsBase t ~ b, All KillRange (Domains t)) =>
t -> t
killRangeN DiscrimTree QName -> Map QName Int -> InstanceTable
InstanceTable DiscrimTree QName
tree Map QName Int
count
instance KillRange Sections where
killRange :: KillRangeT Sections
killRange = (Section -> Section) -> KillRangeT Sections
forall a b. (a -> b) -> Map ModuleName a -> Map ModuleName b
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
fmap Section -> Section
forall a. KillRange a => KillRangeT a
killRange
instance KillRange Definitions where
killRange :: KillRangeT Definitions
killRange = (Definition -> Definition) -> KillRangeT Definitions
forall a b. (a -> b) -> HashMap QName a -> HashMap QName b
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
fmap Definition -> Definition
forall a. KillRange a => KillRangeT a
killRange
instance KillRange RewriteRuleMap where
killRange :: KillRangeT RewriteRuleMap
killRange = ([RewriteRule] -> [RewriteRule]) -> KillRangeT RewriteRuleMap
forall a b. (a -> b) -> HashMap QName a -> HashMap QName b
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
fmap [RewriteRule] -> [RewriteRule]
forall a. KillRange a => KillRangeT a
killRange
instance KillRange Section where
killRange :: Section -> Section
killRange (Section Telescope
tel) = (Telescope -> Section) -> Telescope -> Section
forall t (b :: Bool).
(KILLRANGE t b, IsBase t ~ b, All KillRange (Domains t)) =>
t -> t
killRangeN Telescope -> Section
Section Telescope
tel
instance KillRange InstanceInfo where
killRange :: KillRangeT InstanceInfo
killRange :: KillRangeT InstanceInfo
killRange (InstanceInfo QName
a OverlapMode
b) = (QName -> OverlapMode -> InstanceInfo)
-> QName -> OverlapMode -> InstanceInfo
forall t (b :: Bool).
(KILLRANGE t b, IsBase t ~ b, All KillRange (Domains t)) =>
t -> t
killRangeN QName -> OverlapMode -> InstanceInfo
InstanceInfo QName
a OverlapMode
b
instance KillRange Definition where
killRange :: Definition -> Definition
killRange (Defn ArgInfo
ai QName
name Type
t [Polarity]
pols [Occurrence]
occs [Maybe Name]
gpars [LocalDisplayForm]
displ MutualId
mut CompiledRepresentation
compiled Maybe InstanceInfo
inst Bool
copy Set QName
ma Bool
nc Bool
inj Bool
copat Blocked_
blk Language
lang Defn
def) =
(ArgInfo
-> QName
-> Type
-> [Polarity]
-> [Occurrence]
-> [Maybe Name]
-> [LocalDisplayForm]
-> MutualId
-> CompiledRepresentation
-> Maybe InstanceInfo
-> Bool
-> Set QName
-> Bool
-> Bool
-> Bool
-> Blocked_
-> Language
-> Defn
-> Definition)
-> ArgInfo
-> QName
-> Type
-> [Polarity]
-> [Occurrence]
-> [Maybe Name]
-> [LocalDisplayForm]
-> MutualId
-> CompiledRepresentation
-> Maybe InstanceInfo
-> Bool
-> Set QName
-> Bool
-> Bool
-> Bool
-> Blocked_
-> Language
-> Defn
-> Definition
forall t (b :: Bool).
(KILLRANGE t b, IsBase t ~ b, All KillRange (Domains t)) =>
t -> t
killRangeN ArgInfo
-> QName
-> Type
-> [Polarity]
-> [Occurrence]
-> [Maybe Name]
-> [LocalDisplayForm]
-> MutualId
-> CompiledRepresentation
-> Maybe InstanceInfo
-> Bool
-> Set QName
-> Bool
-> Bool
-> Bool
-> Blocked_
-> Language
-> Defn
-> Definition
Defn ArgInfo
ai QName
name Type
t [Polarity]
pols [Occurrence]
occs [Maybe Name]
gpars [LocalDisplayForm]
displ MutualId
mut CompiledRepresentation
compiled Maybe InstanceInfo
inst Bool
copy Set QName
ma Bool
nc Bool
inj Bool
copat Blocked_
blk Language
lang Defn
def
instance KillRange NumGeneralizableArgs where
killRange :: KillRangeT NumGeneralizableArgs
killRange = KillRangeT NumGeneralizableArgs
forall a. a -> a
id
instance KillRange NLPat where
killRange :: KillRangeT NLPat
killRange (PVar Int
x [Arg Int]
y) = (Int -> [Arg Int] -> NLPat) -> Int -> [Arg Int] -> NLPat
forall t (b :: Bool).
(KILLRANGE t b, IsBase t ~ b, All KillRange (Domains t)) =>
t -> t
killRangeN Int -> [Arg Int] -> NLPat
PVar Int
x [Arg Int]
y
killRange (PDef QName
x PElims
y) = (QName -> PElims -> NLPat) -> QName -> PElims -> NLPat
forall t (b :: Bool).
(KILLRANGE t b, IsBase t ~ b, All KillRange (Domains t)) =>
t -> t
killRangeN QName -> PElims -> NLPat
PDef QName
x PElims
y
killRange (PLam ArgInfo
x Abs NLPat
y) = (ArgInfo -> Abs NLPat -> NLPat) -> ArgInfo -> Abs NLPat -> NLPat
forall t (b :: Bool).
(KILLRANGE t b, IsBase t ~ b, All KillRange (Domains t)) =>
t -> t
killRangeN ArgInfo -> Abs NLPat -> NLPat
PLam ArgInfo
x Abs NLPat
y
killRange (PPi Dom NLPType
x Abs NLPType
y) = (Dom NLPType -> Abs NLPType -> NLPat)
-> Dom NLPType -> Abs NLPType -> NLPat
forall t (b :: Bool).
(KILLRANGE t b, IsBase t ~ b, All KillRange (Domains t)) =>
t -> t
killRangeN Dom NLPType -> Abs NLPType -> NLPat
PPi Dom NLPType
x Abs NLPType
y
killRange (PSort NLPSort
x) = (NLPSort -> NLPat) -> NLPSort -> NLPat
forall t (b :: Bool).
(KILLRANGE t b, IsBase t ~ b, All KillRange (Domains t)) =>
t -> t
killRangeN NLPSort -> NLPat
PSort NLPSort
x
killRange (PBoundVar Int
x PElims
y) = (Int -> PElims -> NLPat) -> Int -> PElims -> NLPat
forall t (b :: Bool).
(KILLRANGE t b, IsBase t ~ b, All KillRange (Domains t)) =>
t -> t
killRangeN Int -> PElims -> NLPat
PBoundVar Int
x PElims
y
killRange (PTerm Term
x) = (Term -> NLPat) -> Term -> NLPat
forall t (b :: Bool).
(KILLRANGE t b, IsBase t ~ b, All KillRange (Domains t)) =>
t -> t
killRangeN Term -> NLPat
PTerm Term
x
instance KillRange NLPType where
killRange :: KillRangeT NLPType
killRange (NLPType NLPSort
s NLPat
a) = (NLPSort -> NLPat -> NLPType) -> NLPSort -> NLPat -> NLPType
forall t (b :: Bool).
(KILLRANGE t b, IsBase t ~ b, All KillRange (Domains t)) =>
t -> t
killRangeN NLPSort -> NLPat -> NLPType
NLPType NLPSort
s NLPat
a
instance KillRange NLPSort where
killRange :: KillRangeT NLPSort
killRange (PUniv Univ
u NLPat
l) = (NLPat -> NLPSort) -> NLPat -> NLPSort
forall t (b :: Bool).
(KILLRANGE t b, IsBase t ~ b, All KillRange (Domains t)) =>
t -> t
killRangeN (Univ -> NLPat -> NLPSort
PUniv Univ
u) NLPat
l
killRange s :: NLPSort
s@(PInf Univ
f Integer
n) = NLPSort
s
killRange NLPSort
PSizeUniv = NLPSort
PSizeUniv
killRange NLPSort
PLockUniv = NLPSort
PLockUniv
killRange NLPSort
PLevelUniv = NLPSort
PLevelUniv
killRange NLPSort
PIntervalUniv = NLPSort
PIntervalUniv
instance KillRange RewriteRule where
killRange :: KillRangeT RewriteRule
killRange (RewriteRule QName
q Telescope
gamma QName
f PElims
es Term
rhs Type
t Bool
c) =
(QName
-> Telescope
-> QName
-> PElims
-> Term
-> Type
-> Bool
-> RewriteRule)
-> QName
-> Telescope
-> QName
-> PElims
-> Term
-> Type
-> Bool
-> RewriteRule
forall t (b :: Bool).
(KILLRANGE t b, IsBase t ~ b, All KillRange (Domains t)) =>
t -> t
killRangeN QName
-> Telescope
-> QName
-> PElims
-> Term
-> Type
-> Bool
-> RewriteRule
RewriteRule QName
q Telescope
gamma QName
f PElims
es Term
rhs Type
t Bool
c
instance KillRange CompiledRepresentation where
killRange :: KillRangeT CompiledRepresentation
killRange = KillRangeT CompiledRepresentation
forall a. a -> a
id
instance KillRange EtaEquality where
killRange :: KillRangeT EtaEquality
killRange = KillRangeT EtaEquality
forall a. a -> a
id
instance KillRange System where
killRange :: System -> System
killRange (System Telescope
tel [(Face, Term)]
sys) = Telescope -> [(Face, Term)] -> System
System (KillRangeT Telescope
forall a. KillRange a => KillRangeT a
killRange Telescope
tel) (KillRangeT [(Face, Term)]
forall a. KillRange a => KillRangeT a
killRange [(Face, Term)]
sys)
instance KillRange ExtLamInfo where
killRange :: ExtLamInfo -> ExtLamInfo
killRange (ExtLamInfo ModuleName
m Bool
b Maybe System
sys) = (ModuleName -> Bool -> Maybe System -> ExtLamInfo)
-> ModuleName -> Bool -> Maybe System -> ExtLamInfo
forall t (b :: Bool).
(KILLRANGE t b, IsBase t ~ b, All KillRange (Domains t)) =>
t -> t
killRangeN ModuleName -> Bool -> Maybe System -> ExtLamInfo
ExtLamInfo ModuleName
m Bool
b Maybe System
sys
instance KillRange FunctionFlag where
killRange :: FunctionFlag -> FunctionFlag
killRange = FunctionFlag -> FunctionFlag
forall a. a -> a
id
instance KillRange CompKit where
killRange :: CompKit -> CompKit
killRange = CompKit -> CompKit
forall a. a -> a
id
instance KillRange ProjectionLikenessMissing where
killRange :: ProjectionLikenessMissing -> ProjectionLikenessMissing
killRange = ProjectionLikenessMissing -> ProjectionLikenessMissing
forall a. a -> a
id
instance KillRange BuiltinSort where
killRange :: KillRangeT BuiltinSort
killRange = KillRangeT BuiltinSort
forall a. a -> a
id
instance KillRange Defn where
killRange :: KillRangeT Defn
killRange Defn
def =
case Defn
def of
Axiom Bool
a -> Bool -> Defn
Axiom Bool
a
DataOrRecSig Int
n -> Int -> Defn
DataOrRecSig Int
n
GeneralizableVar NumGeneralizableArgs
a -> NumGeneralizableArgs -> Defn
GeneralizableVar NumGeneralizableArgs
a
AbstractDefn{} -> Defn
forall a. HasCallStack => a
__IMPOSSIBLE__
Function [Clause]
a Maybe CompiledClauses
b Maybe SplitTree
c Maybe Compiled
d [Clause]
e FunctionInverse
f Maybe [QName]
g Either ProjectionLikenessMissing Projection
h SmallSet FunctionFlag
i Maybe Bool
j Maybe ExtLamInfo
k Maybe QName
l Maybe QName
m IsOpaque
n ->
([Clause]
-> Maybe CompiledClauses
-> Maybe SplitTree
-> Maybe Compiled
-> [Clause]
-> FunctionInverse
-> Maybe [QName]
-> Either ProjectionLikenessMissing Projection
-> SmallSet FunctionFlag
-> Maybe Bool
-> Maybe ExtLamInfo
-> Maybe QName
-> Maybe QName
-> IsOpaque
-> Defn)
-> [Clause]
-> Maybe CompiledClauses
-> Maybe SplitTree
-> Maybe Compiled
-> [Clause]
-> FunctionInverse
-> Maybe [QName]
-> Either ProjectionLikenessMissing Projection
-> SmallSet FunctionFlag
-> Maybe Bool
-> Maybe ExtLamInfo
-> Maybe QName
-> Maybe QName
-> IsOpaque
-> Defn
forall t (b :: Bool).
(KILLRANGE t b, IsBase t ~ b, All KillRange (Domains t)) =>
t -> t
killRangeN [Clause]
-> Maybe CompiledClauses
-> Maybe SplitTree
-> Maybe Compiled
-> [Clause]
-> FunctionInverse
-> Maybe [QName]
-> Either ProjectionLikenessMissing Projection
-> SmallSet FunctionFlag
-> Maybe Bool
-> Maybe ExtLamInfo
-> Maybe QName
-> Maybe QName
-> IsOpaque
-> Defn
Function [Clause]
a Maybe CompiledClauses
b Maybe SplitTree
c Maybe Compiled
d [Clause]
e FunctionInverse
f Maybe [QName]
g Either ProjectionLikenessMissing Projection
h SmallSet FunctionFlag
i Maybe Bool
j Maybe ExtLamInfo
k Maybe QName
l Maybe QName
m IsOpaque
n
Datatype Int
a Int
b Maybe Clause
c [QName]
d Sort
e Maybe [QName]
f IsAbstract
g [QName]
h Maybe QName
i Maybe QName
j -> (Int
-> Int
-> Maybe Clause
-> [QName]
-> Sort
-> Maybe [QName]
-> IsAbstract
-> [QName]
-> Maybe QName
-> Maybe QName
-> Defn)
-> Int
-> Int
-> Maybe Clause
-> [QName]
-> Sort
-> Maybe [QName]
-> IsAbstract
-> [QName]
-> Maybe QName
-> Maybe QName
-> Defn
forall t (b :: Bool).
(KILLRANGE t b, IsBase t ~ b, All KillRange (Domains t)) =>
t -> t
killRangeN Int
-> Int
-> Maybe Clause
-> [QName]
-> Sort
-> Maybe [QName]
-> IsAbstract
-> [QName]
-> Maybe QName
-> Maybe QName
-> Defn
Datatype Int
a Int
b Maybe Clause
c [QName]
d Sort
e Maybe [QName]
f IsAbstract
g [QName]
h Maybe QName
i Maybe QName
j
Record Int
a Maybe Clause
b ConHead
c Bool
d [Dom QName]
e Telescope
f Maybe [QName]
g EtaEquality
h PatternOrCopattern
i Maybe Induction
j Maybe Bool
k IsAbstract
l CompKit
m -> (Int
-> Maybe Clause
-> ConHead
-> Bool
-> [Dom QName]
-> Telescope
-> Maybe [QName]
-> EtaEquality
-> PatternOrCopattern
-> Maybe Induction
-> Maybe Bool
-> IsAbstract
-> CompKit
-> Defn)
-> Int
-> Maybe Clause
-> ConHead
-> Bool
-> [Dom QName]
-> Telescope
-> Maybe [QName]
-> EtaEquality
-> PatternOrCopattern
-> Maybe Induction
-> Maybe Bool
-> IsAbstract
-> CompKit
-> Defn
forall t (b :: Bool).
(KILLRANGE t b, IsBase t ~ b, All KillRange (Domains t)) =>
t -> t
killRangeN Int
-> Maybe Clause
-> ConHead
-> Bool
-> [Dom QName]
-> Telescope
-> Maybe [QName]
-> EtaEquality
-> PatternOrCopattern
-> Maybe Induction
-> Maybe Bool
-> IsAbstract
-> CompKit
-> Defn
Record Int
a Maybe Clause
b ConHead
c Bool
d [Dom QName]
e Telescope
f Maybe [QName]
g EtaEquality
h PatternOrCopattern
i Maybe Induction
j Maybe Bool
k IsAbstract
l CompKit
m
Constructor Int
a Int
b ConHead
c QName
d IsAbstract
e CompKit
f Maybe [QName]
g [IsForced]
h Maybe [Bool]
i Bool
j Bool
k -> (Int
-> Int
-> ConHead
-> QName
-> IsAbstract
-> CompKit
-> Maybe [QName]
-> [IsForced]
-> Maybe [Bool]
-> Bool
-> Bool
-> Defn)
-> Int
-> Int
-> ConHead
-> QName
-> IsAbstract
-> CompKit
-> Maybe [QName]
-> [IsForced]
-> Maybe [Bool]
-> Bool
-> Bool
-> Defn
forall t (b :: Bool).
(KILLRANGE t b, IsBase t ~ b, All KillRange (Domains t)) =>
t -> t
killRangeN Int
-> Int
-> ConHead
-> QName
-> IsAbstract
-> CompKit
-> Maybe [QName]
-> [IsForced]
-> Maybe [Bool]
-> Bool
-> Bool
-> Defn
Constructor Int
a Int
b ConHead
c QName
d IsAbstract
e CompKit
f Maybe [QName]
g [IsForced]
h Maybe [Bool]
i Bool
j Bool
k
Primitive IsAbstract
a PrimitiveId
b [Clause]
c FunctionInverse
d Maybe CompiledClauses
e IsOpaque
f -> (IsAbstract
-> PrimitiveId
-> [Clause]
-> FunctionInverse
-> Maybe CompiledClauses
-> IsOpaque
-> Defn)
-> IsAbstract
-> PrimitiveId
-> [Clause]
-> FunctionInverse
-> Maybe CompiledClauses
-> IsOpaque
-> Defn
forall t (b :: Bool).
(KILLRANGE t b, IsBase t ~ b, All KillRange (Domains t)) =>
t -> t
killRangeN IsAbstract
-> PrimitiveId
-> [Clause]
-> FunctionInverse
-> Maybe CompiledClauses
-> IsOpaque
-> Defn
Primitive IsAbstract
a PrimitiveId
b [Clause]
c FunctionInverse
d Maybe CompiledClauses
e IsOpaque
f
PrimitiveSort BuiltinSort
a Sort
b -> (BuiltinSort -> Sort -> Defn) -> BuiltinSort -> Sort -> Defn
forall t (b :: Bool).
(KILLRANGE t b, IsBase t ~ b, All KillRange (Domains t)) =>
t -> t
killRangeN BuiltinSort -> Sort -> Defn
PrimitiveSort BuiltinSort
a Sort
b
instance KillRange MutualId where
killRange :: MutualId -> MutualId
killRange = MutualId -> MutualId
forall a. a -> a
id
instance KillRange c => KillRange (FunctionInverse' c) where
killRange :: KillRangeT (FunctionInverse' c)
killRange FunctionInverse' c
NotInjective = FunctionInverse' c
forall c. FunctionInverse' c
NotInjective
killRange (Inverse InversionMap c
m) = InversionMap c -> FunctionInverse' c
forall c. InversionMap c -> FunctionInverse' c
Inverse (InversionMap c -> FunctionInverse' c)
-> InversionMap c -> FunctionInverse' c
forall a b. (a -> b) -> a -> b
$ KillRangeT (InversionMap c)
forall k v. (KillRange k, KillRange v) => KillRangeT (Map k v)
killRangeMap InversionMap c
m
instance KillRange TermHead where
killRange :: TermHead -> TermHead
killRange TermHead
SortHead = TermHead
SortHead
killRange TermHead
PiHead = TermHead
PiHead
killRange (ConsHead QName
q) = QName -> TermHead
ConsHead (QName -> TermHead) -> QName -> TermHead
forall a b. (a -> b) -> a -> b
$ KillRangeT QName
forall a. KillRange a => KillRangeT a
killRange QName
q
killRange h :: TermHead
h@VarHead{} = TermHead
h
killRange TermHead
UnknownHead = TermHead
UnknownHead
instance KillRange Projection where
killRange :: KillRangeT Projection
killRange (Projection Maybe QName
a QName
b Arg QName
c Int
d ProjLams
e) = (Maybe QName
-> QName -> Arg QName -> Int -> ProjLams -> Projection)
-> Maybe QName
-> QName
-> Arg QName
-> Int
-> ProjLams
-> Projection
forall t (b :: Bool).
(KILLRANGE t b, IsBase t ~ b, All KillRange (Domains t)) =>
t -> t
killRangeN Maybe QName -> QName -> Arg QName -> Int -> ProjLams -> Projection
Projection Maybe QName
a QName
b Arg QName
c Int
d ProjLams
e
instance KillRange ProjLams where
killRange :: KillRangeT ProjLams
killRange = KillRangeT ProjLams
forall a. a -> a
id
instance KillRange a => KillRange (Open a) where
killRange :: KillRangeT (Open a)
killRange = (a -> a) -> KillRangeT (Open a)
forall a b. (a -> b) -> Open a -> Open b
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
fmap a -> a
forall a. KillRange a => KillRangeT a
killRange
instance KillRange DisplayForm where
killRange :: KillRangeT DisplayForm
killRange (Display Int
n [Elim]
es DisplayTerm
dt) = (Int -> [Elim] -> DisplayTerm -> DisplayForm)
-> Int -> [Elim] -> DisplayTerm -> DisplayForm
forall t (b :: Bool).
(KILLRANGE t b, IsBase t ~ b, All KillRange (Domains t)) =>
t -> t
killRangeN Int -> [Elim] -> DisplayTerm -> DisplayForm
Display Int
n [Elim]
es DisplayTerm
dt
instance KillRange Polarity where
killRange :: KillRangeT Polarity
killRange = KillRangeT Polarity
forall a. a -> a
id
instance KillRange IsForced where
killRange :: KillRangeT IsForced
killRange = KillRangeT IsForced
forall a. a -> a
id
instance KillRange DoGeneralize where
killRange :: DoGeneralize -> DoGeneralize
killRange = DoGeneralize -> DoGeneralize
forall a. a -> a
id
instance KillRange DisplayTerm where
killRange :: KillRangeT DisplayTerm
killRange DisplayTerm
dt =
case DisplayTerm
dt of
DWithApp DisplayTerm
dt List1 DisplayTerm
dts [Elim]
es -> (DisplayTerm -> List1 DisplayTerm -> [Elim] -> DisplayTerm)
-> DisplayTerm -> List1 DisplayTerm -> [Elim] -> DisplayTerm
forall t (b :: Bool).
(KILLRANGE t b, IsBase t ~ b, All KillRange (Domains t)) =>
t -> t
killRangeN DisplayTerm -> List1 DisplayTerm -> [Elim] -> DisplayTerm
DWithApp DisplayTerm
dt List1 DisplayTerm
dts [Elim]
es
DCon ConHead
q ConInfo
ci [Arg DisplayTerm]
dts -> (ConHead -> ConInfo -> [Arg DisplayTerm] -> DisplayTerm)
-> ConHead -> ConInfo -> [Arg DisplayTerm] -> DisplayTerm
forall t (b :: Bool).
(KILLRANGE t b, IsBase t ~ b, All KillRange (Domains t)) =>
t -> t
killRangeN ConHead -> ConInfo -> [Arg DisplayTerm] -> DisplayTerm
DCon ConHead
q ConInfo
ci [Arg DisplayTerm]
dts
DDef QName
q [Elim' DisplayTerm]
dts -> (QName -> [Elim' DisplayTerm] -> DisplayTerm)
-> QName -> [Elim' DisplayTerm] -> DisplayTerm
forall t (b :: Bool).
(KILLRANGE t b, IsBase t ~ b, All KillRange (Domains t)) =>
t -> t
killRangeN QName -> [Elim' DisplayTerm] -> DisplayTerm
DDef QName
q [Elim' DisplayTerm]
dts
DDot' Term
v [Elim]
es -> (Term -> [Elim] -> DisplayTerm) -> Term -> [Elim] -> DisplayTerm
forall t (b :: Bool).
(KILLRANGE t b, IsBase t ~ b, All KillRange (Domains t)) =>
t -> t
killRangeN Term -> [Elim] -> DisplayTerm
DDot' Term
v [Elim]
es
DTerm' Term
v [Elim]
es -> (Term -> [Elim] -> DisplayTerm) -> Term -> [Elim] -> DisplayTerm
forall t (b :: Bool).
(KILLRANGE t b, IsBase t ~ b, All KillRange (Domains t)) =>
t -> t
killRangeN Term -> [Elim] -> DisplayTerm
DTerm' Term
v [Elim]
es
instance KillRange a => KillRange (Closure a) where
killRange :: KillRangeT (Closure a)
killRange = KillRangeT (Closure a)
forall a. a -> a
id
instance NFData NumGeneralizableArgs where
rnf :: NumGeneralizableArgs -> ()
rnf NumGeneralizableArgs
NoGeneralizableArgs = ()
rnf (SomeGeneralizableArgs Int
_) = ()
instance NFData TCErr where
rnf :: TCErr -> ()
rnf (TypeError CallStack
a TCState
b Closure TypeError
c) = CallStack -> ()
forall a. NFData a => a -> ()
rnf CallStack
a () -> () -> ()
forall a b. a -> b -> b
`seq` TCState -> ()
forall a. NFData a => a -> ()
rnf TCState
b () -> () -> ()
forall a b. a -> b -> b
`seq` Closure TypeError -> ()
forall a. NFData a => a -> ()
rnf Closure TypeError
c
rnf (ParserError ParseError
a) = ParseError -> ()
forall a. NFData a => a -> ()
rnf ParseError
a
rnf (GenericException String
a)= String -> ()
forall a. NFData a => a -> ()
rnf String
a
rnf (IOException Maybe TCState
a Range
b IOException
c) = Maybe TCState -> ()
forall a. NFData a => a -> ()
rnf Maybe TCState
a () -> () -> ()
forall a b. a -> b -> b
`seq` Range -> ()
forall a. NFData a => a -> ()
rnf Range
b () -> () -> ()
forall a b. a -> b -> b
`seq` Bool -> ()
forall a. NFData a => a -> ()
rnf (IOException
c IOException -> IOException -> Bool
forall a. Eq a => a -> a -> Bool
== IOException
c)
rnf (PatternErr Blocker
a) = Blocker -> ()
forall a. NFData a => a -> ()
rnf Blocker
a
instance NFData PreScopeState
instance NFData PostScopeState
instance NFData TCState
instance NFData DisambiguatedName
instance NFData MutualBlock
instance NFData OpaqueBlock
instance NFData (BiMap RawTopLevelModuleName ModuleNameHash)
instance NFData PersistentTCState
instance NFData LoadedFileCache
instance NFData TypeCheckAction
instance NFData ModuleCheckMode
instance NFData ModuleInfo
instance NFData ForeignCode
instance NFData Interface
instance NFData a => NFData (Closure a)
instance NFData ProblemConstraint
instance NFData WhyCheckModality
instance NFData Constraint
instance NFData Signature
instance NFData InstanceTable
instance NFData Comparison
instance NFData CompareAs
instance NFData a => NFData (Open a)
instance NFData a => NFData (Judgement a)
instance NFData DoGeneralize
instance NFData GeneralizedValue
instance NFData MetaVariable
instance NFData Listener
instance NFData MetaInstantiation
instance NFData Instantiation
instance NFData RemoteMetaVariable
instance NFData Frozen
instance NFData PrincipalArgTypeMetas
instance NFData TypeCheckingProblem
instance NFData RunMetaOccursCheck
instance NFData MetaInfo
instance NFData InteractionPoint
instance NFData InteractionPoints
instance NFData Overapplied
instance NFData t => NFData (IPBoundary' t)
instance NFData IPClause
instance NFData DisplayForm
instance NFData DisplayTerm
instance NFData NLPat
instance NFData NLPType
instance NFData NLPSort
instance NFData RewriteRule
instance NFData InstanceInfo
instance NFData Definition
instance NFData Polarity
instance NFData IsForced
instance NFData Projection
instance NFData ProjLams
instance NFData CompilerPragma
instance NFData System
instance NFData ExtLamInfo
instance NFData EtaEquality
instance NFData FunctionFlag
instance NFData CompKit
instance NFData AxiomData
instance NFData DataOrRecSigData
instance NFData ProjectionLikenessMissing
instance NFData FunctionData
instance NFData DatatypeData
instance NFData RecordData
instance NFData ConstructorData
instance NFData PrimitiveData
instance NFData PrimitiveSortData
instance NFData Defn
instance NFData Simplification
instance NFData AllowedReduction
instance NFData ReduceDefs
instance NFData PrimFun
instance NFData c => NFData (FunctionInverse' c)
instance NFData TermHead
instance NFData Call
instance NFData BuiltinSort
instance NFData pf => NFData (Builtin pf)
instance NFData HighlightingLevel
instance NFData HighlightingMethod
instance NFData TCEnv
instance NFData LetBinding
instance NFData UnquoteFlags
instance NFData AbstractMode
instance NFData ExpandHidden
instance NFData CandidateKind
instance NFData Candidate
instance NFData Warning
instance NFData RecordFieldWarning
instance NFData TCWarning
instance NFData CallInfo
instance NFData TerminationError
instance NFData SplitError
instance NFData NegativeUnification
instance NFData UnificationFailure
instance NFData UnquoteError
instance NFData TypeError
instance NFData WhyInvalidInstanceType
instance NFData InvalidFileNameReason
instance NFData LHSOrPatSyn
instance NFData DataOrRecordE
instance NFData InductionAndEta
instance NFData IllegalRewriteRuleReason
instance NFData IncorrectTypeForRewriteRelationReason
instance NFData GHCBackendError
instance NFData JSBackendError
instance NFData MissingTypeSignatureInfo
instance NFData WhyNotAHaskellType
instance NFData InteractionError
instance NFData IsAmbiguous
instance NFData CannotQuote
instance NFData ExecError