{-# OPTIONS_GHC -fwarn-unused-binds #-}
module Agda.Interaction.Highlighting.FromAbstract
( runHighlighter
, NameKinds
) where
import Prelude hiding (null)
import Control.Applicative
import Control.Monad.Reader ( MonadReader(..), asks, Reader, runReader )
import qualified Data.Map as Map
import Data.Maybe
import Data.Semigroup ( Semigroup(..) )
import Data.Void ( Void )
import Agda.Interaction.Highlighting.Precise hiding ( singleton )
import qualified Agda.Interaction.Highlighting.Precise as H
import Agda.Interaction.Highlighting.Range ( rToR )
import Agda.Syntax.Abstract ( IsProjP(..) )
import qualified Agda.Syntax.Abstract as A
import Agda.Syntax.Common as Common
import Agda.Syntax.Concrete ( FieldAssignment'(..), TacticAttribute' )
import qualified Agda.Syntax.Concrete.Name as C
import Agda.Syntax.Info ( ModuleInfo(..) )
import Agda.Syntax.Literal
import qualified Agda.Syntax.Position as P
import Agda.Syntax.Position ( Range, HasRange, getRange, noRange )
import Agda.Syntax.Scope.Base ( AbstractName(..), ResolvedName(..), exactConName )
import Agda.Syntax.TopLevelModuleName
import Agda.TypeChecking.Monad
hiding (ModuleInfo, MetaInfo, Primitive, Constructor, Record, Function, Datatype)
import Agda.Utils.Function
import Agda.Utils.Functor
import Agda.Utils.List ( initLast1 )
import Agda.Utils.List1 ( List1 )
import qualified Agda.Utils.List1 as List1
import Agda.Utils.Maybe
import qualified Agda.Utils.Maybe.Strict as Strict
import Agda.Syntax.Common.Pretty
import Agda.Utils.Singleton
import Agda.Utils.Size
runHighlighter
:: Hilite a
=> TopLevelModuleName
-> NameKinds -> a -> HighlightingInfoBuilder
runHighlighter :: forall a.
Hilite a =>
TopLevelModuleName' Range
-> NameKinds -> a -> HighlightingInfoBuilder
runHighlighter TopLevelModuleName' Range
top NameKinds
kinds a
x =
ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> HiliteEnv -> HighlightingInfoBuilder
forall r a. Reader r a -> r -> a
runReader (a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite a
x) (HiliteEnv -> HighlightingInfoBuilder)
-> HiliteEnv -> HighlightingInfoBuilder
forall a b. (a -> b) -> a -> b
$
HiliteEnv
{ hleNameKinds :: NameKinds
hleNameKinds = NameKinds
kinds
, hleCurrentModuleName :: TopLevelModuleName' Range
hleCurrentModuleName = TopLevelModuleName' Range
top
}
data HiliteEnv = HiliteEnv
{ HiliteEnv -> NameKinds
hleNameKinds :: NameKinds
, HiliteEnv -> TopLevelModuleName' Range
hleCurrentModuleName :: TopLevelModuleName
}
type NameKinds = A.QName -> Maybe NameKind
type HiliteM = Reader HiliteEnv
type Hiliter = HiliteM HighlightingInfoBuilder
instance Monoid Hiliter where
mempty :: ReaderT HiliteEnv Identity HighlightingInfoBuilder
mempty = HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. a -> ReaderT HiliteEnv Identity a
forall (f :: * -> *) a. Applicative f => a -> f a
pure HighlightingInfoBuilder
forall a. Monoid a => a
mempty
mappend :: ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
mappend = ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
(<>)
class Hilite a where
hilite :: a -> Hiliter
default hilite :: (Foldable t, Hilite b, t b ~ a) => a -> Hiliter
hilite = (b -> ReaderT HiliteEnv Identity HighlightingInfoBuilder)
-> t b -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall m a. Monoid m => (a -> m) -> t a -> m
forall (t :: * -> *) m a.
(Foldable t, Monoid m) =>
(a -> m) -> t a -> m
foldMap b -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite
instance Hilite a => Hilite [a]
instance Hilite a => Hilite (List1 a)
instance Hilite a => Hilite (Maybe a)
instance Hilite a => Hilite (Ranged a)
instance Hilite a => Hilite (TacticAttribute' a)
instance Hilite a => Hilite (WithHiding a)
instance Hilite Void where
hilite :: Void -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite Void
_ = ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Monoid a => a
mempty
instance (Hilite a, Hilite b) => Hilite (Either a b) where
hilite :: Either a b -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite = (a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder)
-> (b -> ReaderT HiliteEnv Identity HighlightingInfoBuilder)
-> Either a b
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a c b. (a -> c) -> (b -> c) -> Either a b -> c
either a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite b -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite
instance (Hilite a, Hilite b) => Hilite (a, b) where
hilite :: (a, b) -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite (a
a, b
b) = a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite a
a ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> b -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite b
b
instance Hilite A.RecordDirectives where
hilite :: RecordDirectives
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite (RecordDirectives Maybe (Ranged Induction)
_ Maybe (Ranged HasEta0)
_ Maybe Range
_ RecordConName
c) = RecordConName -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite RecordConName
c
instance Hilite A.RecordConName where
hilite :: RecordConName -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite = \case
A.NamedRecCon QName
x -> QName -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite QName
x
A.FreshRecCon{} -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Monoid a => a
mempty
instance Hilite A.Declaration where
hilite :: Declaration -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite = \case
A.Axiom KindOfName
_ax DefInfo
_di ArgInfo
ai Maybe (List1 Occurrence)
_occ QName
x Type
e -> ArgInfo -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl ArgInfo
ai ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> QName -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl QName
x ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> Type -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl Type
e
A.Generalize Set QName
_names DefInfo
_di ArgInfo
ai QName
x Type
e -> ArgInfo -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl ArgInfo
ai ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> QName -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl QName
x ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> Type -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl Type
e
A.Field DefInfo
_di QName
x Arg Type
e -> QName -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hlField QName
x ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> Arg Type -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl Arg Type
e
A.Primitive DefInfo
_di QName
x Arg Type
e -> QName -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl QName
x ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> Arg Type -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl Arg Type
e
A.Mutual MutualInfo
_mi [Declaration]
ds -> [Declaration] -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl [Declaration]
ds
A.Section Range
_r Erased
er ModuleName
x GeneralizeTelescope
tel [Declaration]
ds -> Erased -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl Erased
er ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> ModuleName -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl ModuleName
x ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> GeneralizeTelescope
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl GeneralizeTelescope
tel ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> [Declaration] -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl [Declaration]
ds
A.Apply ModuleInfo
mi Erased
er ModuleName
x ModuleApplication
a ScopeCopyInfo
_ci ImportDirective
dir -> ModuleInfo -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl ModuleInfo
mi ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> Erased -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl Erased
er ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> ModuleName -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl ModuleName
x ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<>
ModuleApplication
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl ModuleApplication
a ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> ImportDirective
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl ImportDirective
dir
A.Import ModuleInfo
mi ModuleName
x ImportDirective
dir -> ModuleInfo -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl ModuleInfo
mi ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> ModuleName -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl ModuleName
x ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> ImportDirective
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl ImportDirective
dir
A.Open ModuleInfo
mi ModuleName
x ImportDirective
dir -> ModuleInfo -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl ModuleInfo
mi ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> ModuleName -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl ModuleName
x ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> ImportDirective
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl ImportDirective
dir
A.FunDef DefInfo
_di QName
x [Clause]
cs -> QName -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl QName
x ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> [Clause] -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl [Clause]
cs
A.DataSig DefInfo
_di Erased
er QName
x GeneralizeTelescope
tel Type
e -> Erased -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl Erased
er ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> QName -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl QName
x ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> GeneralizeTelescope
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl GeneralizeTelescope
tel ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> Type -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl Type
e
A.DataDef DefInfo
_di QName
x UniverseCheck
_uc DataDefParams
pars [Declaration]
cs -> QName -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl QName
x ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> DataDefParams -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl DataDefParams
pars ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> [Declaration] -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl [Declaration]
cs
A.RecSig DefInfo
_di Erased
er QName
x GeneralizeTelescope
tel Type
e -> Erased -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl Erased
er ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> QName -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl QName
x ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> GeneralizeTelescope
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl GeneralizeTelescope
tel ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> Type -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl Type
e
A.RecDef DefInfo
_di QName
x UniverseCheck
_uc RecordDirectives
dir DataDefParams
bs Type
e [Declaration]
ds -> QName -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl QName
x ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> RecordDirectives
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl RecordDirectives
dir ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> DataDefParams -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl DataDefParams
bs ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> Type -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl Type
e ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> [Declaration] -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl [Declaration]
ds
A.PatternSynDef QName
x [WithHiding BindName]
xs Pattern' Void
p -> QName -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl QName
x ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> [WithHiding BindName]
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl [WithHiding BindName]
xs ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> Pattern' Void -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl Pattern' Void
p
A.UnquoteDecl MutualInfo
_mi [DefInfo]
_di [QName]
xs Type
e -> [QName] -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl [QName]
xs ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> Type -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl Type
e
A.UnquoteDef [DefInfo]
_di [QName]
xs Type
e -> [QName] -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl [QName]
xs ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> Type -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl Type
e
A.UnquoteData [DefInfo]
_i QName
xs UniverseCheck
_uc [DefInfo]
_j [QName]
cs Type
e -> QName -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl QName
xs ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> [QName] -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl [QName]
cs ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> Type -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl Type
e
A.ScopedDecl ScopeInfo
s [Declaration]
ds -> [Declaration] -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl [Declaration]
ds
A.Pragma Range
_r Pragma
pragma -> Pragma -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl Pragma
pragma
A.UnfoldingDecl Range
_r [QName]
names -> [QName] -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl [QName]
names
where
hl :: a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl a
a = a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite a
a
hlField :: QName -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hlField QName
x = [Name]
-> Name
-> Maybe Range
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hiliteField (QName -> [Name]
concreteQualifier QName
x) (QName -> Name
concreteBase QName
x) (Range -> Maybe Range
forall a. a -> Maybe a
Just (Range -> Maybe Range) -> Range -> Maybe Range
forall a b. (a -> b) -> a -> b
$ QName -> Range
bindingSite QName
x)
instance Hilite A.Pragma where
hilite :: Pragma -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite = \case
A.OptionsPragma [String]
_strings -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Monoid a => a
mempty
A.BuiltinPragma RString
b ResolvedName
x -> Aspect
-> RString -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
HasRange a =>
Aspect -> a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
singleAspect Aspect
Keyword RString
b ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> ResolvedName -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite ResolvedName
x
A.BuiltinNoDefPragma RString
b KindOfName
k QName
x -> Aspect
-> RString -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
HasRange a =>
Aspect -> a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
singleAspect Aspect
Keyword RString
b ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> Maybe NameKind
-> QName -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hiliteQName (NameKind -> Maybe NameKind
forall a. a -> Maybe a
Just (NameKind -> Maybe NameKind) -> NameKind -> Maybe NameKind
forall a b. (a -> b) -> a -> b
$ KindOfName -> NameKind
kindOfNameToNameKind KindOfName
k) QName
x
A.CompilePragma Ranged BackendName
b QName
x String
_foreign -> Aspect
-> Ranged BackendName
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
HasRange a =>
Aspect -> a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
singleAspect Aspect
Keyword Ranged BackendName
b ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> QName -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite QName
x
A.RewritePragma Range
r [QName]
xs -> Aspect
-> Range -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
HasRange a =>
Aspect -> a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
singleAspect Aspect
Keyword Range
r ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> [QName] -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite [QName]
xs
A.StaticPragma QName
x -> QName -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite QName
x
A.EtaPragma QName
x -> QName -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite QName
x
A.InjectivePragma QName
x -> QName -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite QName
x
A.InjectiveForInferencePragma QName
x -> QName -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite QName
x
A.NotProjectionLikePragma QName
x -> QName -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite QName
x
A.OverlapPragma QName
x OverlapMode
_ -> QName -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite QName
x
A.InlinePragma Bool
_inline QName
x -> QName -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite QName
x
A.DisplayPragma QName
x [NamedArg Pattern]
ps Type
e -> QName -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite QName
x ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> [NamedArg Pattern]
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite [NamedArg Pattern]
ps ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> Type -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite Type
e
instance Hilite A.Expr where
hilite :: Type -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite = \case
A.Var Name
x -> BindName -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl (BindName -> ReaderT HiliteEnv Identity HighlightingInfoBuilder)
-> BindName -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a b. (a -> b) -> a -> b
$ Name -> BindName
A.BindName Name
x
A.Def' QName
q Suffix
_ -> Maybe NameKind
-> QName -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hiliteQName Maybe NameKind
forall a. Maybe a
Nothing QName
q
A.Proj ProjOrigin
_o AmbiguousQName
qs -> Maybe NameKind
-> AmbiguousQName
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hiliteAmbiguousQName Maybe NameKind
forall a. Maybe a
Nothing AmbiguousQName
qs
A.Con AmbiguousQName
qs -> Maybe NameKind
-> AmbiguousQName
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hiliteAmbiguousQName Maybe NameKind
forall a. Maybe a
Nothing AmbiguousQName
qs
A.PatternSyn AmbiguousQName
qs -> AmbiguousQName
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilitePatternSynonym AmbiguousQName
qs
A.Macro QName
q -> Maybe NameKind
-> QName -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hiliteQName (NameKind -> Maybe NameKind
forall a. a -> Maybe a
Just NameKind
Macro) QName
q
A.Lit ExprInfo
_r Literal
l -> Literal -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl Literal
l
A.QuestionMark MetaInfo
_mi InteractionId
_ii -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Monoid a => a
mempty
A.Underscore MetaInfo
_mi -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Monoid a => a
mempty
A.Dot ExprInfo
_r Type
e -> Type -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl Type
e
A.App AppInfo
_r Type
e NamedArg Type
es -> Type -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl Type
e ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> NamedArg Type -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl NamedArg Type
es
A.WithApp ExprInfo
_r Type
e List1 Type
es -> Type -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl Type
e ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> List1 Type -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl List1 Type
es
A.Lam ExprInfo
_r LamBinding
bs Type
e -> LamBinding -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl LamBinding
bs ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> Type -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl Type
e
A.AbsurdLam ExprInfo
_r Hiding
_h -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Monoid a => a
mempty
A.ExtendedLam ExprInfo
_r DefInfo
_di Erased
er QName
_q List1 Clause
cs -> Erased -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl Erased
er ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> List1 Clause -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl List1 Clause
cs
A.Pi ExprInfo
_r Telescope1
tel Type
b -> Telescope1 -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl Telescope1
tel ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> Type -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl Type
b
A.Generalized Set1 QName
_qs Type
e -> Type -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl Type
e
A.Fun ExprInfo
_r Arg Type
a Type
b -> Arg Type -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl Arg Type
a ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> Type -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl Type
b
A.Let ExprInfo
_r List1 LetBinding
bs Type
e -> List1 LetBinding
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl List1 LetBinding
bs ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> Type -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl Type
e
A.Rec RecInfo
_r RecordAssigns
ass -> RecordAssigns -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl RecordAssigns
ass
A.RecUpdate RecInfo
_r Type
e Assigns
ass -> Type -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl Type
e ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> Assigns -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl Assigns
ass
A.ScopedExpr ScopeInfo
_ Type
e -> Type -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl Type
e
A.Quote ExprInfo
_r -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Monoid a => a
mempty
A.QuoteTerm ExprInfo
_r -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Monoid a => a
mempty
A.Unquote ExprInfo
_r -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Monoid a => a
mempty
A.DontCare Type
e -> Type -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl Type
e
where
hl :: a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl a
a = a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite a
a
instance (Hilite a, IsProjP a) => Hilite (A.Pattern' a) where
hilite :: Pattern' a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite = \case
A.VarP BindName
x -> BindName -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl BindName
x
A.ConP ConPatInfo
_i AmbiguousQName
qs NAPs a
es -> AmbiguousQName
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hiliteInductiveConstructor AmbiguousQName
qs ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> NAPs a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl NAPs a
es
A.ProjP PatInfo
_r ProjOrigin
_o AmbiguousQName
qs -> AmbiguousQName
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hiliteProjection AmbiguousQName
qs
A.DefP PatInfo
_r AmbiguousQName
qs NAPs a
es -> AmbiguousQName
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl AmbiguousQName
qs ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> NAPs a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl NAPs a
es
A.WildP PatInfo
_r -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Monoid a => a
mempty
A.AsP PatInfo
_r BindName
x Pattern' a
p -> BindName -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl BindName
x ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> Pattern' a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl Pattern' a
p
A.DotP PatInfo
r a
e -> case a -> Maybe (ProjOrigin, AmbiguousQName)
forall a. IsProjP a => a -> Maybe (ProjOrigin, AmbiguousQName)
isProjP a
e of
Maybe (ProjOrigin, AmbiguousQName)
Nothing -> OtherAspect
-> PatInfo -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
HasRange a =>
OtherAspect
-> a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
singleOtherAspect OtherAspect
DottedPattern PatInfo
r ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl a
e
Just (ProjOrigin
_o, AmbiguousQName
qs) -> AmbiguousQName
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hiliteProjection AmbiguousQName
qs
A.AbsurdP PatInfo
_r -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Monoid a => a
mempty
A.LitP PatInfo
_r Literal
l -> Literal -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl Literal
l
A.PatternSynP PatInfo
_r AmbiguousQName
qs NAPs a
es -> AmbiguousQName
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilitePatternSynonym AmbiguousQName
qs ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> NAPs a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl NAPs a
es
A.RecP ConPatInfo
_r [FieldAssignment' (Pattern' a)]
ps -> [FieldAssignment' (Pattern' a)]
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl [FieldAssignment' (Pattern' a)]
ps
A.EqualP PatInfo
_r List1 (a, a)
ps -> List1 (a, a) -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl List1 (a, a)
ps
A.WithP PatInfo
_ Pattern' a
p -> Pattern' a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl Pattern' a
p
where
hl :: a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl a
a = a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite a
a
instance Hilite Literal where
hilite :: Literal -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite = \case
LitNat{} -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Monoid a => a
mempty
LitWord64{} -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Monoid a => a
mempty
LitFloat{} -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Monoid a => a
mempty
LitString{} -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Monoid a => a
mempty
LitChar{} -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Monoid a => a
mempty
LitQName QName
x -> QName -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite QName
x
LitMeta TopLevelModuleName' Range
_fileName MetaId
_id -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Monoid a => a
mempty
instance Hilite A.LHS where
hilite :: LHS -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite (A.LHS LHSInfo
_r LHSCore
lhs) = LHSCore -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite LHSCore
lhs
instance (Hilite a, IsProjP a) => Hilite (A.LHSCore' a) where
hilite :: LHSCore' a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite = \case
A.LHSHead QName
q [NamedArg (Pattern' a)]
ps -> QName -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite QName
q ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> [NamedArg (Pattern' a)]
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite [NamedArg (Pattern' a)]
ps
A.LHSProj AmbiguousQName
q NamedArg (LHSCore' a)
lhs [NamedArg (Pattern' a)]
ps -> NamedArg (LHSCore' a)
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite NamedArg (LHSCore' a)
lhs ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> AmbiguousQName
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite AmbiguousQName
q ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> [NamedArg (Pattern' a)]
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite [NamedArg (Pattern' a)]
ps
A.LHSWith LHSCore' a
lhs List1 (Arg (Pattern' a))
wps [NamedArg (Pattern' a)]
ps -> LHSCore' a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite LHSCore' a
lhs ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> List1 (Arg (Pattern' a))
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite List1 (Arg (Pattern' a))
wps ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> [NamedArg (Pattern' a)]
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite [NamedArg (Pattern' a)]
ps
instance Hilite A.RHS where
hilite :: RHS -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite = \case
A.RHS Type
e Maybe Expr
_ce -> Type -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl Type
e
RHS
A.AbsurdRHS -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Monoid a => a
mempty
A.WithRHS QName
_q List1 WithExpr
es List1 Clause
cs -> List1 WithExpr
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl List1 WithExpr
es ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> List1 Clause -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl List1 Clause
cs
A.RewriteRHS [RewriteEqn]
eqs [ProblemEq]
strippedPats RHS
rhs WhereDeclarations
wh -> [RewriteEqn] -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl [RewriteEqn]
eqs ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> [ProblemEq] -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl [ProblemEq]
strippedPats ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> RHS -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl RHS
rhs ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> WhereDeclarations
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl WhereDeclarations
wh
where
hl :: a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl a
a = a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite a
a
instance (HasRange n, Hilite p, Hilite e) => Hilite (RewriteEqn' x n p e) where
hilite :: RewriteEqn' x n p e
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite = \case
Rewrite List1 (x, e)
es -> NonEmpty e -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite (NonEmpty e -> ReaderT HiliteEnv Identity HighlightingInfoBuilder)
-> NonEmpty e -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a b. (a -> b) -> a -> b
$ ((x, e) -> e) -> List1 (x, e) -> NonEmpty e
forall a b. (a -> b) -> NonEmpty a -> NonEmpty b
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
fmap (x, e) -> e
forall a b. (a, b) -> b
snd List1 (x, e)
es
Invert x
_x List1 (Named n (p, e))
pes -> List1 (Named n (p, e))
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite List1 (Named n (p, e))
pes
LeftLet List1 (p, e)
pes -> List1 (p, e) -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite List1 (p, e)
pes
instance Hilite a => Hilite (A.Clause' a) where
hilite :: Clause' a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite (A.Clause a
lhs [ProblemEq]
strippedPats RHS
rhs WhereDeclarations
wh Bool
_catchall) =
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite a
lhs ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> [ProblemEq] -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite [ProblemEq]
strippedPats ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> RHS -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite RHS
rhs ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> WhereDeclarations
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite WhereDeclarations
wh
instance Hilite A.ProblemEq where
hilite :: ProblemEq -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite (A.ProblemEq Pattern
p Term
_t Dom Type
_dom) = Pattern -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite Pattern
p
instance Hilite A.WhereDeclarations where
hilite :: WhereDeclarations
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite (A.WhereDecls Maybe ModuleName
m Bool
_ Maybe Declaration
ds) = Maybe ModuleName
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite Maybe ModuleName
m ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> Maybe Declaration
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite Maybe Declaration
ds
instance Hilite A.GeneralizeTelescope where
hilite :: GeneralizeTelescope
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite (A.GeneralizeTel Map QName Name
_gen Telescope
tel) = Telescope -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite Telescope
tel
instance Hilite A.DataDefParams where
hilite :: DataDefParams -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite (A.DataDefParams Set Name
_gen [LamBinding]
pars) = [LamBinding] -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite [LamBinding]
pars
instance Hilite A.ModuleApplication where
hilite :: ModuleApplication
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite = \case
A.SectionApp Telescope
tel ModuleName
x [NamedArg Type]
es -> Telescope -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite Telescope
tel ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> ModuleName -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite ModuleName
x ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> [NamedArg Type]
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite [NamedArg Type]
es
A.RecordModuleInstance ModuleName
x -> ModuleName -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite ModuleName
x
instance Hilite A.LetBinding where
hilite :: LetBinding -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite = \case
A.LetBind LetInfo
_r ArgInfo
ai BindName
x Type
t Type
e -> ArgInfo -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl ArgInfo
ai ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> BindName -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl BindName
x ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> Type -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl Type
t ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> Type -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl Type
e
A.LetAxiom LetInfo
_r ArgInfo
ai BindName
x Type
t -> ArgInfo -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl ArgInfo
ai ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> BindName -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl BindName
x ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> Type -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl Type
t
A.LetPatBind LetInfo
_r Pattern
p Type
e -> Pattern -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl Pattern
p ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> Type -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl Type
e
A.LetApply ModuleInfo
mi Erased
er ModuleName
x ModuleApplication
es ScopeCopyInfo
_c ImportDirective
dir -> ModuleInfo -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl ModuleInfo
mi ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> Erased -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl Erased
er ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> ModuleName -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl ModuleName
x ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<>
ModuleApplication
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl ModuleApplication
es ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> ImportDirective
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl ImportDirective
dir
A.LetOpen ModuleInfo
mi ModuleName
x ImportDirective
dir -> ModuleInfo -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl ModuleInfo
mi ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> ModuleName -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl ModuleName
x ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> ImportDirective
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl ImportDirective
dir
A.LetDeclaredVariable BindName
x -> BindName -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl BindName
x
where
hl :: a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hl a
x = a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite a
x
instance Hilite A.TypedBindingInfo where
hilite :: TypedBindingInfo
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite (A.TypedBindingInfo TacticAttribute
x Bool
_) = TacticAttribute
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite TacticAttribute
x
instance Hilite A.TypedBinding where
hilite :: TypedBinding -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite = \case
A.TBind Range
_r TypedBindingInfo
tac List1 (NamedArg Binder)
binds Type
e -> TypedBindingInfo
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite TypedBindingInfo
tac ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> List1 (NamedArg Binder)
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite List1 (NamedArg Binder)
binds ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> Type -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite Type
e
A.TLet Range
_r List1 LetBinding
binds -> List1 LetBinding
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite List1 LetBinding
binds
instance Hilite A.LamBinding where
hilite :: LamBinding -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite = \case
A.DomainFree TacticAttribute
tac NamedArg Binder
binds -> TacticAttribute
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite TacticAttribute
tac ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> NamedArg Binder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite NamedArg Binder
binds
A.DomainFull TypedBinding
bind -> TypedBinding -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite TypedBinding
bind
instance Hilite a => Hilite (A.Binder' a) where
hilite :: Binder' a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite (A.Binder Maybe Pattern
p BinderNameOrigin
_ a
x) = Maybe Pattern -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite Maybe Pattern
p ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite a
x
instance Hilite A.BindName where
hilite :: BindName -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite (A.BindName Name
x) = Name -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hiliteBound Name
x
instance Hilite a => Hilite (FieldAssignment' a) where
hilite :: FieldAssignment' a
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite (FieldAssignment Name
x a
e) = [Name]
-> Name
-> Maybe Range
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hiliteField [] Name
x Maybe Range
forall a. Maybe a
Nothing ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite a
e
instance (Hilite a, HasRange n) => Hilite (Named n a) where
hilite :: Named n a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite (Named Maybe n
mn a
e)
= ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> (n -> ReaderT HiliteEnv Identity HighlightingInfoBuilder)
-> Maybe n
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall b a. b -> (a -> b) -> Maybe a -> b
maybe ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Monoid a => a
mempty (Aspect -> n -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
HasRange a =>
Aspect -> a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
singleAspect (Aspect -> n -> ReaderT HiliteEnv Identity HighlightingInfoBuilder)
-> Aspect
-> n
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a b. (a -> b) -> a -> b
$ Maybe NameKind -> Bool -> Aspect
Name (NameKind -> Maybe NameKind
forall a. a -> Maybe a
Just NameKind
Argument) Bool
False) Maybe n
mn
ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite a
e
instance Hilite a => Hilite (Arg a) where
hilite :: Arg a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite (Arg ArgInfo
ai a
e) = ArgInfo -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite ArgInfo
ai ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite a
e
instance Hilite ArgInfo where
hilite :: ArgInfo -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite (ArgInfo Hiding
_hiding Modality
modality Origin
_origin FreeVariables
_fv Annotation
_a) = Modality -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite Modality
modality
instance Hilite Modality where
hilite :: Modality -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite (Modality Relevance
_relevance Quantity
quantity Cohesion
_cohesion) = Quantity -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite Quantity
quantity
instance Hilite Quantity where
hilite :: Quantity -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite = Aspect
-> Quantity -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
HasRange a =>
Aspect -> a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
singleAspect Aspect
Symbol
instance Hilite Erased where
hilite :: Erased -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite = Aspect
-> Erased -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
HasRange a =>
Aspect -> a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
singleAspect Aspect
Symbol
instance Hilite ModuleInfo where
hilite :: ModuleInfo -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite (ModuleInfo Range
_r Range
rAsTo Maybe Name
asName Maybe OpenShortHand
_open Maybe ImportDirective
_impDir)
= Aspect
-> Range -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
HasRange a =>
Aspect -> a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
singleAspect Aspect
Symbol Range
rAsTo
ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> (Name -> ReaderT HiliteEnv Identity HighlightingInfoBuilder)
-> Maybe Name
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall b a. b -> (a -> b) -> Maybe a -> b
maybe ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Monoid a => a
mempty Name -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hiliteAsName Maybe Name
asName
where
hiliteAsName :: C.Name -> Hiliter
hiliteAsName :: Name -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hiliteAsName Name
n = [Name]
-> Name
-> Range
-> Maybe Range
-> (Bool -> Aspects)
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hiliteCName [] Name
n Range
forall a. Range' a
noRange Maybe Range
forall a. Maybe a
Nothing ((Bool -> Aspects)
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder)
-> (Bool -> Aspects)
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a b. (a -> b) -> a -> b
$ NameKind -> Bool -> Aspects
nameAsp NameKind
Module
instance (Hilite m, Hilite n, Hilite (RenamingTo m), Hilite (RenamingTo n))
=> Hilite (ImportDirective' m n) where
hilite :: ImportDirective' m n
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite (ImportDirective Range
_r Using' m n
using HidingDirective' m n
hiding RenamingDirective' m n
renaming Maybe KwRange
_ropen) =
Using' m n -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite Using' m n
using ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> HidingDirective' m n
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite HidingDirective' m n
hiding ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> RenamingDirective' m n
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite RenamingDirective' m n
renaming
instance (Hilite m, Hilite n) => Hilite (Using' m n) where
hilite :: Using' m n -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite = \case
Using' m n
UseEverything -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Monoid a => a
mempty
Using [ImportedName' m n]
using -> [ImportedName' m n]
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite [ImportedName' m n]
using
instance (Hilite m, Hilite n, Hilite (RenamingTo m), Hilite (RenamingTo n))
=> Hilite (Renaming' m n) where
hilite :: Renaming' m n -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite (Renaming ImportedName' m n
from ImportedName' m n
to Maybe Fixity
_fixity Range
rangeKwTo)
= ImportedName' m n
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite ImportedName' m n
from
ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> Aspect
-> Range -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
HasRange a =>
Aspect -> a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
singleAspect Aspect
Symbol Range
rangeKwTo
ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Semigroup a => a -> a -> a
<> RenamingTo (ImportedName' m n)
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite (ImportedName' m n -> RenamingTo (ImportedName' m n)
forall a. a -> RenamingTo a
RenamingTo ImportedName' m n
to)
instance (Hilite m, Hilite n) => Hilite (ImportedName' m n) where
hilite :: ImportedName' m n
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite = \case
ImportedModule n
m -> n -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite n
m
ImportedName m
n -> m -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite m
n
instance Hilite DisambiguatedName where
hilite :: DisambiguatedName
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite (DisambiguatedName NameKind
k QName
x) = Maybe NameKind
-> QName -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hiliteQName (NameKind -> Maybe NameKind
forall a. a -> Maybe a
Just NameKind
k) QName
x
instance Hilite ResolvedName where
hilite :: ResolvedName -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite = \case
VarName Name
x BindingSource
_bindSrc -> Name -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hiliteBound Name
x
DefinedName Access
_acc AbstractName
x Suffix
_suffix -> QName -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite (QName -> ReaderT HiliteEnv Identity HighlightingInfoBuilder)
-> QName -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a b. (a -> b) -> a -> b
$ AbstractName -> QName
anameName AbstractName
x
FieldName List1 AbstractName
xs -> AmbiguousQName
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hiliteProjection (AmbiguousQName
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder)
-> AmbiguousQName
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a b. (a -> b) -> a -> b
$ List1 QName -> AmbiguousQName
A.AmbQ (List1 QName -> AmbiguousQName) -> List1 QName -> AmbiguousQName
forall a b. (a -> b) -> a -> b
$ (AbstractName -> QName) -> List1 AbstractName -> List1 QName
forall a b. (a -> b) -> NonEmpty a -> NonEmpty b
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
fmap AbstractName -> QName
anameName List1 AbstractName
xs
ConstructorName Set1 Induction
i List1 AbstractName
xs -> Maybe NameKind
-> AmbiguousQName
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hiliteAmbiguousQName Maybe NameKind
k (AmbiguousQName
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder)
-> AmbiguousQName
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a b. (a -> b) -> a -> b
$ List1 QName -> AmbiguousQName
A.AmbQ (List1 QName -> AmbiguousQName) -> List1 QName -> AmbiguousQName
forall a b. (a -> b) -> a -> b
$ (AbstractName -> QName) -> List1 AbstractName -> List1 QName
forall a b. (a -> b) -> NonEmpty a -> NonEmpty b
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
fmap AbstractName -> QName
anameName List1 AbstractName
xs
where k :: Maybe NameKind
k = KindOfName -> NameKind
kindOfNameToNameKind (KindOfName -> NameKind) -> Maybe KindOfName -> Maybe NameKind
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> Set1 Induction -> Maybe KindOfName
forall (t :: * -> *). Foldable t => t Induction -> Maybe KindOfName
exactConName Set1 Induction
i
PatternSynResName List1 AbstractName
xs -> AmbiguousQName
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilitePatternSynonym (AmbiguousQName
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder)
-> AmbiguousQName
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a b. (a -> b) -> a -> b
$ List1 QName -> AmbiguousQName
A.AmbQ (List1 QName -> AmbiguousQName) -> List1 QName -> AmbiguousQName
forall a b. (a -> b) -> a -> b
$ (AbstractName -> QName) -> List1 AbstractName -> List1 QName
forall a b. (a -> b) -> NonEmpty a -> NonEmpty b
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
fmap AbstractName -> QName
anameName List1 AbstractName
xs
ResolvedName
UnknownName -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Monoid a => a
mempty
instance Hilite A.QName where
hilite :: QName -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite = Maybe NameKind
-> QName -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hiliteQName Maybe NameKind
forall a. Maybe a
Nothing
instance Hilite A.AmbiguousQName where
hilite :: AmbiguousQName
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite = Maybe NameKind
-> AmbiguousQName
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hiliteAmbiguousQName Maybe NameKind
forall a. Maybe a
Nothing
instance Hilite A.ModuleName where
hilite :: ModuleName -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite m :: ModuleName
m@(A.MName [Name]
xs) = (Bool, ModuleName)
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hiliteModule (Bool
isTopLevelModule, ModuleName
m)
where
isTopLevelModule :: Bool
isTopLevelModule =
case (Name -> Maybe (TopLevelModuleName' Range))
-> [Name] -> [TopLevelModuleName' Range]
forall a b. (a -> Maybe b) -> [a] -> [b]
mapMaybe (Range -> Maybe (TopLevelModuleName' Range)
P.rangeModule (Range -> Maybe (TopLevelModuleName' Range))
-> (Name -> Range) -> Name -> Maybe (TopLevelModuleName' Range)
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Name -> Range
A.nameBindingSite) [Name]
xs of
[] -> Bool
False
TopLevelModuleName' Range
top : [TopLevelModuleName' Range]
_ ->
TopLevelModuleName' Range -> RawTopLevelModuleName
rawTopLevelModuleName TopLevelModuleName' Range
top RawTopLevelModuleName -> RawTopLevelModuleName -> Bool
forall a. Eq a => a -> a -> Bool
==
ModuleName -> RawTopLevelModuleName
rawTopLevelModuleNameForModuleName ModuleName
m
newtype RenamingTo a = RenamingTo a
instance Hilite (RenamingTo A.QName) where
hilite :: RenamingTo QName
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite (RenamingTo QName
q) = do
kind <- (HiliteEnv -> NameKinds) -> ReaderT HiliteEnv Identity NameKinds
forall r (m :: * -> *) a. MonadReader r m => (r -> a) -> m a
asks HiliteEnv -> NameKinds
hleNameKinds ReaderT HiliteEnv Identity NameKinds
-> (NameKinds -> Maybe NameKind)
-> ReaderT HiliteEnv Identity (Maybe NameKind)
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> (NameKinds -> NameKinds
forall a b. (a -> b) -> a -> b
$ QName
q)
hiliteAName q False $ nameAsp' kind
instance Hilite (RenamingTo A.ModuleName) where
hilite :: RenamingTo ModuleName
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite (RenamingTo (A.MName [Name]
ns)) = ((Name -> ReaderT HiliteEnv Identity HighlightingInfoBuilder)
-> [Name] -> ReaderT HiliteEnv Identity HighlightingInfoBuilder)
-> [Name]
-> (Name -> ReaderT HiliteEnv Identity HighlightingInfoBuilder)
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a b c. (a -> b -> c) -> b -> a -> c
flip (Name -> ReaderT HiliteEnv Identity HighlightingInfoBuilder)
-> [Name] -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall m a. Monoid m => (a -> m) -> [a] -> m
forall (t :: * -> *) m a.
(Foldable t, Monoid m) =>
(a -> m) -> t a -> m
foldMap [Name]
ns ((Name -> ReaderT HiliteEnv Identity HighlightingInfoBuilder)
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder)
-> (Name -> ReaderT HiliteEnv Identity HighlightingInfoBuilder)
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a b. (a -> b) -> a -> b
$ \ Name
n ->
[Name]
-> Name
-> Range
-> Maybe Range
-> (Bool -> Aspects)
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hiliteCName [] (Name -> Name
A.nameConcrete Name
n) Range
forall a. Range' a
noRange Maybe Range
forall a. Maybe a
Nothing ((Bool -> Aspects)
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder)
-> (Bool -> Aspects)
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a b. (a -> b) -> a -> b
$ NameKind -> Bool -> Aspects
nameAsp NameKind
Module
instance (Hilite (RenamingTo m), Hilite (RenamingTo n))
=> Hilite (RenamingTo (ImportedName' m n)) where
hilite :: RenamingTo (ImportedName' m n)
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite (RenamingTo ImportedName' m n
x) = case ImportedName' m n
x of
ImportedModule n
m -> RenamingTo n -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite (n -> RenamingTo n
forall a. a -> RenamingTo a
RenamingTo n
m)
ImportedName m
n -> RenamingTo m -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a.
Hilite a =>
a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilite (m -> RenamingTo m
forall a. a -> RenamingTo a
RenamingTo m
n)
hiliteQName
:: Maybe NameKind
-> A.QName
-> Hiliter
hiliteQName :: Maybe NameKind
-> QName -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hiliteQName Maybe NameKind
mkind QName
q
| QName -> Bool
isExtendedLambdaName QName
q = ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Monoid a => a
mempty
| QName -> Bool
isAbsurdLambdaName QName
q = ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Monoid a => a
mempty
| Bool
otherwise = do
kind <- Maybe NameKind
-> (NameKind -> ReaderT HiliteEnv Identity (Maybe NameKind))
-> ReaderT HiliteEnv Identity (Maybe NameKind)
-> ReaderT HiliteEnv Identity (Maybe NameKind)
forall a b. Maybe a -> (a -> b) -> b -> b
ifJust Maybe NameKind
mkind (Maybe NameKind -> ReaderT HiliteEnv Identity (Maybe NameKind)
forall a. a -> ReaderT HiliteEnv Identity a
forall (f :: * -> *) a. Applicative f => a -> f a
pure (Maybe NameKind -> ReaderT HiliteEnv Identity (Maybe NameKind))
-> (NameKind -> Maybe NameKind)
-> NameKind
-> ReaderT HiliteEnv Identity (Maybe NameKind)
forall b c a. (b -> c) -> (a -> b) -> a -> c
. NameKind -> Maybe NameKind
forall a. a -> Maybe a
Just) (ReaderT HiliteEnv Identity (Maybe NameKind)
-> ReaderT HiliteEnv Identity (Maybe NameKind))
-> ReaderT HiliteEnv Identity (Maybe NameKind)
-> ReaderT HiliteEnv Identity (Maybe NameKind)
forall a b. (a -> b) -> a -> b
$ (HiliteEnv -> NameKinds) -> ReaderT HiliteEnv Identity NameKinds
forall r (m :: * -> *) a. MonadReader r m => (r -> a) -> m a
asks HiliteEnv -> NameKinds
hleNameKinds ReaderT HiliteEnv Identity NameKinds
-> (NameKinds -> Maybe NameKind)
-> ReaderT HiliteEnv Identity (Maybe NameKind)
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> (NameKinds -> NameKinds
forall a b. (a -> b) -> a -> b
$ QName
q)
hiliteAName q True $ nameAsp' kind
hiliteAmbiguousQName
:: Maybe NameKind
-> A.AmbiguousQName
-> Hiliter
hiliteAmbiguousQName :: Maybe NameKind
-> AmbiguousQName
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hiliteAmbiguousQName Maybe NameKind
mkind (A.AmbQ List1 QName
qs) = do
kind <- Maybe NameKind
-> (NameKind -> ReaderT HiliteEnv Identity (Maybe NameKind))
-> ReaderT HiliteEnv Identity (Maybe NameKind)
-> ReaderT HiliteEnv Identity (Maybe NameKind)
forall a b. Maybe a -> (a -> b) -> b -> b
ifJust Maybe NameKind
mkind (Maybe NameKind -> ReaderT HiliteEnv Identity (Maybe NameKind)
forall a. a -> ReaderT HiliteEnv Identity a
forall (f :: * -> *) a. Applicative f => a -> f a
pure (Maybe NameKind -> ReaderT HiliteEnv Identity (Maybe NameKind))
-> (NameKind -> Maybe NameKind)
-> NameKind
-> ReaderT HiliteEnv Identity (Maybe NameKind)
forall b c a. (b -> c) -> (a -> b) -> a -> c
. NameKind -> Maybe NameKind
forall a. a -> Maybe a
Just) (ReaderT HiliteEnv Identity (Maybe NameKind)
-> ReaderT HiliteEnv Identity (Maybe NameKind))
-> ReaderT HiliteEnv Identity (Maybe NameKind)
-> ReaderT HiliteEnv Identity (Maybe NameKind)
forall a b. (a -> b) -> a -> b
$ do
kinds <- (HiliteEnv -> NameKinds) -> ReaderT HiliteEnv Identity NameKinds
forall r (m :: * -> *) a. MonadReader r m => (r -> a) -> m a
asks HiliteEnv -> NameKinds
hleNameKinds
pure $ listToMaybe $ List1.catMaybes $ fmap kinds qs
flip foldMap qs $ \ QName
q ->
QName
-> Bool
-> (Bool -> Aspects)
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hiliteAName QName
q Bool
include ((Bool -> Aspects)
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder)
-> (Bool -> Aspects)
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a b. (a -> b) -> a -> b
$ Maybe NameKind -> Bool -> Aspects
nameAsp' Maybe NameKind
kind
where
include :: Bool
include = List1 Range -> Bool
forall a. Eq a => List1 a -> Bool
List1.allEqual (List1 Range -> Bool) -> List1 Range -> Bool
forall a b. (a -> b) -> a -> b
$ (QName -> Range) -> List1 QName -> List1 Range
forall a b. (a -> b) -> NonEmpty a -> NonEmpty b
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
fmap QName -> Range
bindingSite List1 QName
qs
hiliteBound :: A.Name -> Hiliter
hiliteBound :: Name -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hiliteBound Name
x =
[Name]
-> Name
-> Range
-> Maybe Range
-> (Bool -> Aspects)
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hiliteCName [] (Name -> Name
A.nameConcrete Name
x) Range
forall a. Range' a
noRange (Range -> Maybe Range
forall a. a -> Maybe a
Just (Range -> Maybe Range) -> Range -> Maybe Range
forall a b. (a -> b) -> a -> b
$ Name -> Range
A.nameBindingSite Name
x) ((Bool -> Aspects)
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder)
-> (Bool -> Aspects)
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a b. (a -> b) -> a -> b
$ NameKind -> Bool -> Aspects
nameAsp NameKind
Bound
hiliteInductiveConstructor :: A.AmbiguousQName -> Hiliter
hiliteInductiveConstructor :: AmbiguousQName
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hiliteInductiveConstructor = Maybe NameKind
-> AmbiguousQName
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hiliteAmbiguousQName (Maybe NameKind
-> AmbiguousQName
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder)
-> Maybe NameKind
-> AmbiguousQName
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a b. (a -> b) -> a -> b
$ NameKind -> Maybe NameKind
forall a. a -> Maybe a
Just (NameKind -> Maybe NameKind) -> NameKind -> Maybe NameKind
forall a b. (a -> b) -> a -> b
$ Induction -> NameKind
Constructor Induction
Inductive
hilitePatternSynonym :: A.AmbiguousQName -> Hiliter
hilitePatternSynonym :: AmbiguousQName
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hilitePatternSynonym = AmbiguousQName
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hiliteInductiveConstructor
hiliteProjection :: A.AmbiguousQName -> Hiliter
hiliteProjection :: AmbiguousQName
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hiliteProjection = Maybe NameKind
-> AmbiguousQName
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hiliteAmbiguousQName (NameKind -> Maybe NameKind
forall a. a -> Maybe a
Just NameKind
Field)
hiliteField :: [C.Name] -> C.Name -> Maybe Range -> Hiliter
hiliteField :: [Name]
-> Name
-> Maybe Range
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hiliteField [Name]
xs Name
x Maybe Range
bindingR = [Name]
-> Name
-> Range
-> Maybe Range
-> (Bool -> Aspects)
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hiliteCName [Name]
xs Name
x Range
forall a. Range' a
noRange Maybe Range
bindingR ((Bool -> Aspects)
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder)
-> (Bool -> Aspects)
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a b. (a -> b) -> a -> b
$ NameKind -> Bool -> Aspects
nameAsp NameKind
Field
hiliteModule :: (Bool, A.ModuleName) -> Hiliter
hiliteModule :: (Bool, ModuleName)
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hiliteModule (Bool
isTopLevelModule, A.MName []) = ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. Monoid a => a
mempty
hiliteModule (Bool
isTopLevelModule, A.MName (Name
n:[Name]
ns)) =
[Name]
-> Name
-> Range
-> Maybe Range
-> (Bool -> Aspects)
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hiliteCName
((Name -> Name) -> [Name] -> [Name]
forall a b. (a -> b) -> [a] -> [b]
map Name -> Name
A.nameConcrete [Name]
ms)
(Name -> Name
A.nameConcrete Name
m)
Range
forall a. Range' a
noRange
Maybe Range
mR
(NameKind -> Bool -> Aspects
nameAsp NameKind
Module)
where
([Name]
ms, Name
m) = Name -> [Name] -> ([Name], Name)
forall a. a -> [a] -> ([a], a)
initLast1 Name
n [Name]
ns
mR :: Maybe Range
mR = Range -> Maybe Range
forall a. a -> Maybe a
Just (Range -> Maybe Range) -> Range -> Maybe Range
forall a b. (a -> b) -> a -> b
$
Bool -> (Range -> Range) -> Range -> Range
forall b a. IsBool b => b -> (a -> a) -> a -> a
applyWhen Bool
isTopLevelModule Range -> Range
P.beginningOfFile (Range -> Range) -> Range -> Range
forall a b. (a -> b) -> a -> b
$
Name -> Range
A.nameBindingSite Name
m
hiliteCName
:: [C.Name]
-> C.Name
-> Range
-> Maybe Range
-> (Bool -> Aspects)
-> Hiliter
hiliteCName :: [Name]
-> Name
-> Range
-> Maybe Range
-> (Bool -> Aspects)
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hiliteCName [Name]
xs Name
x Range
fr Maybe Range
mR Bool -> Aspects
asp = do
env <- ReaderT HiliteEnv Identity HiliteEnv
forall r (m :: * -> *). MonadReader r m => m r
ask
if all (== Just (hleCurrentModuleName env)) moduleNames
then pure $
frFile <>
H.singleton (rToR rs) (aspects { definitionSite = mFilePos })
else mempty
where
aspects :: Aspects
aspects = Bool -> Aspects
asp (Bool -> Aspects) -> Bool -> Aspects
forall a b. (a -> b) -> a -> b
$ Name -> Bool
C.isOperator Name
x
moduleNames :: [Maybe (TopLevelModuleName' Range)]
moduleNames = (Name -> Maybe (Maybe (TopLevelModuleName' Range)))
-> [Name] -> [Maybe (TopLevelModuleName' Range)]
forall a b. (a -> Maybe b) -> [a] -> [b]
mapMaybe (Range -> Maybe (Maybe (TopLevelModuleName' Range))
P.rangeModule' (Range -> Maybe (Maybe (TopLevelModuleName' Range)))
-> (Name -> Range)
-> Name
-> Maybe (Maybe (TopLevelModuleName' Range))
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Name -> Range
forall a. HasRange a => a -> Range
getRange) (Name
x Name -> [Name] -> [Name]
forall a. a -> [a] -> [a]
: [Name]
xs)
frFile :: HighlightingInfoBuilder
frFile = Ranges -> Aspects -> HighlightingInfoBuilder
forall a m. IsBasicRangeMap a m => Ranges -> a -> m
H.singleton (Range -> Ranges
rToR Range
fr) (Aspects -> HighlightingInfoBuilder)
-> Aspects -> HighlightingInfoBuilder
forall a b. (a -> b) -> a -> b
$
Aspects
aspects { definitionSite = notHere <$> mFilePos }
rs :: Range
rs = [Name] -> Range
forall a. HasRange a => a -> Range
getRange (Name
x Name -> [Name] -> [Name]
forall a. a -> [a] -> [a]
: [Name]
xs)
notHere :: DefinitionSite -> DefinitionSite
notHere DefinitionSite
d = DefinitionSite
d { defSiteHere = False }
mFilePos :: Maybe DefinitionSite
mFilePos :: Maybe DefinitionSite
mFilePos = do
r <- Maybe Range
mR
P.Pn { P.srcFile = Strict.Just f, P.posPos = p } <- P.rStart r
mod <- P.rangeFileName f
let qualifiers = Int -> [Name] -> [Name]
forall a. Int -> [a] -> [a]
drop (TopLevelModuleName' Range -> Int
forall a. Sized a => a -> Int
size TopLevelModuleName' Range
mod) [Name]
xs
local = Bool -> (Aspect -> Bool) -> Maybe Aspect -> Bool
forall b a. b -> (a -> b) -> Maybe a -> b
maybe Bool
True Aspect -> Bool
isLocalAspect (Maybe Aspect -> Bool) -> Maybe Aspect -> Bool
forall a b. (a -> b) -> a -> b
$ Aspects -> Maybe Aspect
aspect Aspects
aspects
return $ DefinitionSite
{ defSiteModule = mod
, defSitePos = fromIntegral p
, defSiteHere = r == getRange x
, defSiteAnchor = if local || C.isNoName x || any Common.isUnderscore qualifiers
then Nothing
else Just $ prettyShow $ foldr C.Qual (C.QName x) qualifiers
}
isLocalAspect :: Aspect -> Bool
isLocalAspect :: Aspect -> Bool
isLocalAspect = \case
Name (Just NameKind
kind) Bool
_ -> NameKind -> Bool
isLocal NameKind
kind
Aspect
_ -> Bool
True
isLocal :: NameKind -> Bool
isLocal :: NameKind -> Bool
isLocal = \case
NameKind
Bound -> Bool
True
NameKind
Generalizable -> Bool
True
NameKind
Argument -> Bool
True
Constructor{} -> Bool
False
NameKind
Datatype -> Bool
False
NameKind
Field -> Bool
False
NameKind
Function -> Bool
False
NameKind
Module -> Bool
False
NameKind
Postulate -> Bool
False
NameKind
Primitive -> Bool
False
NameKind
Record -> Bool
False
NameKind
Macro -> Bool
False
hiliteAName
:: A.QName
-> Bool
-> (Bool -> Aspects)
-> Hiliter
hiliteAName :: QName
-> Bool
-> (Bool -> Aspects)
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
hiliteAName QName
x Bool
include Bool -> Aspects
asp = do
currentModule <- (HiliteEnv -> TopLevelModuleName' Range)
-> ReaderT HiliteEnv Identity (TopLevelModuleName' Range)
forall r (m :: * -> *) a. MonadReader r m => (r -> a) -> m a
asks HiliteEnv -> TopLevelModuleName' Range
hleCurrentModuleName
hiliteCName (concreteQualifier x)
(concreteBase x)
(rangeOfFixityDeclaration currentModule)
(if include then Just $ bindingSite x else Nothing)
asp
<> notationFile currentModule
where
rangeOfFixityDeclaration :: TopLevelModuleName' Range -> Range
rangeOfFixityDeclaration TopLevelModuleName' Range
currentModule =
if Range -> Maybe (TopLevelModuleName' Range)
P.rangeModule Range
r Maybe (TopLevelModuleName' Range)
-> Maybe (TopLevelModuleName' Range) -> Bool
forall a. Eq a => a -> a -> Bool
== TopLevelModuleName' Range -> Maybe (TopLevelModuleName' Range)
forall a. a -> Maybe a
Just TopLevelModuleName' Range
currentModule
then Range
r else Range
forall a. Range' a
noRange
where
r :: Range
r = Fixity' -> Range
theNameRange (Fixity' -> Range) -> Fixity' -> Range
forall a b. (a -> b) -> a -> b
$ Name -> Fixity'
A.nameFixity (Name -> Fixity') -> Name -> Fixity'
forall a b. (a -> b) -> a -> b
$ QName -> Name
A.qnameName QName
x
notationFile :: TopLevelModuleName' Range
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
notationFile TopLevelModuleName' Range
currentModule = HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. a -> ReaderT HiliteEnv Identity a
forall (f :: * -> *) a. Applicative f => a -> f a
pure (HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder)
-> HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a b. (a -> b) -> a -> b
$
if Range -> Maybe (TopLevelModuleName' Range)
P.rangeModule (Notation -> Range
forall a. HasRange a => a -> Range
getRange Notation
notation) Maybe (TopLevelModuleName' Range)
-> Maybe (TopLevelModuleName' Range) -> Bool
forall a. Eq a => a -> a -> Bool
== TopLevelModuleName' Range -> Maybe (TopLevelModuleName' Range)
forall a. a -> Maybe a
Just TopLevelModuleName' Range
currentModule
then [HighlightingInfoBuilder] -> HighlightingInfoBuilder
forall a. Monoid a => [a] -> a
mconcat ([HighlightingInfoBuilder] -> HighlightingInfoBuilder)
-> [HighlightingInfoBuilder] -> HighlightingInfoBuilder
forall a b. (a -> b) -> a -> b
$ (NotationPart -> HighlightingInfoBuilder)
-> Notation -> [HighlightingInfoBuilder]
forall a b. (a -> b) -> [a] -> [b]
map NotationPart -> HighlightingInfoBuilder
genPartFile Notation
notation
else HighlightingInfoBuilder
forall a. Monoid a => a
mempty
where
notation :: Notation
notation = Fixity' -> Notation
theNotation (Fixity' -> Notation) -> Fixity' -> Notation
forall a b. (a -> b) -> a -> b
$ Name -> Fixity'
A.nameFixity (Name -> Fixity') -> Name -> Fixity'
forall a b. (a -> b) -> a -> b
$ QName -> Name
A.qnameName QName
x
boundAspect :: Aspects
boundAspect = NameKind -> Bool -> Aspects
nameAsp NameKind
Bound Bool
False
genPartFile :: NotationPart -> HighlightingInfoBuilder
genPartFile (VarPart Range
r Ranged BoundVariablePosition
i) = [Ranges] -> Aspects -> HighlightingInfoBuilder
forall a hl.
(IsBasicRangeMap a hl, Monoid hl) =>
[Ranges] -> a -> hl
several [Range -> Ranges
rToR Range
r, Range -> Ranges
rToR (Range -> Ranges) -> Range -> Ranges
forall a b. (a -> b) -> a -> b
$ Ranged BoundVariablePosition -> Range
forall a. HasRange a => a -> Range
getRange Ranged BoundVariablePosition
i] Aspects
boundAspect
genPartFile (HolePart Range
r NamedArg (Ranged Int)
i) = [Ranges] -> Aspects -> HighlightingInfoBuilder
forall a hl.
(IsBasicRangeMap a hl, Monoid hl) =>
[Ranges] -> a -> hl
several [Range -> Ranges
rToR Range
r, Range -> Ranges
rToR (Range -> Ranges) -> Range -> Ranges
forall a b. (a -> b) -> a -> b
$ NamedArg (Ranged Int) -> Range
forall a. HasRange a => a -> Range
getRange NamedArg (Ranged Int)
i] Aspects
boundAspect
genPartFile WildPart{} = HighlightingInfoBuilder
forall a. Monoid a => a
mempty
genPartFile (IdPart RString
x) = Ranges -> Aspects -> HighlightingInfoBuilder
forall a m. IsBasicRangeMap a m => Ranges -> a -> m
H.singleton (Range -> Ranges
rToR (Range -> Ranges) -> Range -> Ranges
forall a b. (a -> b) -> a -> b
$ RString -> Range
forall a. HasRange a => a -> Range
getRange RString
x) (Bool -> Aspects
asp Bool
False)
singleAspect :: HasRange a => Aspect -> a -> Hiliter
singleAspect :: forall a.
HasRange a =>
Aspect -> a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
singleAspect Aspect
a a
x = HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. a -> ReaderT HiliteEnv Identity a
forall (f :: * -> *) a. Applicative f => a -> f a
pure (HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder)
-> HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a b. (a -> b) -> a -> b
$ Ranges -> Aspects -> HighlightingInfoBuilder
forall a m. IsBasicRangeMap a m => Ranges -> a -> m
H.singleton (Range -> Ranges
rToR (Range -> Ranges) -> Range -> Ranges
forall a b. (a -> b) -> a -> b
$ a -> Range
forall a. HasRange a => a -> Range
getRange a
x) (Aspects -> HighlightingInfoBuilder)
-> Aspects -> HighlightingInfoBuilder
forall a b. (a -> b) -> a -> b
$ Aspects
parserBased { aspect = Just a }
singleOtherAspect :: HasRange a => OtherAspect -> a -> Hiliter
singleOtherAspect :: forall a.
HasRange a =>
OtherAspect
-> a -> ReaderT HiliteEnv Identity HighlightingInfoBuilder
singleOtherAspect OtherAspect
a a
x = HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a. a -> ReaderT HiliteEnv Identity a
forall (f :: * -> *) a. Applicative f => a -> f a
pure (HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder)
-> HighlightingInfoBuilder
-> ReaderT HiliteEnv Identity HighlightingInfoBuilder
forall a b. (a -> b) -> a -> b
$ Ranges -> Aspects -> HighlightingInfoBuilder
forall a m. IsBasicRangeMap a m => Ranges -> a -> m
H.singleton (Range -> Ranges
rToR (Range -> Ranges) -> Range -> Ranges
forall a b. (a -> b) -> a -> b
$ a -> Range
forall a. HasRange a => a -> Range
getRange a
x) (Aspects -> HighlightingInfoBuilder)
-> Aspects -> HighlightingInfoBuilder
forall a b. (a -> b) -> a -> b
$ Aspects
parserBased { otherAspects = singleton a }
nameAsp' :: Maybe NameKind -> Bool -> Aspects
nameAsp' :: Maybe NameKind -> Bool -> Aspects
nameAsp' Maybe NameKind
k Bool
isOp = Aspects
parserBased { aspect = Just $ Name k isOp }
nameAsp :: NameKind -> Bool -> Aspects
nameAsp :: NameKind -> Bool -> Aspects
nameAsp = Maybe NameKind -> Bool -> Aspects
nameAsp' (Maybe NameKind -> Bool -> Aspects)
-> (NameKind -> Maybe NameKind) -> NameKind -> Bool -> Aspects
forall b c a. (b -> c) -> (a -> b) -> a -> c
. NameKind -> Maybe NameKind
forall a. a -> Maybe a
Just
concreteBase :: A.QName -> C.Name
concreteBase :: QName -> Name
concreteBase = Name -> Name
A.nameConcrete (Name -> Name) -> (QName -> Name) -> QName -> Name
forall b c a. (b -> c) -> (a -> b) -> a -> c
. QName -> Name
A.qnameName
concreteQualifier :: A.QName -> [C.Name]
concreteQualifier :: QName -> [Name]
concreteQualifier = (Name -> Name) -> [Name] -> [Name]
forall a b. (a -> b) -> [a] -> [b]
map Name -> Name
A.nameConcrete ([Name] -> [Name]) -> (QName -> [Name]) -> QName -> [Name]
forall b c a. (b -> c) -> (a -> b) -> a -> c
. ModuleName -> [Name]
A.mnameToList (ModuleName -> [Name]) -> (QName -> ModuleName) -> QName -> [Name]
forall b c a. (b -> c) -> (a -> b) -> a -> c
. QName -> ModuleName
A.qnameModule
bindingSite :: A.QName -> Range
bindingSite :: QName -> Range
bindingSite = Name -> Range
A.nameBindingSite (Name -> Range) -> (QName -> Name) -> QName -> Range
forall b c a. (b -> c) -> (a -> b) -> a -> c
. QName -> Name
A.qnameName