{-| Translation from "Agda.Syntax.Concrete" to "Agda.Syntax.Abstract".
    Involves scope analysis,
    figuring out infix operator precedences and tidying up definitions.
-}

module Agda.Syntax.Translation.ConcreteToAbstract
    ( ToAbstract(..), localToAbstract
    , concreteToAbstract_
    , concreteToAbstract
    , TopLevel(..)
    , TopLevelInfo(..)
    , topLevelModuleName
    , importPrimitives
    , checkAttributes
    ) where

import Prelude hiding ( null, (||) )

import Control.Monad        ( (>=>), (<=<), foldM, forM, forM_, zipWithM, zipWithM_ )
import Control.Applicative  ( liftA2, liftA3 )
import Control.Monad.Except ( runExceptT, MonadError(..) )
import Control.Monad.State  ( StateT, execStateT, get, put )
import Control.Monad.Trans.Maybe
import Control.Monad.Trans  ( lift )

import Data.Bifunctor
import Data.Foldable (traverse_)
import Data.Set (Set)
import Data.Map (Map)
import Data.Functor (void)
import qualified Data.List as List
import qualified Data.Set as Set
import qualified Data.Map as Map
import qualified Data.HashSet as HashSet
import Data.Maybe
import Data.Void

import Agda.Syntax.Common
import qualified Agda.Syntax.Common.Pretty as P
import Agda.Syntax.Common.Pretty (render, Pretty, pretty, prettyShow)
import Agda.Syntax.Concrete as C
import Agda.Syntax.Concrete.Attribute as CA
import Agda.Syntax.Concrete.Generic
import Agda.Syntax.Concrete.Operators
import Agda.Syntax.Concrete.Pattern
import Agda.Syntax.Abstract as A
import Agda.Syntax.Abstract.Pattern as A
  ( patternVars, checkPatternLinearity, containsAsPattern, lhsCoreApp, lhsCoreWith, noDotOrEqPattern )
import Agda.Syntax.Abstract.Pretty
import Agda.Syntax.Abstract.UsedNames
  ( allUsedNames )
import qualified Agda.Syntax.Internal as I
import Agda.Syntax.Position
import Agda.Syntax.Literal
import Agda.Syntax.Info as Info
import Agda.Syntax.Concrete.Definitions as C
import Agda.Syntax.Fixity
import Agda.Syntax.Concrete.Fixity (DoWarn(..))
import Agda.Syntax.Notation
import Agda.Syntax.Scope.Base as A
import Agda.Syntax.Scope.Monad
import Agda.Syntax.Translation.AbstractToConcrete (ToConcrete, ConOfAbs)
import Agda.Syntax.DoNotation
import Agda.Syntax.IdiomBrackets
import Agda.Syntax.TopLevelModuleName

import Agda.TypeChecking.Monad.Base hiding (ModuleInfo, MetaInfo)
import Agda.TypeChecking.Monad.Builtin
import Agda.TypeChecking.Monad.Trace (traceCall, setCurrentRange)
import Agda.TypeChecking.Monad.State hiding (topLevelModuleName)
import qualified Agda.TypeChecking.Monad.State as S
import Agda.TypeChecking.Monad.Signature (notUnderOpaque)
import Agda.TypeChecking.Monad.MetaVars (registerInteractionPoint)
import Agda.TypeChecking.Monad.Debug
import Agda.TypeChecking.Monad.Env (insideDotPattern, isInsideDotPattern, getCurrentPath)
import Agda.TypeChecking.Rules.Builtin (isUntypedBuiltin, bindUntypedBuiltin, builtinKindOfName)

import Agda.TypeChecking.Patterns.Abstract (expandPatternSynonyms)
import Agda.TypeChecking.Pretty hiding (pretty, prettyA)
import Agda.TypeChecking.Quote (quotedName)
import Agda.TypeChecking.Opacity
import Agda.TypeChecking.Warnings

import Agda.Interaction.FindFile (checkModuleName, rootNameModule, SourceFile(SourceFile))
-- import Agda.Interaction.Imports  -- for type-checking in ghci
import {-# SOURCE #-} Agda.Interaction.Imports (scopeCheckImport)
import Agda.Interaction.Options
import qualified Agda.Interaction.Options.Lenses as Lens
import Agda.Interaction.Options.Warnings

import qualified Agda.Utils.AssocList as AssocList
import Agda.Utils.Boolean   ( (||), ifThenElse )
import Agda.Utils.CallStack ( HasCallStack, withCurrentCallStack )
import Agda.Utils.Char
import Agda.Utils.Either
import Agda.Utils.FileName
import Agda.Utils.Function ( applyWhen, applyWhenM, applyUnless )
import Agda.Utils.Functor
import Agda.Utils.Lens
import Agda.Utils.List
import Agda.Utils.List1 ( List1, pattern (:|) )
import Agda.Utils.List2 ( List2, pattern List2 )
import qualified Agda.Utils.List1 as List1
import qualified Agda.Utils.Map as Map
import Agda.Utils.Maybe
import Agda.Utils.Monad
import Agda.Utils.Null
import Agda.Utils.Set1 ( Set1 )
import qualified Agda.Utils.Set1 as Set1
import Agda.Utils.Singleton
import Agda.Utils.Tuple

import Agda.Utils.Impossible ( __IMPOSSIBLE__ )
import Agda.ImpossibleTest (impossibleTest, impossibleTestReduceM)
import qualified Agda.Syntax.Common as A

{--------------------------------------------------------------------------
    Exceptions
 --------------------------------------------------------------------------}

notAnExpression :: (HasCallStack, MonadTCError m) => C.Expr -> m a
notAnExpression :: forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
Expr -> m a
notAnExpression = (Expr -> TypeError) -> HasCallStack => Expr -> m a
forall (m :: * -> *) a b.
MonadTCError m =>
(a -> TypeError) -> HasCallStack => a -> m b
locatedTypeError Expr -> TypeError
NotAnExpression

notAValidLetBinding :: (HasCallStack, MonadTCError m) => Maybe NotAValidLetBinding -> m a
notAValidLetBinding :: forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
Maybe NotAValidLetBinding -> m a
notAValidLetBinding = (Maybe NotAValidLetBinding -> TypeError)
-> HasCallStack => Maybe NotAValidLetBinding -> m a
forall (m :: * -> *) a b.
MonadTCError m =>
(a -> TypeError) -> HasCallStack => a -> m b
locatedTypeError Maybe NotAValidLetBinding -> TypeError
NotAValidLetBinding

{--------------------------------------------------------------------------
    Helpers
 --------------------------------------------------------------------------}

newtype RecordConstructorType = RecordConstructorType [C.Declaration]

instance ToAbstract RecordConstructorType where
  type AbsOfCon RecordConstructorType = A.Expr
  toAbstract :: RecordConstructorType -> ScopeM (AbsOfCon RecordConstructorType)
toAbstract (RecordConstructorType [Declaration]
ds) = [Declaration] -> ScopeM Expr
recordConstructorType [Declaration]
ds

-- | Compute the type of the record constructor (with bogus target type)
recordConstructorType :: [C.Declaration] -> ScopeM A.Expr
recordConstructorType :: [Declaration] -> ScopeM Expr
recordConstructorType [Declaration]
decls =
    -- Nicify all declarations since there might be fixity declarations after
    -- the the last field. Use NoWarn to silence fixity warnings. We'll get
    -- them again when scope checking the declarations to build the record
    -- module.
    DoWarn
-> [Declaration]
-> ([NiceDeclaration] -> ScopeM Expr)
-> ScopeM Expr
forall a.
DoWarn
-> [Declaration] -> ([NiceDeclaration] -> ScopeM a) -> ScopeM a
niceDecls DoWarn
NoWarn [Declaration]
decls (([NiceDeclaration] -> ScopeM Expr) -> ScopeM Expr)
-> ([NiceDeclaration] -> ScopeM Expr) -> ScopeM Expr
forall a b. (a -> b) -> a -> b
$ [NiceDeclaration] -> ScopeM Expr
buildType ([NiceDeclaration] -> ScopeM Expr)
-> ([NiceDeclaration] -> [NiceDeclaration])
-> [NiceDeclaration]
-> ScopeM Expr
forall b c a. (b -> c) -> (a -> b) -> a -> c
. [NiceDeclaration] -> [NiceDeclaration]
takeFields
  where
    takeFields :: [NiceDeclaration] -> [NiceDeclaration]
takeFields = (NiceDeclaration -> Bool) -> [NiceDeclaration] -> [NiceDeclaration]
forall a. (a -> Bool) -> [a] -> [a]
List.dropWhileEnd NiceDeclaration -> Bool
notField

    notField :: NiceDeclaration -> Bool
notField NiceField{} = Bool
False
    notField NiceDeclaration
_           = Bool
True

    buildType :: [C.NiceDeclaration] -> ScopeM A.Expr
      -- TODO: Telescope instead of Expr in abstract RecDef
    buildType :: [NiceDeclaration] -> ScopeM Expr
buildType [NiceDeclaration]
ds = do
      -- The constructor target type is computed in the type checker.
      -- For now, we put a dummy expression there.
      -- Andreas, 2022-10-06, issue #6165:
      -- The dummy was builtinSet, but this might not be defined yet.
      let dummy :: Expr
dummy = ExprInfo -> Literal -> Expr
A.Lit ExprInfo
forall a. Null a => a
empty (Literal -> Expr) -> Literal -> Expr
forall a b. (a -> b) -> a -> b
$ Text -> Literal
LitString Text
"TYPE"
      tel   <- [Maybe TypedBinding] -> [TypedBinding]
forall a. [Maybe a] -> [a]
catMaybes ([Maybe TypedBinding] -> [TypedBinding])
-> TCMT IO [Maybe TypedBinding] -> TCMT IO [TypedBinding]
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> (NiceDeclaration -> TCMT IO (Maybe TypedBinding))
-> [NiceDeclaration] -> TCMT IO [Maybe TypedBinding]
forall (t :: * -> *) (m :: * -> *) a b.
(Traversable t, Monad m) =>
(a -> m b) -> t a -> m (t b)
forall (m :: * -> *) a b. Monad m => (a -> m b) -> [a] -> m [b]
mapM NiceDeclaration -> TCMT IO (Maybe TypedBinding)
makeBinding [NiceDeclaration]
ds
      return $ A.mkPi (ExprRange (getRange ds)) tel dummy

    makeBinding :: C.NiceDeclaration -> ScopeM (Maybe A.TypedBinding)
    makeBinding :: NiceDeclaration -> TCMT IO (Maybe TypedBinding)
makeBinding NiceDeclaration
d = do
      let failure :: TCMT IO (Maybe TypedBinding)
failure = TypeError -> TCMT IO (Maybe TypedBinding)
forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
TypeError -> m a
typeError (TypeError -> TCMT IO (Maybe TypedBinding))
-> TypeError -> TCMT IO (Maybe TypedBinding)
forall a b. (a -> b) -> a -> b
$ NiceDeclaration -> TypeError
NotValidBeforeField NiceDeclaration
d
          r :: Range
r       = NiceDeclaration -> Range
forall a. HasRange a => a -> Range
getRange NiceDeclaration
d
          mkLet :: NiceDeclaration -> TCMT IO (Maybe TypedBinding)
mkLet NiceDeclaration
d = TypedBinding -> Maybe TypedBinding
forall a. a -> Maybe a
Just (TypedBinding -> Maybe TypedBinding)
-> (List1 LetBinding -> TypedBinding)
-> List1 LetBinding
-> Maybe TypedBinding
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Range -> List1 LetBinding -> TypedBinding
A.TLet Range
r (List1 LetBinding -> Maybe TypedBinding)
-> TCMT IO (List1 LetBinding) -> TCMT IO (Maybe TypedBinding)
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> LetDef -> ScopeM (AbsOfCon LetDef)
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract (LetDefOrigin -> NiceDeclaration -> LetDef
LetDef LetDefOrigin
RecordLetDef NiceDeclaration
d)
      Range
-> TCMT IO (Maybe TypedBinding) -> TCMT IO (Maybe TypedBinding)
forall (m :: * -> *) x a.
(MonadTrace m, HasRange x) =>
x -> m a -> m a
setCurrentRange Range
r (TCMT IO (Maybe TypedBinding) -> TCMT IO (Maybe TypedBinding))
-> TCMT IO (Maybe TypedBinding) -> TCMT IO (Maybe TypedBinding)
forall a b. (a -> b) -> a -> b
$ case NiceDeclaration
d of

        C.NiceField Range
r Access
pr IsAbstract
ab IsInstance
inst TacticAttribute
tac Name
x Arg Expr
a -> do
          fx  <- Name -> ScopeM Fixity'
getConcreteFixity Name
x
          let bv = Binder -> Named NamedName Binder
forall a name. a -> Named name a
unnamed (BoundName -> Binder
forall a. a -> Binder' a
C.mkBinder (BoundName -> Binder) -> BoundName -> Binder
forall a b. (a -> b) -> a -> b
$ (Name -> Fixity' -> BoundName
C.mkBoundName Name
x Fixity'
fx) { bnameTactic = tac }) Named NamedName Binder -> Arg Expr -> NamedArg Binder
forall a b. a -> Arg b -> Arg a
forall (f :: * -> *) a b. Functor f => a -> f b -> f a
<$ Arg Expr
a
          toAbstract $ C.TBind r (singleton bv) (unArg a)

        -- Public open is allowed and will take effect when scope checking as
        -- proper declarations.
        C.NiceOpen Range
r QName
m ImportDirective
dir -> do
          NiceDeclaration -> TCMT IO (Maybe TypedBinding)
mkLet (NiceDeclaration -> TCMT IO (Maybe TypedBinding))
-> NiceDeclaration -> TCMT IO (Maybe TypedBinding)
forall a b. (a -> b) -> a -> b
$ Range -> QName -> ImportDirective -> NiceDeclaration
C.NiceOpen Range
r QName
m ImportDirective
dir{ publicOpen = Nothing }
        C.NiceModuleMacro Range
r Access
p Erased
e Name
x ModuleApplication
modapp OpenShortHand
open ImportDirective
dir -> do
          NiceDeclaration -> TCMT IO (Maybe TypedBinding)
mkLet (NiceDeclaration -> TCMT IO (Maybe TypedBinding))
-> NiceDeclaration -> TCMT IO (Maybe TypedBinding)
forall a b. (a -> b) -> a -> b
$ Range
-> Access
-> Erased
-> Name
-> ModuleApplication
-> OpenShortHand
-> ImportDirective
-> NiceDeclaration
C.NiceModuleMacro Range
r Access
p Erased
e Name
x ModuleApplication
modapp OpenShortHand
open
                    ImportDirective
dir{ publicOpen = Nothing }

        -- Do some rudimentary matching here to get NotValidBeforeField instead
        -- of NotAValidLetDecl.
        C.NiceMutual KwRange
_ TerminationCheck
_ CoverageCheck
_ PositivityCheck
_
          [ C.FunSig Range
_ Access
_ IsAbstract
_ IsInstance
_ IsMacro
macro ArgInfo
_ TerminationCheck
_ CoverageCheck
_ Name
_ Expr
_
          , C.FunDef Range
_ [Declaration]
_ IsAbstract
abstract IsInstance
_ TerminationCheck
_ CoverageCheck
_ Name
_
             [ C.Clause Name
_ Bool
_ (C.LHS Pattern
_p [] []) (C.RHS Expr
_) WhereClause' [Declaration]
NoWhere [] ]
          ] | IsAbstract
abstract IsAbstract -> IsAbstract -> Bool
forall a. Eq a => a -> a -> Bool
/= IsAbstract
AbstractDef Bool -> Bool -> Bool
&& IsMacro
macro IsMacro -> IsMacro -> Bool
forall a. Eq a => a -> a -> Bool
/= IsMacro
MacroDef -> do
          NiceDeclaration -> TCMT IO (Maybe TypedBinding)
mkLet NiceDeclaration
d

        C.NiceLoneConstructor{} -> TCMT IO (Maybe TypedBinding)
failure
        C.NiceMutual{}        -> TCMT IO (Maybe TypedBinding)
failure
        -- TODO: some of these cases might be __IMPOSSIBLE__
        C.Axiom{}             -> TCMT IO (Maybe TypedBinding)
failure
        C.PrimitiveFunction{} -> TCMT IO (Maybe TypedBinding)
failure
        C.NiceModule{}        -> TCMT IO (Maybe TypedBinding)
failure
        C.NiceImport{}        -> TCMT IO (Maybe TypedBinding)
failure
        C.NicePragma{}        -> TCMT IO (Maybe TypedBinding)
failure
        C.NiceRecSig{}        -> TCMT IO (Maybe TypedBinding)
failure
        C.NiceDataSig{}       -> TCMT IO (Maybe TypedBinding)
failure
        C.NiceFunClause{}     -> TCMT IO (Maybe TypedBinding)
failure
        C.FunSig{}            -> TCMT IO (Maybe TypedBinding)
failure  -- Note: these are bundled with FunDef in NiceMutual
        C.FunDef{}            -> TCMT IO (Maybe TypedBinding)
failure
        C.NiceDataDef{}       -> TCMT IO (Maybe TypedBinding)
failure
        C.NiceRecDef{}        -> TCMT IO (Maybe TypedBinding)
failure
        C.NicePatternSyn{}    -> TCMT IO (Maybe TypedBinding)
failure
        C.NiceGeneralize{}    -> TCMT IO (Maybe TypedBinding)
failure
        C.NiceUnquoteDecl{}   -> TCMT IO (Maybe TypedBinding)
failure
        C.NiceUnquoteDef{}    -> TCMT IO (Maybe TypedBinding)
failure
        C.NiceUnquoteData{}   -> TCMT IO (Maybe TypedBinding)
failure
        C.NiceOpaque{}        -> TCMT IO (Maybe TypedBinding)
failure

checkModuleApplication
  :: C.ModuleApplication
  -> ModuleName
  -> C.Name
  -> C.ImportDirective
  -> ScopeM (A.ModuleApplication, ScopeCopyInfo, A.ImportDirective)

checkModuleApplication :: ModuleApplication
-> ModuleName
-> Name
-> ImportDirective
-> ScopeM (ModuleApplication, ScopeCopyInfo, ImportDirective)
checkModuleApplication (C.SectionApp Range
_ Telescope
tel QName
m [Expr]
es) ModuleName
m0 Name
x ImportDirective
dir' = do
  [Char] -> Int -> TCMT IO Doc -> TCMT IO ()
forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> TCMT IO Doc -> m ()
reportSDoc [Char]
"scope.decl" Int
70 (TCMT IO Doc -> TCMT IO ()) -> TCMT IO Doc -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ [TCMT IO Doc] -> TCMT IO Doc
forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
vcat ([TCMT IO Doc] -> TCMT IO Doc) -> [TCMT IO Doc] -> TCMT IO Doc
forall a b. (a -> b) -> a -> b
$
    [ [Char] -> TCMT IO Doc
forall (m :: * -> *). Applicative m => [Char] -> m Doc
text ([Char] -> TCMT IO Doc) -> [Char] -> TCMT IO Doc
forall a b. (a -> b) -> a -> b
$ [Char]
"scope checking ModuleApplication " [Char] -> [Char] -> [Char]
forall a. [a] -> [a] -> [a]
++ Name -> [Char]
forall a. Pretty a => a -> [Char]
prettyShow Name
x
    ]

  -- For the following, set the current module to be m0.
  ModuleName
-> ScopeM (ModuleApplication, ScopeCopyInfo, ImportDirective)
-> ScopeM (ModuleApplication, ScopeCopyInfo, ImportDirective)
forall (m :: * -> *) a.
(ReadTCState m, MonadTCState m) =>
ModuleName -> m a -> m a
withCurrentModule ModuleName
m0 (ScopeM (ModuleApplication, ScopeCopyInfo, ImportDirective)
 -> ScopeM (ModuleApplication, ScopeCopyInfo, ImportDirective))
-> ScopeM (ModuleApplication, ScopeCopyInfo, ImportDirective)
-> ScopeM (ModuleApplication, ScopeCopyInfo, ImportDirective)
forall a b. (a -> b) -> a -> b
$ do
    -- Parse the raw arguments of the module application. (See issue #1245.)
    args <- Expr -> [Expr] -> ScopeM [NamedArg Expr]
parseArguments (QName -> Expr
C.Ident QName
m) [Expr]
es
    -- Scope check the telescope (introduces bindings!).
    tel' <- catMaybes <$> toAbstract tel
    -- Scope check the old module name and the module args.
    m1    <- toAbstract $ OldModuleName m
    args' <- toAbstractCtx (ArgumentCtx PreferParen) args
    -- Copy the scope associated with m and take the parts actually imported.
    (adir, s) <- applyImportDirectiveM (C.QName x) dir' =<< getNamedScope m1
    (s', copyInfo) <- copyScope m m0 s
    -- Set the current scope to @s'@
    modifyCurrentScope $ const s'
    printScope "mod.inst" 40 "copied source module"
    reportSDoc "scope.mod.inst" 30 $ return $ pretty copyInfo
    let amodapp = [TypedBinding]
-> ModuleName -> [NamedArg Expr] -> ModuleApplication
A.SectionApp [TypedBinding]
tel' ModuleName
m1 [NamedArg Expr]
args'
    reportSDoc "scope.decl" 70 $ vcat $
      [ text $ "scope checked ModuleApplication " ++ prettyShow x
      ]
    reportSDoc "scope.decl" 70 $ vcat $
      [ nest 2 $ prettyA amodapp
      ]
    return (amodapp, copyInfo, adir)

checkModuleApplication (C.RecordModuleInstance Range
_ QName
recN) ModuleName
m0 Name
x ImportDirective
dir' =
  ModuleName
-> ScopeM (ModuleApplication, ScopeCopyInfo, ImportDirective)
-> ScopeM (ModuleApplication, ScopeCopyInfo, ImportDirective)
forall (m :: * -> *) a.
(ReadTCState m, MonadTCState m) =>
ModuleName -> m a -> m a
withCurrentModule ModuleName
m0 (ScopeM (ModuleApplication, ScopeCopyInfo, ImportDirective)
 -> ScopeM (ModuleApplication, ScopeCopyInfo, ImportDirective))
-> ScopeM (ModuleApplication, ScopeCopyInfo, ImportDirective)
-> ScopeM (ModuleApplication, ScopeCopyInfo, ImportDirective)
forall a b. (a -> b) -> a -> b
$ do
    m1 <- OldModuleName -> ScopeM (AbsOfCon OldModuleName)
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract (OldModuleName -> ScopeM (AbsOfCon OldModuleName))
-> OldModuleName -> ScopeM (AbsOfCon OldModuleName)
forall a b. (a -> b) -> a -> b
$ QName -> OldModuleName
OldModuleName QName
recN
    s <- getNamedScope m1
    (adir, s) <- applyImportDirectiveM recN dir' s
    (s', copyInfo) <- copyScope recN m0 s
    modifyCurrentScope $ const s'

    printScope "mod.inst" 40 "copied record module"
    return (A.RecordModuleInstance m1, copyInfo, adir)

-- | @checkModuleMacro mkApply range access concreteName modapp open dir@
--
--   Preserves local variables.

checkModuleMacro
  :: (ToConcrete a, Pretty (ConOfAbs a))
  => (ModuleInfo
      -> Erased
      -> ModuleName
      -> A.ModuleApplication
      -> ScopeCopyInfo
      -> A.ImportDirective
      -> a)
  -> OpenKind
  -> Range
  -> Access
  -> Erased
  -> C.Name
  -> C.ModuleApplication
  -> OpenShortHand
  -> C.ImportDirective
  -> ScopeM a
checkModuleMacro :: forall a.
(ToConcrete a, Pretty (ConOfAbs a)) =>
(ModuleInfo
 -> Erased
 -> ModuleName
 -> ModuleApplication
 -> ScopeCopyInfo
 -> ImportDirective
 -> a)
-> OpenKind
-> Range
-> Access
-> Erased
-> Name
-> ModuleApplication
-> OpenShortHand
-> ImportDirective
-> ScopeM a
checkModuleMacro ModuleInfo
-> Erased
-> ModuleName
-> ModuleApplication
-> ScopeCopyInfo
-> ImportDirective
-> a
apply OpenKind
kind Range
r Access
p Erased
e Name
x ModuleApplication
modapp OpenShortHand
open ImportDirective
dir = do
    [Char] -> Int -> TCMT IO Doc -> TCMT IO ()
forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> TCMT IO Doc -> m ()
reportSDoc [Char]
"scope.decl" Int
70 (TCMT IO Doc -> TCMT IO ()) -> TCMT IO Doc -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ [TCMT IO Doc] -> TCMT IO Doc
forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
vcat ([TCMT IO Doc] -> TCMT IO Doc) -> [TCMT IO Doc] -> TCMT IO Doc
forall a b. (a -> b) -> a -> b
$
      [ [Char] -> TCMT IO Doc
forall (m :: * -> *). Applicative m => [Char] -> m Doc
text ([Char] -> TCMT IO Doc) -> [Char] -> TCMT IO Doc
forall a b. (a -> b) -> a -> b
$ [Char]
"scope checking ModuleMacro " [Char] -> [Char] -> [Char]
forall a. [a] -> [a] -> [a]
++ Name -> [Char]
forall a. Pretty a => a -> [Char]
prettyShow Name
x
      ]
    dir <- OpenShortHand -> ImportDirective -> ScopeM ImportDirective
notPublicWithoutOpen OpenShortHand
open ImportDirective
dir

    m0 <- toAbstract (NewModuleName x)
    reportSDoc "scope.decl" 90 $ "NewModuleName: m0 =" <+> prettyA m0

    printScope "mod.inst" 40 "module macro"

    -- If we're opening a /named/ module, the import directive is
    -- applied to the "open", otherwise to the module itself. However,
    -- "public" is always applied to the "open".
    let (moduleDir, openDir) = case (open, isNoName x) of
          (OpenShortHand
DoOpen,   Bool
False) -> (ImportDirective
forall n m. ImportDirective' n m
defaultImportDir, ImportDirective
dir)
          (OpenShortHand
DoOpen,   Bool
True)  -> ( ImportDirective
dir { publicOpen = Nothing }
                               , ImportDirective
forall n m. ImportDirective' n m
defaultImportDir { publicOpen = publicOpen dir }
                               )
          (OpenShortHand
DontOpen, Bool
_)     -> (ImportDirective
dir, ImportDirective
forall n m. ImportDirective' n m
defaultImportDir)

    -- Restore the locals after module application has been checked.
    (modapp', copyInfo, adir') <- withLocalVars $ checkModuleApplication modapp m0 x moduleDir
    printScope "mod.inst.app" 40 "checkModuleMacro, after checkModuleApplication"

    reportSDoc "scope.decl" 90 $ "after mod app: trying to print m0 ..."
    reportSDoc "scope.decl" 90 $ "after mod app: m0 =" <+> prettyA m0

    bindModule p x m0
    reportSDoc "scope.decl" 90 $ "after bindMod: m0 =" <+> prettyA m0

    printScope "mod.inst.copy.after" 40 "after copying"

    -- Open the module if DoOpen.
    -- Andreas, 2014-09-02: @openModule@ might shadow some locals!
    adir <- case open of
      OpenShortHand
DontOpen -> ImportDirective -> TCMT IO ImportDirective
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return ImportDirective
adir'
      OpenShortHand
DoOpen   -> do
        adir'' <- OpenKind
-> Maybe ModuleName
-> QName
-> ImportDirective
-> TCMT IO ImportDirective
openModule OpenKind
kind (ModuleName -> Maybe ModuleName
forall a. a -> Maybe a
Just ModuleName
m0) (Name -> QName
C.QName Name
x) ImportDirective
openDir
        -- Andreas, 2020-05-14, issue #4656
        -- Keep the more meaningful import directive for highlighting
        -- (the other one is a defaultImportDir).
        return $ if isNoName x then adir' else adir''

    printScope "mod.inst" 40 $ show open
    reportSDoc "scope.decl" 90 $ "after open   : m0 =" <+> prettyA m0

    stripNoNames
    printScope "mod.inst.strip" 30 $ "after stripping"
    reportSDoc "scope.decl" 90 $ "after stripNo: m0 =" <+> prettyA m0

    let m      = ModuleName
m0 ModuleName -> List1 Name -> ModuleName
`withRangesOf` Name -> List1 Name
forall el coll. Singleton el coll => el -> coll
singleton Name
x
        adecl  = ModuleInfo
-> Erased
-> ModuleName
-> ModuleApplication
-> ScopeCopyInfo
-> ImportDirective
-> a
apply ModuleInfo
info Erased
e ModuleName
m ModuleApplication
modapp' ScopeCopyInfo
copyInfo ImportDirective
adir

    reportSDoc "scope.decl" 70 $ vcat $
      [ text $ "scope checked ModuleMacro " ++ prettyShow x
      ]
    reportSLn  "scope.decl" 90 $ "info    = " ++ show info
    reportSLn  "scope.decl" 90 $ "m       = " ++ prettyShow m
    reportSLn  "scope.decl" 90 $ "modapp' = " ++ show modapp'
    reportSDoc "scope.decl" 90 $ return $ pretty copyInfo
    reportSDoc "scope.decl" 70 $ nest 2 $ prettyA adecl
    return adecl
  where
    info :: ModuleInfo
info = ModuleInfo
             { minfoRange :: Range
minfoRange  = Range
r
             , minfoAsName :: Maybe Name
minfoAsName = Maybe Name
forall a. Maybe a
Nothing
             , minfoAsTo :: Range
minfoAsTo   = ImportDirective -> Range
renamingRange ImportDirective
dir
             , minfoOpenShort :: Maybe OpenShortHand
minfoOpenShort = OpenShortHand -> Maybe OpenShortHand
forall a. a -> Maybe a
Just OpenShortHand
open
             , minfoDirective :: Maybe ImportDirective
minfoDirective = ImportDirective -> Maybe ImportDirective
forall a. a -> Maybe a
Just ImportDirective
dir
             }

-- | The @public@ keyword must only be used together with @open@.

notPublicWithoutOpen :: OpenShortHand -> C.ImportDirective -> ScopeM C.ImportDirective
notPublicWithoutOpen :: OpenShortHand -> ImportDirective -> ScopeM ImportDirective
notPublicWithoutOpen OpenShortHand
DoOpen   ImportDirective
dir = ImportDirective -> ScopeM ImportDirective
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return ImportDirective
dir
notPublicWithoutOpen OpenShortHand
DontOpen ImportDirective
dir = do
  Maybe KwRange -> (KwRange -> TCMT IO ()) -> TCMT IO ()
forall (m :: * -> *) a. Monad m => Maybe a -> (a -> m ()) -> m ()
whenJust (ImportDirective -> Maybe KwRange
forall n m. ImportDirective' n m -> Maybe KwRange
publicOpen ImportDirective
dir) ((KwRange -> TCMT IO ()) -> TCMT IO ())
-> (KwRange -> TCMT IO ()) -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ \ KwRange
r ->
    KwRange -> TCMT IO () -> TCMT IO ()
forall (m :: * -> *) x a.
(MonadTrace m, HasRange x) =>
x -> m a -> m a
setCurrentRange KwRange
r (TCMT IO () -> TCMT IO ()) -> TCMT IO () -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ Warning -> TCMT IO ()
forall (m :: * -> *).
(HasCallStack, MonadWarning m) =>
Warning -> m ()
warning Warning
UselessPublic
  ImportDirective -> ScopeM ImportDirective
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return (ImportDirective -> ScopeM ImportDirective)
-> ImportDirective -> ScopeM ImportDirective
forall a b. (a -> b) -> a -> b
$ ImportDirective
dir { publicOpen = Nothing }

-- | Computes the range of all the \"to\" keywords used in a renaming
-- directive.

renamingRange :: C.ImportDirective -> Range
renamingRange :: ImportDirective -> Range
renamingRange = [Range] -> Range
forall a. HasRange a => a -> Range
getRange ([Range] -> Range)
-> (ImportDirective -> [Range]) -> ImportDirective -> Range
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (Renaming' Name Name -> Range) -> [Renaming' Name Name] -> [Range]
forall a b. (a -> b) -> [a] -> [b]
map Renaming' Name Name -> Range
forall n m. Renaming' n m -> Range
renToRange ([Renaming' Name Name] -> [Range])
-> (ImportDirective -> [Renaming' Name Name])
-> ImportDirective
-> [Range]
forall b c a. (b -> c) -> (a -> b) -> a -> c
. ImportDirective -> [Renaming' Name Name]
forall n m. ImportDirective' n m -> RenamingDirective' n m
impRenaming

-- | Scope check a 'NiceOpen'.
checkOpen
  :: Range                -- ^ Range of @open@ statement.
  -> Maybe A.ModuleName   -- ^ Resolution of concrete module name (if already resolved).
  -> C.QName              -- ^ Module to open.
  -> C.ImportDirective    -- ^ Scope modifier.
  -> ScopeM (ModuleInfo, A.ModuleName, A.ImportDirective) -- ^ Arguments of 'A.Open'
checkOpen :: Range
-> Maybe ModuleName
-> QName
-> ImportDirective
-> ScopeM (ModuleInfo, ModuleName, ImportDirective)
checkOpen Range
r Maybe ModuleName
mam QName
x ImportDirective
dir = do
  [Char] -> Int -> TCMT IO Doc -> TCMT IO ()
forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> TCMT IO Doc -> m ()
reportSDoc [Char]
"scope.decl" Int
70 (TCMT IO Doc -> TCMT IO ()) -> TCMT IO Doc -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ do
    cm <- TCMT IO ModuleName
forall (m :: * -> *). ReadTCState m => m ModuleName
getCurrentModule
    vcat $
      [ text   "scope checking NiceOpen " <> return (pretty x)
      , text   "  getCurrentModule       = " <> prettyA cm
      , text $ "  getCurrentModule (raw) = " ++ show cm
      , text $ "  C.ImportDirective      = " ++ prettyShow dir
      ]
  -- Andreas, 2017-01-01, issue #2377: warn about useless `public`
  Maybe KwRange -> (KwRange -> TCMT IO ()) -> TCMT IO ()
forall (m :: * -> *) a. Monad m => Maybe a -> (a -> m ()) -> m ()
whenJust (ImportDirective -> Maybe KwRange
forall n m. ImportDirective' n m -> Maybe KwRange
publicOpen ImportDirective
dir) ((KwRange -> TCMT IO ()) -> TCMT IO ())
-> (KwRange -> TCMT IO ()) -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ \ KwRange
r -> do
    TCMT IO Bool -> TCMT IO () -> TCMT IO ()
forall (m :: * -> *). Monad m => m Bool -> m () -> m ()
whenM ((ModuleName
A.noModuleName ModuleName -> ModuleName -> Bool
forall a. Eq a => a -> a -> Bool
==) (ModuleName -> Bool) -> TCMT IO ModuleName -> TCMT IO Bool
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> TCMT IO ModuleName
forall (m :: * -> *). ReadTCState m => m ModuleName
getCurrentModule) (TCMT IO () -> TCMT IO ()) -> TCMT IO () -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ do
      KwRange -> TCMT IO () -> TCMT IO ()
forall (m :: * -> *) x a.
(MonadTrace m, HasRange x) =>
x -> m a -> m a
setCurrentRange KwRange
r (TCMT IO () -> TCMT IO ()) -> TCMT IO () -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ Warning -> TCMT IO ()
forall (m :: * -> *).
(HasCallStack, MonadWarning m) =>
Warning -> m ()
warning Warning
UselessPublic

  m <- Maybe ModuleName
-> TCMT IO ModuleName
-> (ModuleName -> TCMT IO ModuleName)
-> TCMT IO ModuleName
forall a b. Maybe a -> b -> (a -> b) -> b
caseMaybe Maybe ModuleName
mam (OldModuleName -> ScopeM (AbsOfCon OldModuleName)
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract (QName -> OldModuleName
OldModuleName QName
x)) ModuleName -> TCMT IO ModuleName
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return
  printScope "open" 40 $ "opening " ++ prettyShow x
  adir <- openModule TopOpenModule (Just m) x dir
  printScope "open" 40 $ "result:"
  let minfo = ModuleInfo
        { minfoRange :: Range
minfoRange     = Range
r
        , minfoAsName :: Maybe Name
minfoAsName    = Maybe Name
forall a. Maybe a
Nothing
        , minfoAsTo :: Range
minfoAsTo      = ImportDirective -> Range
renamingRange ImportDirective
dir
        , minfoOpenShort :: Maybe OpenShortHand
minfoOpenShort = Maybe OpenShortHand
forall a. Maybe a
Nothing
        , minfoDirective :: Maybe ImportDirective
minfoDirective = ImportDirective -> Maybe ImportDirective
forall a. a -> Maybe a
Just ImportDirective
dir
        }
  let adecls = [ModuleInfo -> ModuleName -> ImportDirective -> Declaration
A.Open ModuleInfo
minfo ModuleName
m ImportDirective
adir]
  reportSDoc "scope.decl" 70 $ vcat $
    text ( "scope checked NiceOpen " ++ prettyShow x
         ) : map (nest 2 . prettyA) adecls
  return (minfo, m, adir)

-- | Check a literal, issuing an error warning for bad literals.
checkLiteral :: Literal -> ScopeM ()
checkLiteral :: Literal -> TCMT IO ()
checkLiteral = \case
  LitChar   Char
c   -> Bool -> TCMT IO () -> TCMT IO ()
forall b (m :: * -> *). (IsBool b, Monad m) => b -> m () -> m ()
when (Char -> Bool
isSurrogateCodePoint Char
c) (TCMT IO () -> TCMT IO ()) -> TCMT IO () -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ Warning -> TCMT IO ()
forall (m :: * -> *).
(HasCallStack, MonadWarning m) =>
Warning -> m ()
warning (Warning -> TCMT IO ()) -> Warning -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ Char -> Warning
InvalidCharacterLiteral Char
c
  LitNat    Integer
_   -> () -> TCMT IO ()
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return ()
  LitWord64 Word64
_   -> () -> TCMT IO ()
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return ()
  LitFloat  Double
_   -> () -> TCMT IO ()
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return ()
  LitString Text
_   -> () -> TCMT IO ()
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return ()
  LitQName  QName
_   -> () -> TCMT IO ()
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return ()
  LitMeta   TopLevelModuleName' Range
_ MetaId
_ -> () -> TCMT IO ()
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return ()

-- | Check a `record where` expression.
checkRecordWhere :: Range
                 -> [C.Declaration]
                 -> (RecInfo -> A.Assigns -> A.Expr)
                 -> ScopeM A.Expr
checkRecordWhere :: Range
-> [Declaration] -> (RecInfo -> Assigns -> Expr) -> ScopeM Expr
checkRecordWhere Range
r [Declaration]
ds0 RecInfo -> Assigns -> Expr
mkRec = do
  let i :: ExprInfo
i  = Range -> ExprInfo
ExprRange Range
r
      ri :: RecInfo
ri = Range -> RecInfo
recInfoWhere Range
r
  -- We need to rewrite any opens or module applications opened in this
  -- scope into a non-opening declaration and explicit bindings, so that
  -- references inside this body refer unambiguously to things opened
  -- here. Conveniently, this also homogenizes the things that need to be
  -- made fields in the final `A.Rec`. If we get #3801, perhaps this
  -- rewrite won't be necessary any more, but any `using` or `renaming`
  -- clauses will still have to be made into field assignments somehow.
  ds0' <- (Declaration -> TCMT IO [Declaration])
-> [Declaration] -> TCMT IO [Declaration]
forall (t :: * -> *) (m :: * -> *) b a.
(Foldable t, Applicative m, Monoid b) =>
(a -> m b) -> t a -> m b
mapM' Declaration -> TCMT IO [Declaration]
rewriteConcreteOpens [Declaration]
ds0
  List1.ifNull ds0' (return $ A.Rec ri []) $ \ List1 Declaration
ds -> do
    LetDefs -> (AbsOfCon LetDefs -> ScopeM Expr) -> ScopeM Expr
forall c b.
ToAbstract c =>
c -> (AbsOfCon c -> ScopeM b) -> ScopeM b
localToAbstract (LetDefOrigin -> List1 Declaration -> LetDefs
LetDefs LetDefOrigin
RecordWhereLetDef List1 Declaration
ds) ((AbsOfCon LetDefs -> ScopeM Expr) -> ScopeM Expr)
-> (AbsOfCon LetDefs -> ScopeM Expr) -> ScopeM Expr
forall a b. (a -> b) -> a -> b
$ \ AbsOfCon LetDefs
ds' -> do
      names <- [LetBinding] -> (LetBinding -> TCMT IO [Name]) -> TCMT IO [Name]
forall (t :: * -> *) (m :: * -> *) b a.
(Foldable t, Applicative m, Monoid b) =>
t a -> (a -> m b) -> m b
forM' [LetBinding]
AbsOfCon LetDefs
ds' ((LetBinding -> TCMT IO [Name]) -> TCMT IO [Name])
-> (LetBinding -> TCMT IO [Name]) -> TCMT IO [Name]
forall a b. (a -> b) -> a -> b
$ \case
        A.LetBind LetInfo
_ ArgInfo
_ (A.BindName Name
name) Expr
_ Expr
_ -> [Name] -> TCMT IO [Name]
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return [Name
name]
        A.LetPatBind LetInfo
_ Pattern
p Expr
_ -> Pattern -> TCMT IO [Name]
forall {e}. Pattern' e -> TCMT IO [Name]
go Pattern
p
          where
            go :: Pattern' e -> TCMT IO [Name]
go = \case
              A.VarP (A.BindName Name
name)    -> [Name] -> TCMT IO [Name]
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return [Name
name]
              A.ConP ConPatInfo
_ AmbiguousQName
_ NAPs e
args             -> (NamedArg (Pattern' e) -> TCMT IO [Name])
-> NAPs e -> TCMT IO [Name]
forall (t :: * -> *) (m :: * -> *) b a.
(Foldable t, Applicative m, Monoid b) =>
(a -> m b) -> t a -> m b
mapM' (Pattern' e -> TCMT IO [Name]
go (Pattern' e -> TCMT IO [Name])
-> (NamedArg (Pattern' e) -> Pattern' e)
-> NamedArg (Pattern' e)
-> TCMT IO [Name]
forall b c a. (b -> c) -> (a -> b) -> a -> c
. NamedArg (Pattern' e) -> Pattern' e
forall a. NamedArg a -> a
namedArg) NAPs e
args
              A.RecP ConPatInfo
_ [FieldAssignment' (Pattern' e)]
fs                 -> (FieldAssignment' (Pattern' e) -> TCMT IO [Name])
-> [FieldAssignment' (Pattern' e)] -> TCMT IO [Name]
forall (t :: * -> *) (m :: * -> *) b a.
(Foldable t, Applicative m, Monoid b) =>
(a -> m b) -> t a -> m b
mapM' (Pattern' e -> TCMT IO [Name]
go (Pattern' e -> TCMT IO [Name])
-> (FieldAssignment' (Pattern' e) -> Pattern' e)
-> FieldAssignment' (Pattern' e)
-> TCMT IO [Name]
forall b c a. (b -> c) -> (a -> b) -> a -> c
. FieldAssignment' (Pattern' e) -> Pattern' e
forall a. FieldAssignment' a -> a
_exprFieldA) [FieldAssignment' (Pattern' e)]
fs
              A.WildP PatInfo
_                   -> [Name] -> TCMT IO [Name]
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return []
              A.AsP PatInfo
_ (A.BindName Name
name) Pattern' e
p -> (Name
name Name -> [Name] -> [Name]
forall a. a -> [a] -> [a]
:) ([Name] -> [Name]) -> TCMT IO [Name] -> TCMT IO [Name]
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> Pattern' e -> TCMT IO [Name]
go Pattern' e
p
              A.PatternSynP PatInfo
_ AmbiguousQName
_ NAPs e
args      -> (NamedArg (Pattern' e) -> TCMT IO [Name])
-> NAPs e -> TCMT IO [Name]
forall (t :: * -> *) (m :: * -> *) b a.
(Foldable t, Applicative m, Monoid b) =>
(a -> m b) -> t a -> m b
mapM' (Pattern' e -> TCMT IO [Name]
go (Pattern' e -> TCMT IO [Name])
-> (NamedArg (Pattern' e) -> Pattern' e)
-> NamedArg (Pattern' e)
-> TCMT IO [Name]
forall b c a. (b -> c) -> (a -> b) -> a -> c
. NamedArg (Pattern' e) -> Pattern' e
forall a. NamedArg a -> a
namedArg) NAPs e
args
              -- We have run checkValidLetPattern at this point
              A.ProjP{}   -> TCMT IO [Name]
forall a. HasCallStack => a
__IMPOSSIBLE__
              A.DefP{}    -> TCMT IO [Name]
forall a. HasCallStack => a
__IMPOSSIBLE__
              A.DotP{}    -> TCMT IO [Name]
forall a. HasCallStack => a
__IMPOSSIBLE__
              A.AbsurdP{} -> TCMT IO [Name]
forall a. HasCallStack => a
__IMPOSSIBLE__
              A.LitP{}    -> TCMT IO [Name]
forall a. HasCallStack => a
__IMPOSSIBLE__
              A.EqualP{}  -> TCMT IO [Name]
forall a. HasCallStack => a
__IMPOSSIBLE__
              A.WithP{}   -> TCMT IO [Name]
forall a. HasCallStack => a
__IMPOSSIBLE__
        LetBinding
_ -> [Name] -> TCMT IO [Name]
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return []
      let fs = [Name -> Expr -> FieldAssignment' Expr
forall a. Name -> a -> FieldAssignment' a
C.FieldAssignment (Name -> Name
nameCanonical Name
n) (Name -> Expr
A.Var Name
n) | Name
n <- [Name]
names]
      return $ A.mkLet i ds' (mkRec ri fs)
  where
    rewriteConcreteOpens :: C.Declaration -> ScopeM [C.Declaration]
    rewriteConcreteOpens :: Declaration -> TCMT IO [Declaration]
rewriteConcreteOpens = \case
      C.Open Range
_ QName
m ImportDirective
dir -> QName -> ImportDirective -> TCMT IO [Declaration]
handleOpen QName
m ImportDirective
dir
      C.ModuleMacro Range
r Erased
e Name
m ModuleApplication
mapp OpenShortHand
DoOpen ImportDirective
dir -> do
        m' <- case Name
m of
          NoName Range
r NameId
_ -> Range -> Int -> [Char] -> TCMT IO Name
freshConcreteName Range
r Int
0 [Char]
".modulemacro"
          Name
n -> Name -> TCMT IO Name
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return Name
n
        ds <- handleOpen (C.QName m') dir
        return $ C.ModuleMacro r e m' mapp DontOpen defaultImportDir : ds

      C.Private KwRange
r Origin
o [Declaration]
ds -> Declaration -> [Declaration]
forall el coll. Singleton el coll => el -> coll
singleton (Declaration -> [Declaration])
-> ([Declaration] -> Declaration) -> [Declaration] -> [Declaration]
forall b c a. (b -> c) -> (a -> b) -> a -> c
. KwRange -> Origin -> [Declaration] -> Declaration
C.Private KwRange
r Origin
o ([Declaration] -> [Declaration])
-> TCMT IO [Declaration] -> TCMT IO [Declaration]
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> (Declaration -> TCMT IO [Declaration])
-> [Declaration] -> TCMT IO [Declaration]
forall (t :: * -> *) (m :: * -> *) b a.
(Foldable t, Applicative m, Monoid b) =>
(a -> m b) -> t a -> m b
mapM' Declaration -> TCMT IO [Declaration]
rewriteConcreteOpens [Declaration]
ds
      Declaration
d -> [Declaration] -> TCMT IO [Declaration]
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return [Declaration
d]

    handleOpen :: C.QName -> C.ImportDirective -> ScopeM [C.Declaration]
    handleOpen :: QName -> ImportDirective -> TCMT IO [Declaration]
handleOpen QName
m ImportDirective{ Using' Name Name
using :: Using' Name Name
using :: forall n m. ImportDirective' n m -> Using' n m
using, [Renaming' Name Name]
impRenaming :: forall n m. ImportDirective' n m -> RenamingDirective' n m
impRenaming :: [Renaming' Name Name]
impRenaming } = case Using' Name Name
using of
      Using' Name Name
UseEverything ->
        TypeError -> TCMT IO [Declaration]
forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
TypeError -> m a
typeError TypeError
OpenEverythingInRecordWhere
      Using [ImportedName' Name Name]
usingNames ->
        [Declaration] -> TCMT IO [Declaration]
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return ([Declaration] -> TCMT IO [Declaration])
-> ([(Name, Name)] -> [Declaration])
-> [(Name, Name)]
-> TCMT IO [Declaration]
forall b c a. (b -> c) -> (a -> b) -> a -> c
. ((Name, Name) -> Declaration) -> [(Name, Name)] -> [Declaration]
forall a b. (a -> b) -> [a] -> [b]
map (Name, Name) -> Declaration
mkDef ([(Name, Name)] -> TCMT IO [Declaration])
-> [(Name, Name)] -> TCMT IO [Declaration]
forall a b. (a -> b) -> a -> b
$
          [(Name
name, Name
name) | ImportedName Name
name <- [ImportedName' Name Name]
usingNames] [(Name, Name)] -> [(Name, Name)] -> [(Name, Name)]
forall a. [a] -> [a] -> [a]
++
          [(Name
name', Name
name) | Renaming{ renFrom :: forall n m. Renaming' n m -> ImportedName' n m
renFrom = ImportedName Name
name, renTo :: forall n m. Renaming' n m -> ImportedName' n m
renTo = ImportedName Name
name' } <- [Renaming' Name Name]
impRenaming]
        where
          mkDef :: (Name, Name) -> Declaration
mkDef (Name
l, Name
r) = LHS
-> RHS' Expr -> WhereClause' [Declaration] -> Bool -> Declaration
C.FunClause
            (Pattern -> [RewriteEqn] -> [WithExpr] -> LHS
C.LHS (Bool -> QName -> Pattern
C.IdentP Bool
True (Name -> QName
C.QName Name
l)) [] [])
            (Expr -> RHS' Expr
forall e. e -> RHS' e
C.RHS (QName -> Expr
C.Ident (QName -> Name -> QName
C.qualify QName
m Name
r)))
            WhereClause' [Declaration]
forall decls. WhereClause' decls
NoWhere
            Bool
False

{--------------------------------------------------------------------------
    Translation
 --------------------------------------------------------------------------}

concreteToAbstract_ :: ToAbstract c => c -> ScopeM (AbsOfCon c)
concreteToAbstract_ :: forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
concreteToAbstract_ = c -> ScopeM (AbsOfCon c)
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract

concreteToAbstract :: ToAbstract c => ScopeInfo -> c -> ScopeM (AbsOfCon c)
concreteToAbstract :: forall c. ToAbstract c => ScopeInfo -> c -> ScopeM (AbsOfCon c)
concreteToAbstract ScopeInfo
scope c
x = ScopeInfo -> TCMT IO (AbsOfCon c) -> TCMT IO (AbsOfCon c)
forall (m :: * -> *) a. ReadTCState m => ScopeInfo -> m a -> m a
withScope_ ScopeInfo
scope (c -> TCMT IO (AbsOfCon c)
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract c
x)

-- | Things that can be translated to abstract syntax are instances of this
--   class.
class ToAbstract c where
    type AbsOfCon c
    toAbstract :: c -> ScopeM (AbsOfCon c)

-- | This function should be used instead of 'toAbstract' for things that need
--   to keep track of precedences to make sure that we don't forget about it.
toAbstractCtx :: ToAbstract c => Precedence -> c-> ScopeM (AbsOfCon c)
toAbstractCtx :: forall c. ToAbstract c => Precedence -> c -> ScopeM (AbsOfCon c)
toAbstractCtx Precedence
ctx c
c = Precedence -> TCMT IO (AbsOfCon c) -> TCMT IO (AbsOfCon c)
forall (m :: * -> *) a. ReadTCState m => Precedence -> m a -> m a
withContextPrecedence Precedence
ctx (TCMT IO (AbsOfCon c) -> TCMT IO (AbsOfCon c))
-> TCMT IO (AbsOfCon c) -> TCMT IO (AbsOfCon c)
forall a b. (a -> b) -> a -> b
$ c -> TCMT IO (AbsOfCon c)
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract c
c

--UNUSED Liang-Ting Chen 2019-07-16
--toAbstractTopCtx :: ToAbstract c a => c -> ScopeM a
--toAbstractTopCtx = toAbstractCtx TopCtx

toAbstractHiding :: (LensHiding h, ToAbstract c) => h -> c -> ScopeM (AbsOfCon c)
toAbstractHiding :: forall h c.
(LensHiding h, ToAbstract c) =>
h -> c -> ScopeM (AbsOfCon c)
toAbstractHiding h
h | h -> Bool
forall a. LensHiding a => a -> Bool
visible h
h = c -> ScopeM (AbsOfCon c)
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract -- don't change precedence if visible
toAbstractHiding h
_             = Precedence -> c -> ScopeM (AbsOfCon c)
forall c. ToAbstract c => Precedence -> c -> ScopeM (AbsOfCon c)
toAbstractCtx Precedence
TopCtx

--UNUSED Liang-Ting Chen 2019-07-16
--setContextCPS :: Precedence -> (a -> ScopeM b) ->
--                 ((a -> ScopeM b) -> ScopeM b) -> ScopeM b
--setContextCPS p ret f = do
--  old <- useScope scopePrecedence
--  withContextPrecedence p $ f $ \ x -> setContextPrecedence old >> ret x
--
--localToAbstractCtx :: ToAbstract c =>
--                     Precedence -> c -> (AbsOfCon -> ScopeM (AbsOfCon c)) -> ScopeM (AbsOfCon c)
--localToAbstractCtx ctx c ret = setContextCPS ctx ret (localToAbstract c)

-- | This operation does not affect the scope, i.e. the original scope
--   is restored upon completion.
localToAbstract :: ToAbstract c => c -> (AbsOfCon c -> ScopeM b) -> ScopeM b
localToAbstract :: forall c b.
ToAbstract c =>
c -> (AbsOfCon c -> ScopeM b) -> ScopeM b
localToAbstract c
x AbsOfCon c -> ScopeM b
ret = (b, ScopeInfo) -> b
forall a b. (a, b) -> a
fst ((b, ScopeInfo) -> b) -> TCMT IO (b, ScopeInfo) -> ScopeM b
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> c -> (AbsOfCon c -> ScopeM b) -> TCMT IO (b, ScopeInfo)
forall c b.
ToAbstract c =>
c -> (AbsOfCon c -> ScopeM b) -> ScopeM (b, ScopeInfo)
localToAbstract' c
x AbsOfCon c -> ScopeM b
ret

-- | Like 'localToAbstract' but returns the scope after the completion of the
--   second argument.
localToAbstract' :: ToAbstract c => c -> (AbsOfCon c -> ScopeM b) -> ScopeM (b, ScopeInfo)
localToAbstract' :: forall c b.
ToAbstract c =>
c -> (AbsOfCon c -> ScopeM b) -> ScopeM (b, ScopeInfo)
localToAbstract' c
x AbsOfCon c -> ScopeM b
ret = do
  scope <- TCMT IO ScopeInfo
forall (m :: * -> *). ReadTCState m => m ScopeInfo
getScope
  withScope scope $ ret =<< toAbstract x

instance ToAbstract () where
  type AbsOfCon () = ()
  toAbstract :: () -> ScopeM (AbsOfCon ())
toAbstract = () -> TCMT IO ()
() -> ScopeM (AbsOfCon ())
forall a. a -> TCMT IO a
forall (f :: * -> *) a. Applicative f => a -> f a
pure

instance (ToAbstract c1, ToAbstract c2) => ToAbstract (c1, c2) where
  type AbsOfCon (c1, c2) = (AbsOfCon c1, AbsOfCon c2)
  toAbstract :: (c1, c2) -> ScopeM (AbsOfCon (c1, c2))
toAbstract (c1
x,c2
y) = (,) (AbsOfCon c1 -> AbsOfCon c2 -> (AbsOfCon c1, AbsOfCon c2))
-> TCMT IO (AbsOfCon c1)
-> TCMT IO (AbsOfCon c2 -> (AbsOfCon c1, AbsOfCon c2))
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> c1 -> TCMT IO (AbsOfCon c1)
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract c1
x TCMT IO (AbsOfCon c2 -> (AbsOfCon c1, AbsOfCon c2))
-> TCMT IO (AbsOfCon c2) -> TCMT IO (AbsOfCon c1, AbsOfCon c2)
forall a b. TCMT IO (a -> b) -> TCMT IO a -> TCMT IO b
forall (f :: * -> *) a b. Applicative f => f (a -> b) -> f a -> f b
<*> c2 -> TCMT IO (AbsOfCon c2)
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract c2
y

instance (ToAbstract c1, ToAbstract c2, ToAbstract c3) => ToAbstract (c1, c2, c3) where
  type AbsOfCon (c1, c2, c3) = (AbsOfCon c1, AbsOfCon c2, AbsOfCon c3)
  toAbstract :: (c1, c2, c3) -> ScopeM (AbsOfCon (c1, c2, c3))
toAbstract (c1
x,c2
y,c3
z) = (AbsOfCon c1, (AbsOfCon c2, AbsOfCon c3))
-> (AbsOfCon c1, AbsOfCon c2, AbsOfCon c3)
forall {a} {b} {c}. (a, (b, c)) -> (a, b, c)
flatten ((AbsOfCon c1, (AbsOfCon c2, AbsOfCon c3))
 -> (AbsOfCon c1, AbsOfCon c2, AbsOfCon c3))
-> TCMT IO (AbsOfCon c1, (AbsOfCon c2, AbsOfCon c3))
-> TCMT IO (AbsOfCon c1, AbsOfCon c2, AbsOfCon c3)
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> (c1, (c2, c3)) -> ScopeM (AbsOfCon (c1, (c2, c3)))
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract (c1
x,(c2
y,c3
z))
    where
      flatten :: (a, (b, c)) -> (a, b, c)
flatten (a
x,(b
y,c
z)) = (a
x,b
y,c
z)

instance ToAbstract c => ToAbstract [c] where
  type AbsOfCon [c] = [AbsOfCon c]
  toAbstract :: [c] -> ScopeM (AbsOfCon [c])
toAbstract = (c -> TCMT IO (AbsOfCon c)) -> [c] -> TCMT IO [AbsOfCon c]
forall (t :: * -> *) (m :: * -> *) a b.
(Traversable t, Monad m) =>
(a -> m b) -> t a -> m (t b)
forall (m :: * -> *) a b. Monad m => (a -> m b) -> [a] -> m [b]
mapM c -> TCMT IO (AbsOfCon c)
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract

instance ToAbstract c => ToAbstract (List1 c) where
  type AbsOfCon (List1 c) = List1 (AbsOfCon c)
  toAbstract :: List1 c -> ScopeM (AbsOfCon (List1 c))
toAbstract = (c -> TCMT IO (AbsOfCon c))
-> List1 c -> TCMT IO (NonEmpty (AbsOfCon c))
forall (t :: * -> *) (m :: * -> *) a b.
(Traversable t, Monad m) =>
(a -> m b) -> t a -> m (t b)
forall (m :: * -> *) a b.
Monad m =>
(a -> m b) -> NonEmpty a -> m (NonEmpty b)
mapM c -> TCMT IO (AbsOfCon c)
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract

instance (ToAbstract c1, ToAbstract c2) => ToAbstract (Either c1 c2) where
  type AbsOfCon (Either c1 c2) = Either (AbsOfCon c1) (AbsOfCon c2)
  toAbstract :: Either c1 c2 -> ScopeM (AbsOfCon (Either c1 c2))
toAbstract = (c1 -> TCMT IO (AbsOfCon c1))
-> (c2 -> TCMT IO (AbsOfCon c2))
-> Either c1 c2
-> TCMT IO (Either (AbsOfCon c1) (AbsOfCon c2))
forall (f :: * -> *) a c b d.
Functor f =>
(a -> f c) -> (b -> f d) -> Either a b -> f (Either c d)
traverseEither c1 -> TCMT IO (AbsOfCon c1)
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract c2 -> TCMT IO (AbsOfCon c2)
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract

instance ToAbstract c => ToAbstract (Maybe c) where
  type AbsOfCon (Maybe c) = Maybe (AbsOfCon c)
  toAbstract :: Maybe c -> ScopeM (AbsOfCon (Maybe c))
toAbstract = (c -> TCMT IO (AbsOfCon c))
-> Maybe c -> TCMT IO (Maybe (AbsOfCon c))
forall (t :: * -> *) (f :: * -> *) a b.
(Traversable t, Applicative f) =>
(a -> f b) -> t a -> f (t b)
forall (f :: * -> *) a b.
Applicative f =>
(a -> f b) -> Maybe a -> f (Maybe b)
traverse c -> TCMT IO (AbsOfCon c)
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract

-- Names ------------------------------------------------------------------

data NewName a = NewName
  { forall a. NewName a -> BindingSource
newBinder   :: A.BindingSource -- what kind of binder?
  , forall a. NewName a -> a
newName     :: a
  } deriving ((forall a b. (a -> b) -> NewName a -> NewName b)
-> (forall a b. a -> NewName b -> NewName a) -> Functor NewName
forall a b. a -> NewName b -> NewName a
forall a b. (a -> b) -> NewName a -> NewName b
forall (f :: * -> *).
(forall a b. (a -> b) -> f a -> f b)
-> (forall a b. a -> f b -> f a) -> Functor f
$cfmap :: forall a b. (a -> b) -> NewName a -> NewName b
fmap :: forall a b. (a -> b) -> NewName a -> NewName b
$c<$ :: forall a b. a -> NewName b -> NewName a
<$ :: forall a b. a -> NewName b -> NewName a
Functor)

data OldQName = OldQName
  C.QName
    -- ^ Concrete name to be resolved.
  (Maybe (Set1 A.Name))
    -- ^ If a set is given, then the first name must
    --   correspond to one of the names in the set.

-- | We sometimes do not want to fail hard if the name is not actually
--   in scope because we have a strategy to recover from this problem
--   (e.g. drop the offending COMPILE pragma)
data MaybeOldQName = MaybeOldQName OldQName

-- | Wrapper for a concrete name that we already bound to an 'A.Def'.
--
newtype OldName a = OldName a

-- | Wrapper to resolve a name to a 'ResolvedName' (rather than an 'A.Expr').
data ResolveQName = ResolveQName C.QName

-- | Wrapper to resolve a name in a pattern.
data PatName = PatName
  C.QName
    -- ^ Concrete name to be resolved in a pattern.
  (Maybe (Set1 A.Name))
    -- ^ If a set is given, then the first name must correspond to one
    --   of the names in the set.
  Hiding
    -- ^ If pattern variable is hidden, its status is indicated in 'Hiding'.
  DisplayLHS
    -- ^ If we parse the lhs of a 'DisplayPragma',
    --   names of arbitrary definitions count as constructors.

instance ToAbstract (NewName C.Name) where
  type AbsOfCon (NewName C.Name) = A.Name
  toAbstract :: NewName Name -> ScopeM (AbsOfCon (NewName Name))
toAbstract (NewName BindingSource
b Name
x) = do
    y <- Name -> ScopeM Name
freshAbstractName_ Name
x
    bindVariable b x y
    return y

instance ToAbstract (NewName C.BoundName) where
  type AbsOfCon (NewName C.BoundName) = A.BindName
  toAbstract :: NewName BoundName -> ScopeM (AbsOfCon (NewName BoundName))
toAbstract NewName{ newBinder :: forall a. NewName a -> BindingSource
newBinder = BindingSource
b, newName :: forall a. NewName a -> a
newName = BName{ boundName :: BoundName -> Name
boundName = Name
x, bnameFixity :: BoundName -> Fixity'
bnameFixity = Fixity'
fx }} = do
    y <- Fixity' -> Name -> ScopeM Name
freshAbstractName Fixity'
fx Name
x
    bindVariable b x y
    return $ A.BindName y

instance ToAbstract OldQName where
  type AbsOfCon OldQName = A.Expr
  toAbstract :: OldQName -> ScopeM (AbsOfCon OldQName)
toAbstract q :: OldQName
q@(OldQName QName
x Maybe (Set1 Name)
_) =
    ScopeM (AbsOfCon OldQName)
-> TCMT IO (Maybe (AbsOfCon OldQName))
-> ScopeM (AbsOfCon OldQName)
forall (m :: * -> *) a. Monad m => m a -> m (Maybe a) -> m a
fromMaybeM (QName -> ScopeM Expr
forall a. QName -> TCM a
notInScopeError QName
x) (TCMT IO (Maybe (AbsOfCon OldQName)) -> ScopeM (AbsOfCon OldQName))
-> TCMT IO (Maybe (AbsOfCon OldQName))
-> ScopeM (AbsOfCon OldQName)
forall a b. (a -> b) -> a -> b
$ MaybeOldQName -> ScopeM (AbsOfCon MaybeOldQName)
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract (OldQName -> MaybeOldQName
MaybeOldQName OldQName
q)

instance ToAbstract MaybeOldQName where
  type AbsOfCon MaybeOldQName = Maybe A.Expr
  toAbstract :: MaybeOldQName -> ScopeM (AbsOfCon MaybeOldQName)
toAbstract (MaybeOldQName (OldQName QName
x Maybe (Set1 Name)
ns)) = do
    qx <- KindsOfNames -> Maybe (Set1 Name) -> QName -> ScopeM ResolvedName
resolveName' KindsOfNames
allKindsOfNames Maybe (Set1 Name)
ns QName
x
    reportSLn "scope.name" 30 $ "resolved " ++ prettyShow x ++ ": " ++ prettyShow qx
    case qx of
      VarName Name
x' BindingSource
_         -> Maybe Expr -> TCMT IO (Maybe Expr)
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return (Maybe Expr -> TCMT IO (Maybe Expr))
-> Maybe Expr -> TCMT IO (Maybe Expr)
forall a b. (a -> b) -> a -> b
$ Expr -> Maybe Expr
forall a. a -> Maybe a
Just (Expr -> Maybe Expr) -> Expr -> Maybe Expr
forall a b. (a -> b) -> a -> b
$ Name -> Expr
A.Var Name
x'
      DefinedName Access
_ AbstractName
d Suffix
suffix -> do
        QName -> TCMT IO ()
forall (m :: * -> *).
(MonadWarning m, ReadTCState m) =>
QName -> m ()
raiseWarningsOnUsage (QName -> TCMT IO ()) -> QName -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ AbstractName -> QName
anameName AbstractName
d
        -- then we take note of generalized names used
        case AbstractName -> KindOfName
anameKind AbstractName
d of
          KindOfName
GeneralizeName -> do
            gvs <- Lens' TCState (Maybe (Set QName)) -> TCMT IO (Maybe (Set QName))
forall (m :: * -> *) a. ReadTCState m => Lens' TCState a -> m a
useTC (Maybe (Set QName) -> f (Maybe (Set QName)))
-> TCState -> f TCState
Lens' TCState (Maybe (Set QName))
stGeneralizedVars
            case gvs of   -- Subtle: Use (left-biased) union instead of insert to keep the old name if
                          -- already present. This way we can sort by source location when generalizing
                          -- (Issue 3354).
                Just Set QName
s -> (Maybe (Set QName) -> f (Maybe (Set QName)))
-> TCState -> f TCState
Lens' TCState (Maybe (Set QName))
stGeneralizedVars Lens' TCState (Maybe (Set QName))
-> Maybe (Set QName) -> TCMT IO ()
forall (m :: * -> *) a.
MonadTCState m =>
Lens' TCState a -> a -> m ()
`setTCLens` Set QName -> Maybe (Set QName)
forall a. a -> Maybe a
Just (Set QName
s Set QName -> Set QName -> Set QName
forall a. Ord a => Set a -> Set a -> Set a
`Set.union` QName -> Set QName
forall a. a -> Set a
Set.singleton (AbstractName -> QName
anameName AbstractName
d))
                Maybe (Set QName)
Nothing -> TypeError -> TCMT IO ()
forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
TypeError -> m a
typeError (TypeError -> TCMT IO ()) -> TypeError -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ QName -> TypeError
GeneralizeNotSupportedHere (QName -> TypeError) -> QName -> TypeError
forall a b. (a -> b) -> a -> b
$ AbstractName -> QName
anameName AbstractName
d
          KindOfName
DisallowedGeneralizeName -> TypeError -> TCMT IO ()
forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
TypeError -> m a
typeError (TypeError -> TCMT IO ()) -> TypeError -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ QName -> TypeError
GeneralizedVarInLetOpenedModule (QName -> TypeError) -> QName -> TypeError
forall a b. (a -> b) -> a -> b
$ AbstractName -> QName
anameName AbstractName
d
          KindOfName
_ -> () -> TCMT IO ()
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return ()
        -- and then we return the name
        Maybe Expr -> TCMT IO (Maybe Expr)
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return (Maybe Expr -> TCMT IO (Maybe Expr))
-> Maybe Expr -> TCMT IO (Maybe Expr)
forall a b. (a -> b) -> a -> b
$ Suffix -> Expr -> Maybe Expr
withSuffix Suffix
suffix (Expr -> Maybe Expr) -> Expr -> Maybe Expr
forall a b. (a -> b) -> a -> b
$ AbstractName -> Expr
forall a. NameToExpr a => a -> Expr
nameToExpr AbstractName
d
        where
          withSuffix :: Suffix -> Expr -> Maybe Expr
withSuffix Suffix
NoSuffix   Expr
e         = Expr -> Maybe Expr
forall a. a -> Maybe a
Just Expr
e
          withSuffix s :: Suffix
s@Suffix{} (A.Def QName
x) = Expr -> Maybe Expr
forall a. a -> Maybe a
Just (Expr -> Maybe Expr) -> Expr -> Maybe Expr
forall a b. (a -> b) -> a -> b
$ QName -> Suffix -> Expr
A.Def' QName
x Suffix
s
          withSuffix Suffix
_          Expr
_         = Maybe Expr
forall a. Maybe a
Nothing

      FieldName     List1 AbstractName
ds     -> (AmbiguousQName -> Expr)
-> List1 AbstractName -> TCMT IO (Maybe Expr)
ambiguous (ProjOrigin -> AmbiguousQName -> Expr
A.Proj ProjOrigin
ProjPrefix) List1 AbstractName
ds
      ConstructorName Set1 Induction
_ List1 AbstractName
ds -> (AmbiguousQName -> Expr)
-> List1 AbstractName -> TCMT IO (Maybe Expr)
ambiguous AmbiguousQName -> Expr
A.Con List1 AbstractName
ds
      PatternSynResName List1 AbstractName
ds -> (AmbiguousQName -> Expr)
-> List1 AbstractName -> TCMT IO (Maybe Expr)
ambiguous AmbiguousQName -> Expr
A.PatternSyn List1 AbstractName
ds
      ResolvedName
UnknownName          -> do
        [Char] -> Int -> [Char] -> TCMT IO ()
forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> [Char] -> m ()
reportSLn [Char]
"scope.name.unknown" Int
80 ([Char] -> TCMT IO ()) -> [Char] -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ [Char]
"resolved : unknown " [Char] -> [Char] -> [Char]
forall a. [a] -> [a] -> [a]
++ QName -> [Char]
forall a. Pretty a => a -> [Char]
prettyShow QName
x
        Maybe Expr -> TCMT IO (Maybe Expr)
forall a. a -> TCMT IO a
forall (f :: * -> *) a. Applicative f => a -> f a
pure Maybe Expr
forall a. Maybe a
Nothing
    where
      ambiguous :: (AmbiguousQName -> A.Expr) -> List1 AbstractName -> ScopeM (Maybe A.Expr)
      ambiguous :: (AmbiguousQName -> Expr)
-> List1 AbstractName -> TCMT IO (Maybe Expr)
ambiguous AmbiguousQName -> Expr
f List1 AbstractName
ds = do
        let xs :: List1 QName
xs = (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
ds
        List1 QName -> TCMT IO ()
raiseWarningsOnUsageIfUnambiguous List1 QName
xs
        Maybe Expr -> TCMT IO (Maybe Expr)
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return (Maybe Expr -> TCMT IO (Maybe Expr))
-> Maybe Expr -> TCMT IO (Maybe Expr)
forall a b. (a -> b) -> a -> b
$ Expr -> Maybe Expr
forall a. a -> Maybe a
Just (Expr -> Maybe Expr) -> Expr -> Maybe Expr
forall a b. (a -> b) -> a -> b
$ AmbiguousQName -> Expr
f (AmbiguousQName -> Expr) -> AmbiguousQName -> Expr
forall a b. (a -> b) -> a -> b
$ List1 QName -> AmbiguousQName
AmbQ List1 QName
xs

      -- Note: user warnings on ambiguous names will be raised by the type checker,
      -- see 'storeDisambiguatedName'.
      raiseWarningsOnUsageIfUnambiguous :: List1 A.QName -> ScopeM ()
      raiseWarningsOnUsageIfUnambiguous :: List1 QName -> TCMT IO ()
raiseWarningsOnUsageIfUnambiguous = \case
        QName
x :| [] -> QName -> TCMT IO ()
forall (m :: * -> *).
(MonadWarning m, ReadTCState m) =>
QName -> m ()
raiseWarningsOnUsage QName
x
        List1 QName
_       -> () -> TCMT IO ()
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return ()


instance ToAbstract ResolveQName where
  type AbsOfCon ResolveQName = ResolvedName
  toAbstract :: ResolveQName -> ScopeM (AbsOfCon ResolveQName)
toAbstract (ResolveQName QName
x) = QName -> ScopeM ResolvedName
resolveName QName
x ScopeM ResolvedName
-> (ResolvedName -> ScopeM ResolvedName) -> ScopeM ResolvedName
forall a b. TCMT IO a -> (a -> TCMT IO b) -> TCMT IO b
forall (m :: * -> *) a b. Monad m => m a -> (a -> m b) -> m b
>>= \case
    ResolvedName
UnknownName -> QName -> ScopeM ResolvedName
forall a. QName -> TCM a
notInScopeError QName
x
    ResolvedName
q -> ResolvedName -> ScopeM ResolvedName
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return ResolvedName
q

-- | A name resolved in a pattern.
data APatName
  = VarPatName A.Name
      -- ^ Pattern variable.
  | ConPatName (List1 AbstractName)
      -- ^ A (possibly ambiguous) constructor.
      --   When parsing a 'C.DisplayPragma', this can be the name of a definition.
  | PatternSynPatName (List1 AbstractName)
      -- ^ A (possibly ambiguous) pattern synonym.
  | DefPatName AbstractName
      -- ^ A defined name, only possible when checking a 'C.DisplayPragma'.

instance ToAbstract PatName where
  type AbsOfCon PatName = APatName
  toAbstract :: PatName -> ScopeM (AbsOfCon PatName)
toAbstract (PatName QName
x Maybe (Set1 Name)
ns Hiding
h DisplayLHS
displayLhs) = do
    [Char] -> Int -> [Char] -> TCMT IO ()
forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> [Char] -> m ()
reportSLn [Char]
"scope.pat" Int
30 ([Char] -> TCMT IO ()) -> [Char] -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ [Char]
"checking pattern name: " [Char] -> [Char] -> [Char]
forall a. [a] -> [a] -> [a]
++ QName -> [Char]
forall a. Pretty a => a -> [Char]
prettyShow QName
x
    let kinds :: [KindOfName]
kinds = DisplayLHS
-> ([KindOfName] -> [KindOfName]) -> [KindOfName] -> [KindOfName]
forall b a. IsBool b => b -> (a -> a) -> a -> a
applyWhen DisplayLHS
displayLhs ([KindOfName]
defNameKinds [KindOfName] -> [KindOfName] -> [KindOfName]
forall a. [a] -> [a] -> [a]
++) [KindOfName]
conLikeNameKinds
    rx <- KindsOfNames -> Maybe (Set1 Name) -> QName -> ScopeM ResolvedName
resolveName' ([KindOfName] -> KindsOfNames
someKindsOfNames [KindOfName]
kinds) Maybe (Set1 Name)
ns QName
x
          -- Andreas, 2013-03-21 ignore conflicting names which cannot
          -- be meant since we are in a pattern
          -- Andreas, 2020-04-11 CoConName:
          -- coinductive constructors will be rejected later, in the type checker
    reportSLn "scope.pat" 40 $ "resolved as " ++ prettyShow rx
    case rx of
      ConstructorName Set1 Induction
_ List1 AbstractName
ds -> List1 AbstractName -> APatName
ConPatName List1 AbstractName
ds APatName -> TCMT IO () -> TCMT IO APatName
forall a b. a -> TCMT IO b -> TCMT IO a
forall (f :: * -> *) a b. Functor f => a -> f b -> f a
<$ do
        [Char] -> Int -> [Char] -> TCMT IO ()
forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> [Char] -> m ()
reportSLn [Char]
"scope.pat" Int
30 ([Char] -> TCMT IO ()) -> [Char] -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ [Char]
"it was a con: " [Char] -> [Char] -> [Char]
forall a. [a] -> [a] -> [a]
++ List1 QName -> [Char]
forall a. Pretty a => a -> [Char]
prettyShow ((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
ds)
      PatternSynResName List1 AbstractName
ds -> List1 AbstractName -> APatName
PatternSynPatName List1 AbstractName
ds APatName -> TCMT IO () -> TCMT IO APatName
forall a b. a -> TCMT IO b -> TCMT IO a
forall (f :: * -> *) a b. Functor f => a -> f b -> f a
<$ do
        [Char] -> Int -> [Char] -> TCMT IO ()
forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> [Char] -> m ()
reportSLn [Char]
"scope.pat" Int
30 ([Char] -> TCMT IO ()) -> [Char] -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ [Char]
"it was a pat syn: " [Char] -> [Char] -> [Char]
forall a. [a] -> [a] -> [a]
++ List1 QName -> [Char]
forall a. Pretty a => a -> [Char]
prettyShow ((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
ds)
      DefinedName Access
_ AbstractName
d Suffix
suffix | DisplayLHS
YesDisplayLHS <- DisplayLHS
displayLhs, Suffix -> Bool
forall a. Null a => a -> Bool
null Suffix
suffix -> AbstractName -> APatName
DefPatName AbstractName
d APatName -> TCMT IO () -> TCMT IO APatName
forall a b. a -> TCMT IO b -> TCMT IO a
forall (f :: * -> *) a b. Functor f => a -> f b -> f a
<$ do
        [Char] -> Int -> [Char] -> TCMT IO ()
forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> [Char] -> m ()
reportSLn [Char]
"scope.pat" Int
30 ([Char] -> TCMT IO ()) -> [Char] -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ [Char]
"it was a def: " [Char] -> [Char] -> [Char]
forall a. [a] -> [a] -> [a]
++ QName -> [Char]
forall a. Pretty a => a -> [Char]
prettyShow (AbstractName -> QName
anameName AbstractName
d)
      ResolvedName
_ -> case QName
x of
        C.QName Name
y -> Name -> APatName
VarPatName (Name -> APatName) -> ScopeM Name -> TCMT IO APatName
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> Hiding -> Name -> ScopeM Name
bindPatternVariable Hiding
h Name
y
        C.Qual{}  -> TypeError -> TCMT IO APatName
forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
TypeError -> m a
typeError (TypeError -> TCMT IO APatName) -> TypeError -> TCMT IO APatName
forall a b. (a -> b) -> a -> b
$ Pattern -> TypeError
InvalidPattern (Pattern -> TypeError) -> Pattern -> TypeError
forall a b. (a -> b) -> a -> b
$ Bool -> QName -> Pattern
C.IdentP Bool
True QName
x

-- | Translate and possibly bind a pattern variable
--   (which could have been bound before due to non-linearity).
bindPatternVariable :: Hiding -> C.Name -> ScopeM A.Name
bindPatternVariable :: Hiding -> Name -> ScopeM Name
bindPatternVariable Hiding
h Name
x = do
  y <- (Name -> [(Name, LocalVar)] -> Maybe LocalVar
forall a b. Eq a => a -> [(a, b)] -> Maybe b
AssocList.lookup Name
x ([(Name, LocalVar)] -> Maybe LocalVar)
-> TCMT IO [(Name, LocalVar)] -> TCMT IO (Maybe LocalVar)
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> TCMT IO [(Name, LocalVar)]
getVarsToBind) TCMT IO (Maybe LocalVar)
-> (Maybe LocalVar -> ScopeM Name) -> ScopeM Name
forall a b. TCMT IO a -> (a -> TCMT IO b) -> TCMT IO b
forall (m :: * -> *) a b. Monad m => m a -> (a -> m b) -> m b
>>= \case
    Just (LocalVar Name
y BindingSource
_ [AbstractName]
_) -> do
      [Char] -> Int -> [Char] -> TCMT IO ()
forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> [Char] -> m ()
reportSLn [Char]
"scope.pat" Int
30 ([Char] -> TCMT IO ()) -> [Char] -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ [Char]
"it was a old var: " [Char] -> [Char] -> [Char]
forall a. [a] -> [a] -> [a]
++ Name -> [Char]
forall a. Pretty a => a -> [Char]
prettyShow Name
x
      Name -> ScopeM Name
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return (Name -> ScopeM Name) -> Name -> ScopeM Name
forall a b. (a -> b) -> a -> b
$ Range -> Name -> Name
forall a. SetRange a => Range -> a -> a
setRange (Name -> Range
forall a. HasRange a => a -> Range
getRange Name
x) Name
y
    Maybe LocalVar
Nothing -> do
      [Char] -> Int -> [Char] -> TCMT IO ()
forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> [Char] -> m ()
reportSLn [Char]
"scope.pat" Int
30 ([Char] -> TCMT IO ()) -> [Char] -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ [Char]
"it was a new var: " [Char] -> [Char] -> [Char]
forall a. [a] -> [a] -> [a]
++ Name -> [Char]
forall a. Pretty a => a -> [Char]
prettyShow Name
x
      Name -> ScopeM Name
freshAbstractName_ Name
x
  addVarToBind x $ LocalVar y (PatternBound h) []
  return y

class ToQName a where
  toQName :: a -> C.QName

instance ToQName C.Name  where toQName :: Name -> QName
toQName = Name -> QName
C.QName
instance ToQName C.QName where toQName :: QName -> QName
toQName = QName -> QName
forall a. a -> a
id

-- | Should be a defined name.
instance ToQName a => ToAbstract (OldName a) where
  type AbsOfCon (OldName a) = A.QName
  toAbstract :: OldName a -> ScopeM (AbsOfCon (OldName a))
toAbstract (OldName a
x) = do
    QName -> ScopeM ResolvedName
resolveName (a -> QName
forall a. ToQName a => a -> QName
toQName a
x) ScopeM ResolvedName
-> (ResolvedName -> TCMT IO QName) -> TCMT IO QName
forall a b. TCMT IO a -> (a -> TCMT IO b) -> TCMT IO b
forall (m :: * -> *) a b. Monad m => m a -> (a -> m b) -> m b
>>= \case
      DefinedName Access
_ AbstractName
d Suffix
NoSuffix -> QName -> TCMT IO QName
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return (QName -> TCMT IO QName) -> QName -> TCMT IO QName
forall a b. (a -> b) -> a -> b
$ AbstractName -> QName
anameName AbstractName
d
      DefinedName Access
_ AbstractName
d Suffix{} -> TCMT IO QName
forall a. HasCallStack => a
__IMPOSSIBLE__
      VarName{}                -> TCMT IO QName
forall a. HasCallStack => a
__IMPOSSIBLE__
      ResolvedName
UnknownName              -> TCMT IO QName
forall a. HasCallStack => a
__IMPOSSIBLE__
      -- We can get the cases below for DISPLAY pragmas
      ConstructorName Set1 Induction
_ List1 AbstractName
ds -> QName -> TCMT IO QName
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return (QName -> TCMT IO QName) -> QName -> TCMT IO QName
forall a b. (a -> b) -> a -> b
$ AbstractName -> QName
anameName (List1 AbstractName -> AbstractName
forall a. NonEmpty a -> a
List1.head List1 AbstractName
ds)   -- We'll throw out this one, so it doesn't matter which one we pick
      FieldName List1 AbstractName
ds         -> QName -> TCMT IO QName
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return (QName -> TCMT IO QName) -> QName -> TCMT IO QName
forall a b. (a -> b) -> a -> b
$ AbstractName -> QName
anameName (List1 AbstractName -> AbstractName
forall a. NonEmpty a -> a
List1.head List1 AbstractName
ds)
      PatternSynResName List1 AbstractName
ds -> QName -> TCMT IO QName
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return (QName -> TCMT IO QName) -> QName -> TCMT IO QName
forall a b. (a -> b) -> a -> b
$ AbstractName -> QName
anameName (List1 AbstractName -> AbstractName
forall a. NonEmpty a -> a
List1.head List1 AbstractName
ds)

newtype NewModuleName      = NewModuleName      C.Name
newtype NewModuleQName     = NewModuleQName     C.QName
newtype OldModuleName      = OldModuleName      C.QName

freshQModule :: A.ModuleName -> C.Name -> ScopeM A.ModuleName
freshQModule :: ModuleName -> Name -> TCMT IO ModuleName
freshQModule ModuleName
m Name
x = ModuleName -> ModuleName -> ModuleName
A.qualifyM ModuleName
m (ModuleName -> ModuleName)
-> (Name -> ModuleName) -> Name -> ModuleName
forall b c a. (b -> c) -> (a -> b) -> a -> c
. List1 Name -> ModuleName
mnameFromList1 (List1 Name -> ModuleName)
-> (Name -> List1 Name) -> Name -> ModuleName
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Name -> List1 Name
forall el coll. Singleton el coll => el -> coll
singleton (Name -> ModuleName) -> ScopeM Name -> TCMT IO ModuleName
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> Name -> ScopeM Name
freshAbstractName_ Name
x

checkForModuleClash :: C.Name -> ScopeM ()
checkForModuleClash :: Name -> TCMT IO ()
checkForModuleClash Name
x = do
  ms :: [AbstractModule] <- QName -> ScopeInfo -> [AbstractModule]
forall a. InScope a => QName -> ScopeInfo -> [a]
scopeLookup (Name -> QName
C.QName Name
x) (ScopeInfo -> [AbstractModule])
-> TCMT IO ScopeInfo -> TCMT IO [AbstractModule]
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> TCMT IO ScopeInfo
forall (m :: * -> *). ReadTCState m => m ScopeInfo
getScope
  List1.unlessNull ms \ List1 AbstractModule
ms -> do
    [Char] -> Int -> [Char] -> TCMT IO ()
forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> [Char] -> m ()
reportSLn [Char]
"scope.clash" Int
40 ([Char] -> TCMT IO ()) -> [Char] -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ [Char]
"clashing modules ms = " [Char] -> [Char] -> [Char]
forall a. [a] -> [a] -> [a]
++ List1 AbstractModule -> [Char]
forall a. Pretty a => a -> [Char]
prettyShow List1 AbstractModule
ms
    [Char] -> Int -> [Char] -> TCMT IO ()
forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> [Char] -> m ()
reportSLn [Char]
"scope.clash" Int
60 ([Char] -> TCMT IO ()) -> [Char] -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ [Char]
"clashing modules ms = " [Char] -> [Char] -> [Char]
forall a. [a] -> [a] -> [a]
++ List1 AbstractModule -> [Char]
forall a. Show a => a -> [Char]
show List1 AbstractModule
ms
    Name -> TCMT IO () -> TCMT IO ()
forall (m :: * -> *) x a.
(MonadTrace m, HasRange x) =>
x -> m a -> m a
setCurrentRange Name
x (TCMT IO () -> TCMT IO ()) -> TCMT IO () -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$
      TypeError -> TCMT IO ()
forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
TypeError -> m a
typeError (TypeError -> TCMT IO ()) -> TypeError -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ Name -> List1 ModuleName -> TypeError
ShadowedModule Name
x (List1 ModuleName -> TypeError) -> List1 ModuleName -> TypeError
forall a b. (a -> b) -> a -> b
$ (AbstractModule -> ModuleName)
-> List1 AbstractModule -> List1 ModuleName
forall a b. (a -> b) -> NonEmpty a -> NonEmpty b
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
fmap ((ModuleName -> Name -> ModuleName
forall t u. (SetRange t, HasRange u) => t -> u -> t
`withRangeOf` Name
x) (ModuleName -> ModuleName)
-> (AbstractModule -> ModuleName) -> AbstractModule -> ModuleName
forall b c a. (b -> c) -> (a -> b) -> a -> c
. AbstractModule -> ModuleName
amodName) List1 AbstractModule
ms

instance ToAbstract NewModuleName where
  type AbsOfCon NewModuleName = A.ModuleName
  toAbstract :: NewModuleName -> ScopeM (AbsOfCon NewModuleName)
toAbstract (NewModuleName Name
x) = do
    Name -> TCMT IO ()
checkForModuleClash Name
x
    m <- TCMT IO ModuleName
forall (m :: * -> *). ReadTCState m => m ModuleName
getCurrentModule
    y <- freshQModule m x
    createModule Nothing y
    return y

instance ToAbstract NewModuleQName where
  type AbsOfCon NewModuleQName = A.ModuleName
  toAbstract :: NewModuleQName -> ScopeM (AbsOfCon NewModuleQName)
toAbstract (NewModuleQName QName
m) = ModuleName -> QName -> TCMT IO ModuleName
toAbs ModuleName
noModuleName QName
m
    where
      toAbs :: ModuleName -> QName -> TCMT IO ModuleName
toAbs ModuleName
m (C.QName Name
x)  = do
        y <- ModuleName -> Name -> TCMT IO ModuleName
freshQModule ModuleName
m Name
x
        createModule Nothing y
        return y
      toAbs ModuleName
m (C.Qual Name
x QName
q) = do
        m' <- ModuleName -> Name -> TCMT IO ModuleName
freshQModule ModuleName
m Name
x
        toAbs m' q

instance ToAbstract OldModuleName where
  type AbsOfCon OldModuleName = A.ModuleName

  toAbstract :: OldModuleName -> ScopeM (AbsOfCon OldModuleName)
toAbstract (OldModuleName QName
q) = QName
-> ScopeM (AbsOfCon OldModuleName)
-> ScopeM (AbsOfCon OldModuleName)
forall (m :: * -> *) x a.
(MonadTrace m, HasRange x) =>
x -> m a -> m a
setCurrentRange QName
q (ScopeM (AbsOfCon OldModuleName)
 -> ScopeM (AbsOfCon OldModuleName))
-> ScopeM (AbsOfCon OldModuleName)
-> ScopeM (AbsOfCon OldModuleName)
forall a b. (a -> b) -> a -> b
$ do
    AbstractModule -> ModuleName
amodName (AbstractModule -> ModuleName)
-> TCMT IO AbstractModule -> TCMT IO ModuleName
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> QName -> TCMT IO AbstractModule
resolveModule QName
q

-- Expressions ------------------------------------------------------------
--UNUSED Liang-Ting Chen 2019-07-16
---- | Peel off 'C.HiddenArg' and represent it as an 'NamedArg'.
--mkNamedArg :: C.Expr -> NamedArg C.Expr
--mkNamedArg (C.HiddenArg   _ e) = Arg (hide         defaultArgInfo) e
--mkNamedArg (C.InstanceArg _ e) = Arg (makeInstance defaultArgInfo) e
--mkNamedArg e                   = Arg defaultArgInfo $ unnamed e

-- | Peel off 'C.HiddenArg' and represent it as an 'Arg', throwing away any name.
mkArg' :: ArgInfo -> C.Expr -> Arg C.Expr
mkArg' :: ArgInfo -> Expr -> Arg Expr
mkArg' ArgInfo
info (C.HiddenArg   Range
_ Named NamedName Expr
e) = ArgInfo -> Expr -> Arg Expr
forall e. ArgInfo -> e -> Arg e
Arg (ArgInfo -> ArgInfo
forall a. LensHiding a => a -> a
hide         ArgInfo
info) (Expr -> Arg Expr) -> Expr -> Arg Expr
forall a b. (a -> b) -> a -> b
$ Named NamedName Expr -> Expr
forall name a. Named name a -> a
namedThing Named NamedName Expr
e
mkArg' ArgInfo
info (C.InstanceArg Range
_ Named NamedName Expr
e) = ArgInfo -> Expr -> Arg Expr
forall e. ArgInfo -> e -> Arg e
Arg (ArgInfo -> ArgInfo
forall a. LensHiding a => a -> a
makeInstance ArgInfo
info) (Expr -> Arg Expr) -> Expr -> Arg Expr
forall a b. (a -> b) -> a -> b
$ Named NamedName Expr -> Expr
forall name a. Named name a -> a
namedThing Named NamedName Expr
e
mkArg' ArgInfo
info Expr
e                   = ArgInfo -> Expr -> Arg Expr
forall e. ArgInfo -> e -> Arg e
Arg (Hiding -> ArgInfo -> ArgInfo
forall a. LensHiding a => Hiding -> a -> a
setHiding Hiding
NotHidden ArgInfo
info) Expr
e

inferParenPreference :: C.Expr -> ParenPreference
inferParenPreference :: Expr -> ParenPreference
inferParenPreference C.Paren{} = ParenPreference
PreferParen
inferParenPreference Expr
_         = ParenPreference
PreferParenless

-- | Parse a possibly dotted and braced @C.Expr@ as @A.Expr@,
--   interpreting dots as relevance and braces as hiding.
--   Only accept a layer of dotting/bracing if the respective accumulator is @Nothing@.
toAbstractDotHiding :: Maybe Relevance -> Maybe Hiding -> Precedence -> C.Expr -> ScopeM (A.Expr, Relevance, Hiding)
toAbstractDotHiding :: Maybe Relevance
-> Maybe Hiding
-> Precedence
-> Expr
-> ScopeM (Expr, Relevance, Hiding)
toAbstractDotHiding Maybe Relevance
mr Maybe Hiding
mh Precedence
prec Expr
e = do
    [Char] -> Int -> [Char] -> TCMT IO ()
forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> [Char] -> m ()
reportSLn [Char]
"scope.irrelevance" Int
100 ([Char] -> TCMT IO ()) -> [Char] -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ [Char]
"toAbstractDotHiding: " [Char] -> [Char] -> [Char]
forall a. [a] -> [a] -> [a]
++ Doc -> [Char]
forall a. Doc a -> [Char]
render (Expr -> Doc
forall a. Pretty a => a -> Doc
pretty Expr
e)
    Call
-> ScopeM (Expr, Relevance, Hiding)
-> ScopeM (Expr, Relevance, Hiding)
forall a. Call -> TCMT IO a -> TCMT IO a
forall (m :: * -> *) a. MonadTrace m => Call -> m a -> m a
traceCall (Expr -> Call
ScopeCheckExpr Expr
e) (ScopeM (Expr, Relevance, Hiding)
 -> ScopeM (Expr, Relevance, Hiding))
-> ScopeM (Expr, Relevance, Hiding)
-> ScopeM (Expr, Relevance, Hiding)
forall a b. (a -> b) -> a -> b
$ case Expr
e of

      C.RawApp Range
_ List2 Expr
es     -> Maybe Relevance
-> Maybe Hiding
-> Precedence
-> Expr
-> ScopeM (Expr, Relevance, Hiding)
toAbstractDotHiding Maybe Relevance
mr Maybe Hiding
mh Precedence
prec (Expr -> ScopeM (Expr, Relevance, Hiding))
-> TCMT IO Expr -> ScopeM (Expr, Relevance, Hiding)
forall (m :: * -> *) a b. Monad m => (a -> m b) -> m a -> m b
=<< List2 Expr -> TCMT IO Expr
parseApplication List2 Expr
es
      C.Paren Range
_ Expr
e       -> Maybe Relevance
-> Maybe Hiding
-> Precedence
-> Expr
-> ScopeM (Expr, Relevance, Hiding)
toAbstractDotHiding Maybe Relevance
mr Maybe Hiding
mh Precedence
TopCtx Expr
e

      C.Dot KwRange
kwr Expr
e
        | Maybe Relevance
Nothing <- Maybe Relevance
mr -> Maybe Relevance
-> Maybe Hiding
-> Precedence
-> Expr
-> ScopeM (Expr, Relevance, Hiding)
toAbstractDotHiding (Relevance -> Maybe Relevance
forall a. a -> Maybe a
Just (Relevance -> Maybe Relevance) -> Relevance -> Maybe Relevance
forall a b. (a -> b) -> a -> b
$ OriginIrrelevant -> Relevance
Irrelevant (OriginIrrelevant -> Relevance) -> OriginIrrelevant -> Relevance
forall a b. (a -> b) -> a -> b
$ Range -> OriginIrrelevant
OIrrDot (Range -> OriginIrrelevant) -> Range -> OriginIrrelevant
forall a b. (a -> b) -> a -> b
$ KwRange -> Range
forall a. HasRange a => a -> Range
getRange KwRange
kwr) Maybe Hiding
mh Precedence
prec Expr
e

      C.DoubleDot KwRange
kwr Expr
e
        | Maybe Relevance
Nothing <- Maybe Relevance
mr -> Maybe Relevance
-> Maybe Hiding
-> Precedence
-> Expr
-> ScopeM (Expr, Relevance, Hiding)
toAbstractDotHiding (Relevance -> Maybe Relevance
forall a. a -> Maybe a
Just (Relevance -> Maybe Relevance) -> Relevance -> Maybe Relevance
forall a b. (a -> b) -> a -> b
$ OriginShapeIrrelevant -> Relevance
ShapeIrrelevant (OriginShapeIrrelevant -> Relevance)
-> OriginShapeIrrelevant -> Relevance
forall a b. (a -> b) -> a -> b
$ Range -> OriginShapeIrrelevant
OShIrrDotDot (Range -> OriginShapeIrrelevant) -> Range -> OriginShapeIrrelevant
forall a b. (a -> b) -> a -> b
$ KwRange -> Range
forall a. HasRange a => a -> Range
getRange KwRange
kwr) Maybe Hiding
mh Precedence
prec Expr
e

      C.HiddenArg Range
_ (Named Maybe NamedName
Nothing Expr
e)
        | Maybe Hiding
Nothing <- Maybe Hiding
mh -> Maybe Relevance
-> Maybe Hiding
-> Precedence
-> Expr
-> ScopeM (Expr, Relevance, Hiding)
toAbstractDotHiding Maybe Relevance
mr (Hiding -> Maybe Hiding
forall a. a -> Maybe a
Just Hiding
Hidden) Precedence
TopCtx Expr
e

      C.InstanceArg Range
_ (Named Maybe NamedName
Nothing Expr
e)
        | Maybe Hiding
Nothing <- Maybe Hiding
mh -> Maybe Relevance
-> Maybe Hiding
-> Precedence
-> Expr
-> ScopeM (Expr, Relevance, Hiding)
toAbstractDotHiding Maybe Relevance
mr (Hiding -> Maybe Hiding
forall a. a -> Maybe a
Just (Hiding -> Maybe Hiding) -> Hiding -> Maybe Hiding
forall a b. (a -> b) -> a -> b
$ Overlappable -> Hiding
Instance Overlappable
NoOverlap) Precedence
TopCtx Expr
e

      Expr
e                 -> (, Relevance -> Maybe Relevance -> Relevance
forall a. a -> Maybe a -> a
fromMaybe Relevance
relevant Maybe Relevance
mr, Hiding -> Maybe Hiding -> Hiding
forall a. a -> Maybe a -> a
fromMaybe Hiding
NotHidden Maybe Hiding
mh) (Expr -> (Expr, Relevance, Hiding))
-> ScopeM Expr -> ScopeM (Expr, Relevance, Hiding)
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$>
                             Precedence -> Expr -> ScopeM (AbsOfCon Expr)
forall c. ToAbstract c => Precedence -> c -> ScopeM (AbsOfCon c)
toAbstractCtx Precedence
prec Expr
e

-- | Translate concrete expression under at least one binder into nested
--   lambda abstraction in abstract syntax.
toAbstractLam :: Range -> List1 C.LamBinding -> C.Expr -> Precedence -> ScopeM A.Expr
toAbstractLam :: Range
-> NonEmpty (LamBinding' TypedBinding)
-> Expr
-> Precedence
-> ScopeM Expr
toAbstractLam Range
r NonEmpty (LamBinding' TypedBinding)
bs Expr
e Precedence
ctx = do
  -- Translate the binders
  lvars0 <- TCMT IO [(Name, LocalVar)]
forall (m :: * -> *). ReadTCState m => m [(Name, LocalVar)]
getLocalVars
  localToAbstract (fmap (C.DomainFull . makeDomainFull) bs) $ \ AbsOfCon (NonEmpty (LamBinding' TypedBinding))
bs -> do
    lvars1 <- TCMT IO [(Name, LocalVar)]
forall (m :: * -> *). ReadTCState m => m [(Name, LocalVar)]
getLocalVars
    checkNoShadowing lvars0 lvars1
    -- Translate the body
    e <- toAbstractCtx ctx e
    -- We have at least one binder.  Get first @b@ and rest @bs@.
    return $ case List1.catMaybes bs of
      -- Andreas, 2020-06-18
      -- There is a pathological case in which we end up without binder:
      --   λ (let
      --        mutual -- warning: empty mutual block
      --     ) -> Set
      []   -> Expr
e
      LamBinding
b:[LamBinding]
bs -> ExprInfo -> LamBinding -> Expr -> Expr
A.Lam (Range -> ExprInfo
ExprRange Range
r) LamBinding
b (Expr -> Expr) -> Expr -> Expr
forall a b. (a -> b) -> a -> b
$ (LamBinding -> Expr -> Expr) -> Expr -> [LamBinding] -> Expr
forall a b. (a -> b -> b) -> b -> [a] -> b
forall (t :: * -> *) a b.
Foldable t =>
(a -> b -> b) -> b -> t a -> b
foldr LamBinding -> Expr -> Expr
mkLam Expr
e [LamBinding]
bs
  where
    mkLam :: LamBinding -> Expr -> Expr
mkLam LamBinding
b Expr
e = ExprInfo -> LamBinding -> Expr -> Expr
A.Lam (Range -> ExprInfo
ExprRange (Range -> ExprInfo) -> Range -> ExprInfo
forall a b. (a -> b) -> a -> b
$ LamBinding -> Expr -> Range
forall u t. (HasRange u, HasRange t) => u -> t -> Range
fuseRange LamBinding
b Expr
e) LamBinding
b Expr
e

-- | Scope check extended lambda expression.
scopeCheckExtendedLam ::
  Range -> Erased -> List1 C.LamClause -> ScopeM A.Expr
scopeCheckExtendedLam :: Range -> Erased -> List1 LamClause -> ScopeM Expr
scopeCheckExtendedLam Range
r Erased
e List1 LamClause
cs = do
  TCMT IO Bool -> TCMT IO () -> TCMT IO ()
forall (m :: * -> *). Monad m => m Bool -> m () -> m ()
whenM TCMT IO Bool
isInsideDotPattern (TCMT IO () -> TCMT IO ()) -> TCMT IO () -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ TypeError -> TCMT IO ()
forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
TypeError -> m a
typeError (TypeError -> TCMT IO ()) -> TypeError -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ NotAllowedInDotPatterns -> TypeError
NotAllowedInDotPatterns NotAllowedInDotPatterns
PatternLambdas

  -- Find an unused name for the extended lambda definition.
  cname <- Range -> Int -> [Char] -> TCMT IO Name
freshConcreteName Range
r Int
0 [Char]
extendedLambdaName
  name  <- freshAbstractName_ cname
  a <- asksTC (^. lensIsAbstract)
  reportSDoc "scope.extendedLambda" 30 $ vcat
    [ text $ "new extended lambda name (" ++ show a ++ "): " ++ prettyShow name
    ]
  verboseS "scope.extendedLambda" 60 $ do
    forM_ cs $ \ LamClause
c -> do
      [Char] -> Int -> [Char] -> TCMT IO ()
forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> [Char] -> m ()
reportSLn [Char]
"scope.extendedLambda" Int
60 ([Char] -> TCMT IO ()) -> [Char] -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ [Char]
"extended lambda lhs: " [Char] -> [Char] -> [Char]
forall a. [a] -> [a] -> [a]
++ [Pattern] -> [Char]
forall a. Show a => a -> [Char]
show (LamClause -> [Pattern]
C.lamLHS LamClause
c)
  qname <- qualifyName_ name
  bindName privateAccessInserted FunName cname qname

  -- Andreas, 2019-08-20
  -- Keep the following __IMPOSSIBLE__, which is triggered by -v scope.decl.trace:80,
  -- for testing issue #4016.
  d <- C.FunDef r [] a NotInstanceDef __IMPOSSIBLE__ __IMPOSSIBLE__ cname . List1.toList <$> do
          forM cs $ \ (LamClause [Pattern]
ps RHS' Expr
rhs Bool
ca) -> do
            let p :: Pattern
p   = List1 Pattern -> Pattern
C.rawAppP (List1 Pattern -> Pattern) -> List1 Pattern -> Pattern
forall a b. (a -> b) -> a -> b
$
                        KillRangeT Pattern
forall a. KillRange a => KillRangeT a
killRange (Bool -> QName -> Pattern
IdentP Bool
True (QName -> Pattern) -> QName -> Pattern
forall a b. (a -> b) -> a -> b
$ Name -> QName
C.QName Name
cname) Pattern -> [Pattern] -> List1 Pattern
forall a. a -> [a] -> NonEmpty a
:| [Pattern]
ps
            let lhs :: LHS
lhs = Pattern -> [RewriteEqn] -> [WithExpr] -> LHS
C.LHS Pattern
p [] []
            Clause -> TCMT IO Clause
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return (Clause -> TCMT IO Clause) -> Clause -> TCMT IO Clause
forall a b. (a -> b) -> a -> b
$ Name
-> Bool
-> LHS
-> RHS' Expr
-> WhereClause' [Declaration]
-> [Clause]
-> Clause
C.Clause Name
cname Bool
ca LHS
lhs RHS' Expr
rhs WhereClause' [Declaration]
forall decls. WhereClause' decls
NoWhere []
  scdef <- toAbstract d

  -- Create the abstract syntax for the extended lambda.
  case scdef of
    A.ScopedDecl ScopeInfo
si [A.FunDef DefInfo
di QName
qname' [Clause]
cs] -> do
      ScopeInfo -> TCMT IO ()
setScope ScopeInfo
si  -- This turns into an A.ScopedExpr si $ A.ExtendedLam...
      Expr -> ScopeM Expr
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return (Expr -> ScopeM Expr) -> Expr -> ScopeM Expr
forall a b. (a -> b) -> a -> b
$
        ExprInfo -> DefInfo -> Erased -> QName -> List1 Clause -> Expr
A.ExtendedLam (Range -> ExprInfo
ExprRange Range
r) DefInfo
di Erased
e QName
qname' (List1 Clause -> Expr) -> List1 Clause -> Expr
forall a b. (a -> b) -> a -> b
$
        List1 Clause -> [Clause] -> List1 Clause
forall a. List1 a -> [a] -> List1 a
List1.fromListSafe List1 Clause
forall a. HasCallStack => a
__IMPOSSIBLE__ [Clause]
cs
    Declaration
_ -> ScopeM Expr
forall a. HasCallStack => a
__IMPOSSIBLE__

-- | Scope check an expression.

instance ToAbstract C.Expr where
  type AbsOfCon C.Expr = A.Expr

  toAbstract :: Expr -> ScopeM (AbsOfCon Expr)
toAbstract Expr
e =
    Call -> ScopeM (AbsOfCon Expr) -> ScopeM (AbsOfCon Expr)
forall a. Call -> TCMT IO a -> TCMT IO a
forall (m :: * -> *) a. MonadTrace m => Call -> m a -> m a
traceCall (Expr -> Call
ScopeCheckExpr Expr
e) (ScopeM (AbsOfCon Expr) -> ScopeM (AbsOfCon Expr))
-> ScopeM (AbsOfCon Expr) -> ScopeM (AbsOfCon Expr)
forall a b. (a -> b) -> a -> b
$ ScopeM Expr -> ScopeM Expr
forall (m :: * -> *). ReadTCState m => m Expr -> m Expr
annotateExpr (ScopeM Expr -> ScopeM Expr) -> ScopeM Expr -> ScopeM Expr
forall a b. (a -> b) -> a -> b
$ case Expr
e of

  -- Names
      Ident QName
x -> OldQName -> ScopeM (AbsOfCon OldQName)
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract (QName -> Maybe (Set1 Name) -> OldQName
OldQName QName
x Maybe (Set1 Name)
forall a. Maybe a
Nothing)
      KnownIdent NameKind
_ QName
x -> OldQName -> ScopeM (AbsOfCon OldQName)
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract (QName -> Maybe (Set1 Name) -> OldQName
OldQName QName
x Maybe (Set1 Name)
forall a. Maybe a
Nothing)
      -- Just discard the syntax highlighting information.

  -- Literals
      C.Lit Range
r Literal
l -> do
        Literal -> TCMT IO ()
checkLiteral Literal
l
        case Literal
l of
          LitNat Integer
n -> do
            let builtin :: TCMT IO (Maybe Term)
builtin | Integer
n Integer -> Integer -> Bool
forall a. Ord a => a -> a -> Bool
< Integer
0     = Term -> Maybe Term
forall a. a -> Maybe a
Just (Term -> Maybe Term) -> TCMT IO Term -> TCMT IO (Maybe Term)
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> TCMT IO Term
forall (m :: * -> *).
(HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) =>
m Term
primFromNeg    -- negative literals are only allowed if FROMNEG is defined
                        | Bool
otherwise = Maybe Term -> TCMT IO (Maybe Term)
ensureInScope (Maybe Term -> TCMT IO (Maybe Term))
-> TCMT IO (Maybe Term) -> TCMT IO (Maybe Term)
forall (m :: * -> *) a b. Monad m => (a -> m b) -> m a -> m b
=<< BuiltinId -> TCMT IO (Maybe Term)
forall (m :: * -> *). HasBuiltins m => BuiltinId -> m (Maybe Term)
getBuiltin' BuiltinId
builtinFromNat
            TCMT IO (Maybe Term)
builtin TCMT IO (Maybe Term) -> (Maybe Term -> ScopeM Expr) -> ScopeM Expr
forall a b. TCMT IO a -> (a -> TCMT IO b) -> TCMT IO b
forall (m :: * -> *) a b. Monad m => m a -> (a -> m b) -> m b
>>= \case
              Just (I.Def QName
q Elims
_) -> Expr -> ScopeM Expr
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return (Expr -> ScopeM Expr) -> Expr -> ScopeM Expr
forall a b. (a -> b) -> a -> b
$ QName -> Expr -> Expr
mkApp QName
q (Expr -> Expr) -> Expr -> Expr
forall a b. (a -> b) -> a -> b
$ ExprInfo -> Literal -> Expr
A.Lit ExprInfo
i (Literal -> Expr) -> Literal -> Expr
forall a b. (a -> b) -> a -> b
$ Integer -> Literal
LitNat (Integer -> Literal) -> Integer -> Literal
forall a b. (a -> b) -> a -> b
$ Integer -> Integer
forall a. Num a => a -> a
abs Integer
n
              Maybe Term
_                -> Expr -> ScopeM Expr
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return Expr
alit

          LitString Text
s -> do
            BuiltinId -> TCMT IO (Maybe Term)
forall (m :: * -> *). HasBuiltins m => BuiltinId -> m (Maybe Term)
getBuiltin' BuiltinId
builtinFromString TCMT IO (Maybe Term)
-> (Maybe Term -> TCMT IO (Maybe Term)) -> TCMT IO (Maybe Term)
forall a b. TCMT IO a -> (a -> TCMT IO b) -> TCMT IO b
forall (m :: * -> *) a b. Monad m => m a -> (a -> m b) -> m b
>>= Maybe Term -> TCMT IO (Maybe Term)
ensureInScope TCMT IO (Maybe Term) -> (Maybe Term -> ScopeM Expr) -> ScopeM Expr
forall a b. TCMT IO a -> (a -> TCMT IO b) -> TCMT IO b
forall (m :: * -> *) a b. Monad m => m a -> (a -> m b) -> m b
>>= \case
              Just (I.Def QName
q Elims
_) -> Expr -> ScopeM Expr
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return (Expr -> ScopeM Expr) -> Expr -> ScopeM Expr
forall a b. (a -> b) -> a -> b
$ QName -> Expr -> Expr
mkApp QName
q Expr
alit
              Maybe Term
_                -> Expr -> ScopeM Expr
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return Expr
alit

          Literal
_ -> Expr -> ScopeM Expr
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return Expr
alit
        where
        i :: ExprInfo
i       = Range -> ExprInfo
ExprRange Range
r
        alit :: Expr
alit    = ExprInfo -> Literal -> Expr
A.Lit ExprInfo
i Literal
l
        mkApp :: QName -> Expr -> Expr
mkApp QName
q = AppInfo -> Expr -> NamedArg Expr -> Expr
A.App (Range -> AppInfo
defaultAppInfo Range
r) (QName -> Expr
A.Def QName
q) (NamedArg Expr -> Expr) -> (Expr -> NamedArg Expr) -> Expr -> Expr
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Expr -> NamedArg Expr
forall a. a -> NamedArg a
defaultNamedArg

        -- #4925: Require fromNat/fromNeg to be in scope *unqualified* for literal overloading to
        -- apply.
        ensureInScope :: Maybe I.Term -> ScopeM (Maybe I.Term)
        ensureInScope :: Maybe Term -> TCMT IO (Maybe Term)
ensureInScope v :: Maybe Term
v@(Just (I.Def QName
q Elims
_)) =
          TCMT IO Bool
-> TCMT IO (Maybe Term)
-> TCMT IO (Maybe Term)
-> TCMT IO (Maybe Term)
forall (m :: * -> *) a. Monad m => m Bool -> m a -> m a -> m a
ifM (QName -> ScopeInfo -> Bool
isNameInScopeUnqualified QName
q (ScopeInfo -> Bool) -> TCMT IO ScopeInfo -> TCMT IO Bool
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> TCMT IO ScopeInfo
forall (m :: * -> *). ReadTCState m => m ScopeInfo
getScope) (Maybe Term -> TCMT IO (Maybe Term)
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return Maybe Term
v) (Maybe Term -> TCMT IO (Maybe Term)
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return Maybe Term
forall a. Maybe a
Nothing)
        ensureInScope Maybe Term
_ = Maybe Term -> TCMT IO (Maybe Term)
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return Maybe Term
forall a. Maybe a
Nothing

  -- Meta variables
      C.QuestionMark Range
r Maybe Int
n -> do
        scope <- TCMT IO ScopeInfo
forall (m :: * -> *). ReadTCState m => m ScopeInfo
getScope
        -- Andreas, 2014-04-06 create interaction point.
        ii <- registerInteractionPoint True r n
        let info = MetaInfo
             { metaRange :: Range
metaRange  = Range
r
             , metaScope :: ScopeInfo
metaScope  = ScopeInfo
scope
             , metaNumber :: Maybe MetaId
metaNumber = Maybe MetaId
forall a. Maybe a
Nothing
             , metaNameSuggestion :: [Char]
metaNameSuggestion = [Char]
""
             , metaKind :: MetaKind
metaKind   = MetaKind
UnificationMeta
             }
        return $ A.QuestionMark info ii
      C.Underscore Range
r Maybe [Char]
n -> do
        scope <- TCMT IO ScopeInfo
forall (m :: * -> *). ReadTCState m => m ScopeInfo
getScope
        return $ A.Underscore $ MetaInfo
                    { metaRange  = r
                    , metaScope  = scope
                    , metaNumber = __IMPOSSIBLE__ =<< n
                    , metaNameSuggestion = fromMaybe "" n
                    , metaKind   = UnificationMeta
                    }

  -- Raw application
      C.RawApp Range
r List2 Expr
es -> do
        e <- List2 Expr -> TCMT IO Expr
parseApplication List2 Expr
es
        toAbstract e

  -- Application
      C.App Range
r Expr
e1 NamedArg Expr
e2 -> do
        -- Andreas, 2021-02-10, issue #3289: reject @e {.p}@ and @e ⦃ .p ⦄@.

        -- Raise an error if argument is a C.Dot with Hiding info.
        case NamedArg Expr -> Expr
forall a. NamedArg a -> a
namedArg NamedArg Expr
e2 of
          C.Dot{} | NamedArg Expr -> Bool
forall a. LensHiding a => a -> Bool
notVisible NamedArg Expr
e2 -> NamedArg Expr -> TCMT IO () -> TCMT IO ()
forall (m :: * -> *) x a.
(MonadTrace m, HasRange x) =>
x -> m a -> m a
setCurrentRange NamedArg Expr
e2 (TCMT IO () -> TCMT IO ()) -> TCMT IO () -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ TypeError -> TCMT IO ()
forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
TypeError -> m a
typeError (TypeError -> TCMT IO ()) -> TypeError -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ NamedArg Expr -> TypeError
IllegalHidingInPostfixProjection NamedArg Expr
e2
          Expr
_ -> () -> TCMT IO ()
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return ()

        let parenPref :: ParenPreference
parenPref = Expr -> ParenPreference
inferParenPreference (NamedArg Expr -> Expr
forall a. NamedArg a -> a
namedArg NamedArg Expr
e2)
            info :: AppInfo
info = (Range -> AppInfo
defaultAppInfo Range
r) { appOrigin = UserWritten, appParens = parenPref }
        e1 <- Precedence -> Expr -> ScopeM (AbsOfCon Expr)
forall c. ToAbstract c => Precedence -> c -> ScopeM (AbsOfCon c)
toAbstractCtx Precedence
FunctionCtx Expr
e1
        e2 <- toAbstractCtx (ArgumentCtx parenPref) e2
        return $ A.App info e1 e2

  -- Operator application
      C.OpApp Range
r QName
op Set1 Name
ns OpAppArgs
es -> QName -> Set1 Name -> OpAppArgs -> ScopeM Expr
toAbstractOpApp QName
op Set1 Name
ns OpAppArgs
es
      C.KnownOpApp NameKind
_ Range
r QName
op Set1 Name
ns OpAppArgs
es -> QName -> Set1 Name -> OpAppArgs -> ScopeM Expr
toAbstractOpApp QName
op Set1 Name
ns OpAppArgs
es

  -- With application
      C.WithApp Range
r Expr
e List1 Expr
es -> do
        e  <- Precedence -> Expr -> ScopeM (AbsOfCon Expr)
forall c. ToAbstract c => Precedence -> c -> ScopeM (AbsOfCon c)
toAbstractCtx Precedence
WithFunCtx Expr
e
        es <- mapM (toAbstractCtx WithArgCtx) es
        return $ A.WithApp (ExprRange r) e es

  -- Misplaced hidden argument. We can treat these as parentheses and
  -- raise an error-warning
      C.HiddenArg Range
_ Named NamedName Expr
e' -> do
        Warning -> TCMT IO ()
forall (m :: * -> *).
(HasCallStack, MonadWarning m) =>
Warning -> m ()
warning (Expr -> Warning
HiddenNotInArgumentPosition Expr
e)
        Expr -> ScopeM (AbsOfCon Expr)
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract (Named NamedName Expr -> Expr
forall name a. Named name a -> a
namedThing Named NamedName Expr
e')

      C.InstanceArg Range
_ Named NamedName Expr
e' -> do
        Warning -> TCMT IO ()
forall (m :: * -> *).
(HasCallStack, MonadWarning m) =>
Warning -> m ()
warning (Expr -> Warning
InstanceNotInArgumentPosition Expr
e)
        Expr -> ScopeM (AbsOfCon Expr)
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract (Named NamedName Expr -> Expr
forall name a. Named name a -> a
namedThing Named NamedName Expr
e')

  -- Lambda
      C.AbsurdLam Range
r Hiding
h -> Expr -> ScopeM Expr
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return (Expr -> ScopeM Expr) -> Expr -> ScopeM Expr
forall a b. (a -> b) -> a -> b
$ ExprInfo -> Hiding -> Expr
A.AbsurdLam (Range -> ExprInfo
ExprRange Range
r) Hiding
h

      C.Lam Range
r NonEmpty (LamBinding' TypedBinding)
bs Expr
e -> Range
-> NonEmpty (LamBinding' TypedBinding)
-> Expr
-> Precedence
-> ScopeM Expr
toAbstractLam Range
r NonEmpty (LamBinding' TypedBinding)
bs Expr
e Precedence
TopCtx

  -- Extended Lambda
      C.ExtendedLam Range
r Erased
e List1 LamClause
cs -> Range -> Erased -> List1 LamClause -> ScopeM Expr
scopeCheckExtendedLam Range
r Erased
e List1 LamClause
cs

  -- Relevant and irrelevant non-dependent function type
      C.Fun Range
r (Arg ArgInfo
info1 Expr
e1) Expr
e2 -> do
        let arg :: Arg Expr
arg = ArgInfo -> Expr -> Arg Expr
mkArg' ArgInfo
info1 Expr
e1
        let mr :: Maybe Relevance
mr = case Arg Expr -> Relevance
forall a. LensRelevance a => a -> Relevance
getRelevance Arg Expr
arg of
              Relevant{} -> Maybe Relevance
forall a. Maybe a
Nothing
              Relevance
r -> Relevance -> Maybe Relevance
forall a. a -> Maybe a
Just Relevance
r
        let mh :: Maybe Hiding
mh = case Arg Expr -> Hiding
forall a. LensHiding a => a -> Hiding
getHiding Arg Expr
arg of
              Hiding
NotHidden -> Maybe Hiding
forall a. Maybe a
Nothing
              Hiding
h         -> Hiding -> Maybe Hiding
forall a. a -> Maybe a
Just Hiding
h
        Arg info (e1', rel, hid) <- (Expr -> ScopeM (Expr, Relevance, Hiding))
-> Arg Expr -> TCMT IO (Arg (Expr, Relevance, Hiding))
forall (t :: * -> *) (f :: * -> *) a b.
(Traversable t, Applicative f) =>
(a -> f b) -> t a -> f (t b)
forall (f :: * -> *) a b.
Applicative f =>
(a -> f b) -> Arg a -> f (Arg b)
traverse (Maybe Relevance
-> Maybe Hiding
-> Precedence
-> Expr
-> ScopeM (Expr, Relevance, Hiding)
toAbstractDotHiding Maybe Relevance
mr Maybe Hiding
mh Precedence
FunctionSpaceDomainCtx) Arg Expr
arg
        let updRel = Bool -> (ArgInfo -> ArgInfo) -> ArgInfo -> ArgInfo
forall b a. IsBool b => b -> (a -> a) -> a -> a
applyUnless (Relevance -> Bool
forall a. LensRelevance a => a -> Bool
isRelevant Relevance
rel) ((ArgInfo -> ArgInfo) -> ArgInfo -> ArgInfo)
-> (ArgInfo -> ArgInfo) -> ArgInfo -> ArgInfo
forall a b. (a -> b) -> a -> b
$ Relevance -> ArgInfo -> ArgInfo
forall a. LensRelevance a => Relevance -> a -> a
setRelevance Relevance
rel
        let updHid = case Hiding
hid of
              Hiding
NotHidden -> ArgInfo -> ArgInfo
forall a. a -> a
id
              Hiding
hid       -> Hiding -> ArgInfo -> ArgInfo
forall a. LensHiding a => Hiding -> a -> a
setHiding Hiding
hid
        A.Fun (ExprRange r) (Arg (updRel $ updHid info) e1') <$> toAbstractCtx TopCtx e2

  -- Dependent function type
      e0 :: Expr
e0@(C.Pi Telescope1
tel Expr
e) -> do
        lvars0 <- TCMT IO [(Name, LocalVar)]
forall (m :: * -> *). ReadTCState m => m [(Name, LocalVar)]
getLocalVars
        localToAbstract tel $ \AbsOfCon Telescope1
tel -> do
          lvars1 <- TCMT IO [(Name, LocalVar)]
forall (m :: * -> *). ReadTCState m => m [(Name, LocalVar)]
getLocalVars
          checkNoShadowing lvars0 lvars1
          e <- toAbstractCtx TopCtx e
          let info = Range -> ExprInfo
ExprRange (Expr -> Range
forall a. HasRange a => a -> Range
getRange Expr
e0)
          return $ A.mkPi info (List1.catMaybes tel) e

  -- Let
      e0 :: Expr
e0@(C.Let Range
_ List1 Declaration
ds (Just Expr
e)) ->
        TCMT IO Bool -> ScopeM Expr -> ScopeM Expr -> ScopeM Expr
forall (m :: * -> *) a. Monad m => m Bool -> m a -> m a -> m a
ifM TCMT IO Bool
isInsideDotPattern (TypeError -> ScopeM Expr
forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
TypeError -> m a
typeError (TypeError -> ScopeM Expr) -> TypeError -> ScopeM Expr
forall a b. (a -> b) -> a -> b
$ NotAllowedInDotPatterns -> TypeError
NotAllowedInDotPatterns NotAllowedInDotPatterns
LetExpressions) {-else-} do
        LetDefs -> (AbsOfCon LetDefs -> ScopeM Expr) -> ScopeM Expr
forall c b.
ToAbstract c =>
c -> (AbsOfCon c -> ScopeM b) -> ScopeM b
localToAbstract (LetDefOrigin -> List1 Declaration -> LetDefs
LetDefs LetDefOrigin
ExprLetDef List1 Declaration
ds) ((AbsOfCon LetDefs -> ScopeM Expr) -> ScopeM Expr)
-> (AbsOfCon LetDefs -> ScopeM Expr) -> ScopeM Expr
forall a b. (a -> b) -> a -> b
$ \AbsOfCon LetDefs
ds' -> do
          e <- Precedence -> Expr -> ScopeM (AbsOfCon Expr)
forall c. ToAbstract c => Precedence -> c -> ScopeM (AbsOfCon c)
toAbstractCtx Precedence
TopCtx Expr
e
          let info = Range -> ExprInfo
ExprRange (Expr -> Range
forall a. HasRange a => a -> Range
getRange Expr
e0)
          return $ A.mkLet info ds' e
      C.Let Range
_ List1 Declaration
_ Maybe Expr
Nothing -> TypeError -> ScopeM Expr
forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
TypeError -> m a
typeError (TypeError -> ScopeM Expr) -> TypeError -> ScopeM Expr
forall a b. (a -> b) -> a -> b
$ NotAValidLetExpression -> TypeError
NotAValidLetExpression NotAValidLetExpression
MissingBody

  -- Record construction
      C.Rec Range
r RecordAssignments
fs  -> do
        fs' <- Precedence
-> RecordAssignments -> ScopeM (AbsOfCon RecordAssignments)
forall c. ToAbstract c => Precedence -> c -> ScopeM (AbsOfCon c)
toAbstractCtx Precedence
TopCtx RecordAssignments
fs
        let ds'  = [ LetBinding
d | Right (ModuleName
_, Just LetBinding
d) <- [Either (FieldAssignment' Expr) (ModuleName, Maybe LetBinding)]
fs' ]
            fs'' = (Either (FieldAssignment' Expr) (ModuleName, Maybe LetBinding)
 -> Either (FieldAssignment' Expr) ModuleName)
-> [Either (FieldAssignment' Expr) (ModuleName, Maybe LetBinding)]
-> RecordAssigns
forall a b. (a -> b) -> [a] -> [b]
map (((ModuleName, Maybe LetBinding) -> ModuleName)
-> Either (FieldAssignment' Expr) (ModuleName, Maybe LetBinding)
-> Either (FieldAssignment' Expr) ModuleName
forall b d a. (b -> d) -> Either a b -> Either a d
mapRight (ModuleName, Maybe LetBinding) -> ModuleName
forall a b. (a, b) -> a
fst) [Either (FieldAssignment' Expr) (ModuleName, Maybe LetBinding)]
fs'
            i    = Range -> ExprInfo
ExprRange Range
r
            ri   = Range -> RecInfo
recInfoBrace Range
r
        return $ A.mkLet i ds' (A.Rec ri fs'')
      C.RecWhere Range
r [Declaration]
ds -> Range
-> [Declaration] -> (RecInfo -> Assigns -> Expr) -> ScopeM Expr
checkRecordWhere Range
r [Declaration]
ds \ RecInfo
i Assigns
fs -> RecInfo -> RecordAssigns -> Expr
A.Rec RecInfo
i (RecordAssigns -> Expr) -> RecordAssigns -> Expr
forall a b. (a -> b) -> a -> b
$ (FieldAssignment' Expr
 -> Either (FieldAssignment' Expr) ModuleName)
-> Assigns -> RecordAssigns
forall a b. (a -> b) -> [a] -> [b]
map FieldAssignment' Expr -> Either (FieldAssignment' Expr) ModuleName
forall a b. a -> Either a b
Left Assigns
fs

  -- Record update
      C.RecUpdate Range
r Expr
e [FieldAssignment]
fs -> do
        RecInfo -> Expr -> Assigns -> Expr
A.RecUpdate (Range -> RecInfo
recInfoBrace Range
r) (Expr -> Assigns -> Expr)
-> ScopeM Expr -> TCMT IO (Assigns -> Expr)
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> Expr -> ScopeM (AbsOfCon Expr)
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract Expr
e TCMT IO (Assigns -> Expr) -> TCMT IO Assigns -> ScopeM Expr
forall a b. TCMT IO (a -> b) -> TCMT IO a -> TCMT IO b
forall (f :: * -> *) a b. Applicative f => f (a -> b) -> f a -> f b
<*> Precedence
-> [FieldAssignment] -> ScopeM (AbsOfCon [FieldAssignment])
forall c. ToAbstract c => Precedence -> c -> ScopeM (AbsOfCon c)
toAbstractCtx Precedence
TopCtx [FieldAssignment]
fs
      C.RecUpdateWhere Range
r Expr
e [Declaration]
ds -> do
        e <- Expr -> ScopeM (AbsOfCon Expr)
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract Expr
e
        checkRecordWhere r ds \ RecInfo
i Assigns
fs -> RecInfo -> Expr -> Assigns -> Expr
A.RecUpdate RecInfo
i Expr
e Assigns
fs

  -- Parenthesis
      C.Paren Range
_ Expr
e -> Precedence -> Expr -> ScopeM (AbsOfCon Expr)
forall c. ToAbstract c => Precedence -> c -> ScopeM (AbsOfCon c)
toAbstractCtx Precedence
TopCtx Expr
e

  -- Idiom brackets
      C.IdiomBrackets Range
r [Expr]
es ->
        Precedence -> Expr -> ScopeM (AbsOfCon Expr)
forall c. ToAbstract c => Precedence -> c -> ScopeM (AbsOfCon c)
toAbstractCtx Precedence
TopCtx (Expr -> ScopeM Expr) -> TCMT IO Expr -> ScopeM Expr
forall (m :: * -> *) a b. Monad m => (a -> m b) -> m a -> m b
=<< Range -> [Expr] -> TCMT IO Expr
parseIdiomBracketsSeq Range
r [Expr]
es

  -- Do notation
      C.DoBlock Range
r List1 DoStmt
ss ->
        Precedence -> Expr -> ScopeM (AbsOfCon Expr)
forall c. ToAbstract c => Precedence -> c -> ScopeM (AbsOfCon c)
toAbstractCtx Precedence
TopCtx (Expr -> ScopeM Expr) -> TCMT IO Expr -> ScopeM Expr
forall (m :: * -> *) a b. Monad m => (a -> m b) -> m a -> m b
=<< Range -> List1 DoStmt -> TCMT IO Expr
desugarDoNotation Range
r List1 DoStmt
ss

  -- Post-fix projections
      e0 :: Expr
e0@(C.Dot KwRange
_kwr Expr
e) -> ExprInfo -> Expr -> Expr
A.Dot (Range -> ExprInfo
ExprRange (Range -> ExprInfo) -> Range -> ExprInfo
forall a b. (a -> b) -> a -> b
$ Expr -> Range
forall a. HasRange a => a -> Range
getRange Expr
e0) (Expr -> Expr) -> ScopeM Expr -> ScopeM Expr
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> Expr -> ScopeM (AbsOfCon Expr)
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract Expr
e

  -- Pattern things
      C.As Range
_ Name
_ Expr
_ -> Expr -> ScopeM Expr
forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
Expr -> m a
notAnExpression Expr
e
      C.Absurd Range
_ -> Expr -> ScopeM Expr
forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
Expr -> m a
notAnExpression Expr
e

  -- Impossible things
      C.Equal{} -> [Char] -> ScopeM Expr
forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
[Char] -> m a
syntaxError [Char]
"unexpected '='" -- triggered by 'f = (x = e)'
      C.Ellipsis Range
_ -> [Char] -> ScopeM Expr
forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
[Char] -> m a
syntaxError [Char]
"unexpected '...'"  -- triggered by 'f = ...'
      C.DoubleDot KwRange
_ Expr
_ -> ScopeM Expr
forall a. HasCallStack => a
__IMPOSSIBLE__

  -- Quoting
      C.Quote Range
r -> Expr -> ScopeM Expr
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return (Expr -> ScopeM Expr) -> Expr -> ScopeM Expr
forall a b. (a -> b) -> a -> b
$ ExprInfo -> Expr
A.Quote (Range -> ExprInfo
ExprRange Range
r)
      C.QuoteTerm Range
r -> Expr -> ScopeM Expr
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return (Expr -> ScopeM Expr) -> Expr -> ScopeM Expr
forall a b. (a -> b) -> a -> b
$ ExprInfo -> Expr
A.QuoteTerm (Range -> ExprInfo
ExprRange Range
r)
      C.Unquote Range
r -> Expr -> ScopeM Expr
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return (Expr -> ScopeM Expr) -> Expr -> ScopeM Expr
forall a b. (a -> b) -> a -> b
$ ExprInfo -> Expr
A.Unquote (Range -> ExprInfo
ExprRange Range
r)

      C.Tactic Range
r Expr
e -> [Char] -> ScopeM Expr
forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
[Char] -> m a
syntaxError [Char]
"'tactic' can only appear in attributes"

  -- DontCare
      C.DontCare Expr
e -> Expr -> Expr
A.DontCare (Expr -> Expr) -> ScopeM Expr -> ScopeM Expr
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> Expr -> ScopeM (AbsOfCon Expr)
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract Expr
e

  -- forall-generalize
      C.Generalized Expr
e -> do
        (s, e) <- ScopeM Expr -> ScopeM (Set QName, Expr)
forall a. ScopeM a -> ScopeM (Set QName, a)
collectGeneralizables (ScopeM Expr -> ScopeM (Set QName, Expr))
-> ScopeM Expr -> ScopeM (Set QName, Expr)
forall a b. (a -> b) -> a -> b
$ Expr -> ScopeM (AbsOfCon Expr)
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract Expr
e
        pure $ A.generalized s e

instance ToAbstract C.ModuleAssignment where
  type AbsOfCon C.ModuleAssignment = (A.ModuleName, Maybe A.LetBinding)
  toAbstract :: ModuleAssignment -> ScopeM (AbsOfCon ModuleAssignment)
toAbstract (C.ModuleAssignment QName
m [Expr]
es ImportDirective
i)
    | [Expr] -> Bool
forall a. Null a => a -> Bool
null [Expr]
es Bool -> Bool -> Bool
&& ImportDirective -> Bool
forall n m. ImportDirective' n m -> Bool
isDefaultImportDir ImportDirective
i = (, Maybe LetBinding
forall a. Maybe a
Nothing) (ModuleName -> (ModuleName, Maybe LetBinding))
-> TCMT IO ModuleName -> TCMT IO (ModuleName, Maybe LetBinding)
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> OldModuleName -> ScopeM (AbsOfCon OldModuleName)
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract (QName -> OldModuleName
OldModuleName QName
m)
    | Bool
otherwise = do
        x <- Range -> NameId -> Name
C.NoName (QName -> Range
forall a. HasRange a => a -> Range
getRange QName
m) (NameId -> Name) -> TCMT IO NameId -> TCMT IO Name
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> TCMT IO NameId
forall i (m :: * -> *). MonadFresh i m => m i
fresh
        r <- checkModuleMacro LetApply LetOpenModule
               (getRange (m, es, i)) PublicAccess defaultErased x
               (C.SectionApp (getRange (m , es)) [] m es)
               DontOpen i
        case r of
          LetApply ModuleInfo
_ Erased
_ ModuleName
m' ModuleApplication
_ ScopeCopyInfo
_ ImportDirective
_ -> (ModuleName, Maybe LetBinding)
-> TCMT IO (ModuleName, Maybe LetBinding)
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return (ModuleName
m', LetBinding -> Maybe LetBinding
forall a. a -> Maybe a
Just LetBinding
r)
          LetBinding
_ -> TCMT IO (ModuleName, Maybe LetBinding)
forall a. HasCallStack => a
__IMPOSSIBLE__

instance ToAbstract c => ToAbstract (FieldAssignment' c) where
  type AbsOfCon (FieldAssignment' c) = FieldAssignment' (AbsOfCon c)

  toAbstract :: FieldAssignment' c -> ScopeM (AbsOfCon (FieldAssignment' c))
toAbstract = (c -> TCMT IO (AbsOfCon c))
-> FieldAssignment' c -> TCMT IO (FieldAssignment' (AbsOfCon c))
forall (t :: * -> *) (f :: * -> *) a b.
(Traversable t, Applicative f) =>
(a -> f b) -> t a -> f (t b)
forall (f :: * -> *) a b.
Applicative f =>
(a -> f b) -> FieldAssignment' a -> f (FieldAssignment' b)
traverse c -> TCMT IO (AbsOfCon c)
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract

instance ToAbstract (C.Binder' (NewName C.BoundName)) where
  type AbsOfCon (C.Binder' (NewName C.BoundName)) = A.Binder

  toAbstract :: Binder' (NewName BoundName)
-> ScopeM (AbsOfCon (Binder' (NewName BoundName)))
toAbstract (C.Binder Maybe Pattern
p BinderNameOrigin
o NewName BoundName
n) = do
    let name :: Name
name = BoundName -> Name
C.boundName (BoundName -> Name) -> BoundName -> Name
forall a b. (a -> b) -> a -> b
$ NewName BoundName -> BoundName
forall a. NewName a -> a
newName NewName BoundName
n

    -- If we do have a pattern then the variable needs to be inserted
    -- so we do need a proper internal name for it.
    --
    -- Amy, 2024-10-18: If we generated a name, then mark the binder
    -- name as being inserted.
    (n, o) <- if Bool -> Bool
not (Name -> Bool
forall a. IsNoName a => a -> Bool
isNoName Name
name Bool -> Bool -> Bool
&& Maybe Pattern -> Bool
forall a. Maybe a -> Bool
isJust Maybe Pattern
p) then (NewName BoundName, BinderNameOrigin)
-> TCMT IO (NewName BoundName, BinderNameOrigin)
forall a. a -> TCMT IO a
forall (f :: * -> *) a. Applicative f => a -> f a
pure (NewName BoundName
n, BinderNameOrigin
o) else do
      n' <- Range -> Int -> [Char] -> TCMT IO Name
freshConcreteName (BoundName -> Range
forall a. HasRange a => a -> Range
getRange (BoundName -> Range) -> BoundName -> Range
forall a b. (a -> b) -> a -> b
$ NewName BoundName -> BoundName
forall a. NewName a -> a
newName NewName BoundName
n) Int
0 [Char]
patternInTeleName
      pure (fmap (\ BoundName
n -> BoundName
n { C.boundName = n' }) n, InsertedBinderName)

    n <- toAbstract n
    -- Expand puns if optHiddenArgumentPuns is True.
    p <- traverse expandPunsOpt p
    -- Actually parsing the pattern, checking it is linear,
    -- and bind its variables
    p <- traverse parsePattern p
    p <- toAbstract p
    checkPatternLinearity p $ \List1 Name
ys ->
      TypeError -> TCMT IO ()
forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
TypeError -> m a
typeError (TypeError -> TCMT IO ()) -> TypeError -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ List1 Name -> TypeError
RepeatedVariablesInPattern List1 Name
ys
    bindVarsToBind
    p <- toAbstract p
    pure $ A.Binder p o n

instance ToAbstract C.LamBinding where
  type AbsOfCon C.LamBinding = Maybe A.LamBinding

  toAbstract :: LamBinding' TypedBinding
-> ScopeM (AbsOfCon (LamBinding' TypedBinding))
toAbstract (C.DomainFree NamedArg Binder
x)  = do
    tac <- (Expr -> ScopeM Expr)
-> TacticAttribute -> TCMT IO (TacticAttribute' Expr)
forall (t :: * -> *) (f :: * -> *) a b.
(Traversable t, Applicative f) =>
(a -> f b) -> t a -> f (t b)
forall (f :: * -> *) a b.
Applicative f =>
(a -> f b) -> TacticAttribute' a -> f (TacticAttribute' b)
traverse Expr -> ScopeM Expr
Expr -> ScopeM (AbsOfCon Expr)
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract (TacticAttribute -> TCMT IO (TacticAttribute' Expr))
-> TacticAttribute -> TCMT IO (TacticAttribute' Expr)
forall a b. (a -> b) -> a -> b
$ BoundName -> TacticAttribute
bnameTactic (BoundName -> TacticAttribute) -> BoundName -> TacticAttribute
forall a b. (a -> b) -> a -> b
$ Binder -> BoundName
forall a. Binder' a -> a
C.binderName (Binder -> BoundName) -> Binder -> BoundName
forall a b. (a -> b) -> a -> b
$ NamedArg Binder -> Binder
forall a. NamedArg a -> a
namedArg NamedArg Binder
x
    Just . A.DomainFree tac <$> toAbstract (updateNamedArg (fmap $ NewName LambdaBound) x)
  toAbstract (C.DomainFull TypedBinding
tb) = (TypedBinding -> LamBinding)
-> Maybe TypedBinding -> Maybe LamBinding
forall a b. (a -> b) -> Maybe a -> Maybe b
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
fmap TypedBinding -> LamBinding
A.DomainFull (Maybe TypedBinding -> Maybe LamBinding)
-> TCMT IO (Maybe TypedBinding) -> TCMT IO (Maybe LamBinding)
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> TypedBinding -> ScopeM (AbsOfCon TypedBinding)
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract TypedBinding
tb

makeDomainFull :: C.LamBinding -> C.TypedBinding
makeDomainFull :: LamBinding' TypedBinding -> TypedBinding
makeDomainFull (C.DomainFull TypedBinding
b) = TypedBinding
b
makeDomainFull (C.DomainFree NamedArg Binder
x) = Range -> List1 (NamedArg Binder) -> Expr -> TypedBinding
forall e. Range -> List1 (NamedArg Binder) -> e -> TypedBinding' e
C.TBind Range
r (NamedArg Binder -> List1 (NamedArg Binder)
forall el coll. Singleton el coll => el -> coll
singleton NamedArg Binder
x) (Expr -> TypedBinding) -> Expr -> TypedBinding
forall a b. (a -> b) -> a -> b
$ Range -> Maybe [Char] -> Expr
C.Underscore Range
r Maybe [Char]
forall a. Maybe a
Nothing
  where r :: Range
r = NamedArg Binder -> Range
forall a. HasRange a => a -> Range
getRange NamedArg Binder
x

instance ToAbstract C.TypedBinding where
  type AbsOfCon C.TypedBinding = Maybe A.TypedBinding

  toAbstract :: TypedBinding -> ScopeM (AbsOfCon TypedBinding)
toAbstract (C.TBind Range
r List1 (NamedArg Binder)
xs Expr
t) = do
    t' <- Precedence -> Expr -> ScopeM (AbsOfCon Expr)
forall c. ToAbstract c => Precedence -> c -> ScopeM (AbsOfCon c)
toAbstractCtx Precedence
TopCtx Expr
t
    tac <- C.TacticAttribute <$> do
     traverse toAbstract $
      -- Invariant: all tactics are the same
      -- (distributed in the parser, TODO: don't)
      case List1.mapMaybe (theTacticAttribute . bnameTactic . C.binderName . namedArg) xs of
        []      -> Maybe (Ranged Expr)
forall a. Maybe a
Nothing
        Ranged Expr
tac : [Ranged Expr]
_ -> Ranged Expr -> Maybe (Ranged Expr)
forall a. a -> Maybe a
Just Ranged Expr
tac

    let fin = (NamedArg Binder -> Bool) -> List1 (NamedArg Binder) -> Bool
forall (t :: * -> *) a. Foldable t => (a -> Bool) -> t a -> Bool
all (BoundName -> Bool
bnameIsFinite (BoundName -> Bool)
-> (NamedArg Binder -> BoundName) -> NamedArg Binder -> Bool
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Binder -> BoundName
forall a. Binder' a -> a
C.binderName (Binder -> BoundName)
-> (NamedArg Binder -> Binder) -> NamedArg Binder -> BoundName
forall b c a. (b -> c) -> (a -> b) -> a -> c
. NamedArg Binder -> Binder
forall a. NamedArg a -> a
namedArg) List1 (NamedArg Binder)
xs
    xs' <- toAbstract $ fmap (updateNamedArg (fmap $ NewName LambdaBound)) xs

    return $ Just $ A.TBind r (TypedBindingInfo tac fin) xs' t'
  toAbstract (C.TLet Range
r List1 Declaration
ds) = Range -> [LetBinding] -> Maybe TypedBinding
A.mkTLet Range
r ([LetBinding] -> Maybe TypedBinding)
-> TCMT IO [LetBinding] -> TCMT IO (Maybe TypedBinding)
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> LetDefs -> ScopeM (AbsOfCon LetDefs)
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract (LetDefOrigin -> List1 Declaration -> LetDefs
LetDefs LetDefOrigin
ExprLetDef List1 Declaration
ds)

-- | Scope check a module (top level function).
--
scopeCheckNiceModule
  :: Range
  -> Access
  -> Erased
  -> C.Name
  -> C.Telescope
  -> ScopeM [A.Declaration]
  -> ScopeM A.Declaration
       -- ^ The returned declaration is an 'A.Section'.
scopeCheckNiceModule :: Range
-> Access
-> Erased
-> Name
-> Telescope
-> ScopeM [Declaration]
-> TCMT IO Declaration
scopeCheckNiceModule Range
r Access
p Erased
e Name
name Telescope
tel ScopeM [Declaration]
checkDs
  | Telescope -> Bool
telHasOpenStmsOrModuleMacros Telescope
tel = do
      -- Andreas, 2013-12-10:
      -- If the module telescope contains open statements
      -- or module macros (Issue 1299),
      -- add an extra anonymous module around the current one.
      -- Otherwise, the open statements would create
      -- identifiers in the parent scope of the current module.
      -- But open statements in the module telescope should
      -- only affect the current module!
      Range
-> Access
-> Erased
-> Name
-> Telescope
-> ScopeM [Declaration]
-> TCMT IO Declaration
scopeCheckNiceModule Range
forall a. Range' a
noRange Access
p Erased
e Name
noName_ [] (ScopeM [Declaration] -> TCMT IO Declaration)
-> ScopeM [Declaration] -> TCMT IO Declaration
forall a b. (a -> b) -> a -> b
$ Declaration -> [Declaration]
forall el coll. Singleton el coll => el -> coll
singleton (Declaration -> [Declaration])
-> TCMT IO Declaration -> ScopeM [Declaration]
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$>
        Access -> TCMT IO Declaration
scopeCheckNiceModule_ Access
PublicAccess  -- See #4350

  | Bool
otherwise = do
        Access -> TCMT IO Declaration
scopeCheckNiceModule_ Access
p
  where
    -- The actual workhorse:
    scopeCheckNiceModule_ :: Access -> ScopeM A.Declaration
    scopeCheckNiceModule_ :: Access -> TCMT IO Declaration
scopeCheckNiceModule_ Access
p = do

      -- Check whether we are dealing with an anonymous module.
      -- This corresponds to a Coq/LEGO section.
      (name, p', open) <- do
        if Name -> Bool
forall a. IsNoName a => a -> Bool
isNoName Name
name then do
          (i :: NameId) <- TCMT IO NameId
forall i (m :: * -> *). MonadFresh i m => m i
fresh
          return (C.NoName (getRange name) i, privateAccessInserted, True)
         else (Name, Access, Bool) -> TCMT IO (Name, Access, Bool)
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return (Name
name, Access
p, Bool
False)

      -- Check and bind the module, using the supplied check for its contents.
      aname <- toAbstract (NewModuleName name)
      d <- snd <$> do
        scopeCheckModule r e (C.QName name) aname tel checkDs
      bindModule p' name aname

      -- If the module was anonymous open it public
      -- unless it's private, in which case we just open it (#2099)
      when open $
       void $ -- We can discard the returned default A.ImportDirective.
        openModule TopOpenModule (Just aname) (C.QName name) $
          defaultImportDir { publicOpen = boolToMaybe (p == PublicAccess) empty }
      return d

-- | Check whether a telescope has open declarations or module macros.
telHasOpenStmsOrModuleMacros :: C.Telescope -> Bool
telHasOpenStmsOrModuleMacros :: Telescope -> Bool
telHasOpenStmsOrModuleMacros = (TypedBinding -> Bool) -> Telescope -> Bool
forall (t :: * -> *) a. Foldable t => (a -> Bool) -> t a -> Bool
any TypedBinding -> Bool
forall {e}. TypedBinding' e -> Bool
yesBind
  where
    yesBind :: TypedBinding' e -> Bool
yesBind C.TBind{}     = Bool
False
    yesBind (C.TLet Range
_ List1 Declaration
ds) = (Declaration -> Bool) -> List1 Declaration -> Bool
forall (t :: * -> *) a. Foldable t => (a -> Bool) -> t a -> Bool
any Declaration -> Bool
yes List1 Declaration
ds
    yes :: Declaration -> Bool
yes C.ModuleMacro{}   = Bool
True
    yes C.Open{}          = Bool
True
    yes C.Import{}        = Bool
True -- not __IMPOSSIBLE__, see Issue #1718
      -- However, it does not matter what we return here, as this will
      -- become an error later: "Not a valid let-declaration".
      -- (Andreas, 2015-11-17)
    yes (C.Mutual   KwRange
_ [Declaration]
ds) = (Declaration -> Bool) -> [Declaration] -> Bool
forall (t :: * -> *) a. Foldable t => (a -> Bool) -> t a -> Bool
any Declaration -> Bool
yes [Declaration]
ds
    yes (C.Abstract KwRange
_ [Declaration]
ds) = (Declaration -> Bool) -> [Declaration] -> Bool
forall (t :: * -> *) a. Foldable t => (a -> Bool) -> t a -> Bool
any Declaration -> Bool
yes [Declaration]
ds
    yes (C.Private KwRange
_ Origin
_ [Declaration]
ds) = (Declaration -> Bool) -> [Declaration] -> Bool
forall (t :: * -> *) a. Foldable t => (a -> Bool) -> t a -> Bool
any Declaration -> Bool
yes [Declaration]
ds
    yes Declaration
_                 = Bool
False

{- UNUSED
telHasLetStms :: C.Telescope -> Bool
telHasLetStms = any isLetBind
  where
    isLetBind C.TBind{} = False
    isLetBind C.TLet{}  = True
-}

-- | We for now disallow let-bindings in @data@ and @record@ telescopes.
--   This due "nested datatypes"; there is no easy interpretation of
--   @
--      data D (A : Set) (open M A) (b : B) : Set where
--        c : D (A × A) b → D A b
--   @
--   where @B@ is brought in scope by @open M A@.

class EnsureNoLetStms a where
  ensureNoLetStms :: a -> ScopeM ()

  default ensureNoLetStms :: (Foldable t, EnsureNoLetStms b, t b ~ a) => a -> ScopeM ()
  ensureNoLetStms = (b -> TCMT IO ()) -> t b -> TCMT IO ()
forall (t :: * -> *) (f :: * -> *) a b.
(Foldable t, Applicative f) =>
(a -> f b) -> t a -> f ()
traverse_ b -> TCMT IO ()
forall a. EnsureNoLetStms a => a -> TCMT IO ()
ensureNoLetStms

instance EnsureNoLetStms C.Binder where
  ensureNoLetStms :: Binder -> TCMT IO ()
ensureNoLetStms arg :: Binder
arg@(C.Binder Maybe Pattern
p BinderNameOrigin
_ BoundName
n) =
    Bool -> TCMT IO () -> TCMT IO ()
forall b (m :: * -> *). (IsBool b, Monad m) => b -> m () -> m ()
when (Maybe Pattern -> Bool
forall a. Maybe a -> Bool
isJust Maybe Pattern
p) (TCMT IO () -> TCMT IO ()) -> TCMT IO () -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ TypeError -> TCMT IO ()
forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
TypeError -> m a
typeError (TypeError -> TCMT IO ()) -> TypeError -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ Binder -> TypeError
IllegalPatternInTelescope Binder
arg

instance EnsureNoLetStms C.TypedBinding where
  ensureNoLetStms :: TypedBinding -> TCMT IO ()
ensureNoLetStms = \case
    tb :: TypedBinding
tb@C.TLet{}    -> TypeError -> TCMT IO ()
forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
TypeError -> m a
typeError (TypeError -> TCMT IO ()) -> TypeError -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ TypedBinding -> TypeError
IllegalLetInTelescope TypedBinding
tb
    C.TBind Range
_ List1 (NamedArg Binder)
xs Expr
_ -> (NamedArg Binder -> TCMT IO ())
-> List1 (NamedArg Binder) -> TCMT IO ()
forall (t :: * -> *) (f :: * -> *) a b.
(Foldable t, Applicative f) =>
(a -> f b) -> t a -> f ()
traverse_ (Binder -> TCMT IO ()
forall a. EnsureNoLetStms a => a -> TCMT IO ()
ensureNoLetStms (Binder -> TCMT IO ())
-> (NamedArg Binder -> Binder) -> NamedArg Binder -> TCMT IO ()
forall b c a. (b -> c) -> (a -> b) -> a -> c
. NamedArg Binder -> Binder
forall a. NamedArg a -> a
namedArg) List1 (NamedArg Binder)
xs

instance EnsureNoLetStms a => EnsureNoLetStms (LamBinding' a) where
  ensureNoLetStms :: LamBinding' a -> TCMT IO ()
ensureNoLetStms = \case
    -- GA: DO NOT use traverse here: `LamBinding'` only uses its parameter in
    --     the DomainFull constructor so we would miss out on some potentially
    --     illegal lets! Cf. #4402
    C.DomainFree NamedArg Binder
a -> NamedArg Binder -> TCMT IO ()
forall a. EnsureNoLetStms a => a -> TCMT IO ()
ensureNoLetStms NamedArg Binder
a
    C.DomainFull a
a -> a -> TCMT IO ()
forall a. EnsureNoLetStms a => a -> TCMT IO ()
ensureNoLetStms a
a

instance EnsureNoLetStms a => EnsureNoLetStms (Named_ a) where
instance EnsureNoLetStms a => EnsureNoLetStms (NamedArg a) where
instance EnsureNoLetStms a => EnsureNoLetStms [a] where


-- | Returns the scope inside the checked module.
scopeCheckModule
  :: Range                   -- ^ The range of the module.
  -> Erased                  -- ^ Is the module erased?
  -> C.QName                 -- ^ The concrete name of the module.
  -> A.ModuleName            -- ^ The abstract name of the module.
  -> C.Telescope             -- ^ The module telescope.
  -> ScopeM [A.Declaration]  -- ^ The code for checking the module contents.
  -> ScopeM (ScopeInfo, A.Declaration)
       -- ^ The returned declaration is an 'A.Section'.
scopeCheckModule :: Range
-> Erased
-> QName
-> ModuleName
-> Telescope
-> ScopeM [Declaration]
-> TCMT IO (ScopeInfo, Declaration)
scopeCheckModule Range
r Erased
e QName
x ModuleName
qm Telescope
tel ScopeM [Declaration]
checkDs = do
  [Char] -> Int -> [Char] -> TCMT IO ()
printScope [Char]
"module" Int
40 ([Char] -> TCMT IO ()) -> [Char] -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ [Char]
"checking module " [Char] -> [Char] -> [Char]
forall a. [a] -> [a] -> [a]
++ QName -> [Char]
forall a. Pretty a => a -> [Char]
prettyShow QName
x
  -- Andreas, 2013-12-10: Telescope does not live in the new module
  -- but its parent, so check it before entering the new module.
  -- This is important for Nicolas Pouillard's open parametrized modules
  -- statements inside telescopes.
  res <- TCMT IO (ScopeInfo, Declaration)
-> TCMT IO (ScopeInfo, Declaration)
forall a. ScopeM a -> ScopeM a
withLocalVars (TCMT IO (ScopeInfo, Declaration)
 -> TCMT IO (ScopeInfo, Declaration))
-> TCMT IO (ScopeInfo, Declaration)
-> TCMT IO (ScopeInfo, Declaration)
forall a b. (a -> b) -> a -> b
$ do
    tel <- GenTel -> ScopeM (AbsOfCon GenTel)
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract (Telescope -> GenTel
GenTel Telescope
tel)
    withCurrentModule qm $ do
      -- pushScope m
      -- qm <- getCurrentModule
      printScope "module" 40 $ "inside module " ++ prettyShow x
      ds    <- checkDs
      scope <- getScope
      return (scope, A.Section r e (qm `withRangesOfQ` x) tel ds)

  -- Binding is done by the caller
  printScope "module" 40 $ "after module " ++ prettyShow x
  return res

-- | Temporary data type to scope check a file.
data TopLevel a = TopLevel
  { forall a. TopLevel a -> AbsolutePath
topLevelPath           :: AbsolutePath
    -- ^ The file path from which we loaded this module.
  , forall a. TopLevel a -> TopLevelModuleName' Range
topLevelExpectedName   :: TopLevelModuleName
    -- ^ The expected module name
    --   (coming from the import statement that triggered scope checking this file).
  , forall a. TopLevel a -> a
topLevelTheThing       :: a
    -- ^ The file content.
  }

data TopLevelInfo = TopLevelInfo
        { TopLevelInfo -> [Declaration]
topLevelDecls :: [A.Declaration]
        , TopLevelInfo -> ScopeInfo
topLevelScope :: ScopeInfo  -- ^ as seen from inside the module
        }

-- | The top-level module name.

topLevelModuleName :: TopLevelInfo -> A.ModuleName
topLevelModuleName :: TopLevelInfo -> ModuleName
topLevelModuleName = (ScopeInfo -> Lens' ScopeInfo ModuleName -> ModuleName
forall o i. o -> Lens' o i -> i
^. (ModuleName -> f ModuleName) -> ScopeInfo -> f ScopeInfo
Lens' ScopeInfo ModuleName
scopeCurrent) (ScopeInfo -> ModuleName)
-> (TopLevelInfo -> ScopeInfo) -> TopLevelInfo -> ModuleName
forall b c a. (b -> c) -> (a -> b) -> a -> c
. TopLevelInfo -> ScopeInfo
topLevelScope

-- | Top-level declarations are always
--   @
--     (import|open)*         -- a bunch of possibly opened imports
--     module ThisModule ...  -- the top-level module of this file
--   @
instance ToAbstract (TopLevel [C.Declaration]) where
    type AbsOfCon (TopLevel [C.Declaration]) = TopLevelInfo

    toAbstract :: TopLevel [Declaration]
-> ScopeM (AbsOfCon (TopLevel [Declaration]))
toAbstract (TopLevel AbsolutePath
file TopLevelModuleName' Range
expectedMName [Declaration]
ds) =
      -- A file is a bunch of preliminary decls (imports etc.)
      -- plus a single module decl.
      case [Declaration] -> ([Declaration], [Declaration])
C.spanAllowedBeforeModule [Declaration]
ds of

        -- If there are declarations after the top-level module
        -- we have to report a parse error here.
        ([Declaration]
_, C.Module{} : Declaration
d : [Declaration]
_) -> Declaration
-> ScopeM (AbsOfCon (TopLevel [Declaration]))
-> ScopeM (AbsOfCon (TopLevel [Declaration]))
forall (m :: * -> *) x a.
(MonadTrace m, HasRange x) =>
x -> m a -> m a
setCurrentRange Declaration
d (ScopeM (AbsOfCon (TopLevel [Declaration]))
 -> ScopeM (AbsOfCon (TopLevel [Declaration])))
-> ScopeM (AbsOfCon (TopLevel [Declaration]))
-> ScopeM (AbsOfCon (TopLevel [Declaration]))
forall a b. (a -> b) -> a -> b
$ TypeError -> TCMT IO TopLevelInfo
forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
TypeError -> m a
typeError TypeError
DeclarationsAfterTopLevelModule

        -- Otherwise, proceed.
        ([Declaration]
outsideDecls, [ C.Module Range
r Erased
e QName
m0 Telescope
tel [Declaration]
insideDecls ]) -> do
          -- If the module name is _ compute the name from the file path
          (m, top) <- if QName -> Bool
forall a. IsNoName a => a -> Bool
isNoName QName
m0
                then do
                  -- Andreas, 2017-07-28, issue #1077
                  -- Check if the insideDecls end in a single module which has the same
                  -- name as the file.  In this case, it is highly likely that the user
                  -- put some non-allowed declarations before the top-level module in error.
                  -- Andreas, 2017-10-19, issue #2808
                  -- Widen this check to:
                  -- If the first module of the insideDecls has the same name as the file,
                  -- report an error.
                  case ((Declaration -> Bool)
 -> [Declaration] -> ([Declaration], [Declaration]))
-> [Declaration]
-> (Declaration -> Bool)
-> ([Declaration], [Declaration])
forall a b c. (a -> b -> c) -> b -> a -> c
flip (Declaration -> Bool)
-> [Declaration] -> ([Declaration], [Declaration])
forall a. (a -> Bool) -> [a] -> ([a], [a])
span [Declaration]
insideDecls ((Declaration -> Bool) -> ([Declaration], [Declaration]))
-> (Declaration -> Bool) -> ([Declaration], [Declaration])
forall a b. (a -> b) -> a -> b
$ \case { C.Module{} -> Bool
False; Declaration
_ -> Bool
True } of
                    ([Declaration]
ds0, (C.Module Range
_ Erased
_ QName
m1 Telescope
_ [Declaration]
_ : [Declaration]
_))
                       | QName -> RawTopLevelModuleName
rawTopLevelModuleNameForQName QName
m1 RawTopLevelModuleName -> RawTopLevelModuleName -> Bool
forall a. Eq a => a -> a -> Bool
==
                         TopLevelModuleName' Range -> RawTopLevelModuleName
rawTopLevelModuleName TopLevelModuleName' Range
expectedMName
                         -- If the anonymous module comes from the user,
                         -- the range cannot be the beginningOfFile.
                         -- That is the range if the parser inserted the anon. module.
                       , Range
r Range -> Range -> Bool
forall a. Eq a => a -> a -> Bool
== Range -> Range
beginningOfFile ([Declaration] -> Range
forall a. HasRange a => a -> Range
getRange [Declaration]
insideDecls) -> do

                         -- GA #4888: We know we are in a bad place. But we still scopecheck
                         -- the initial segment on the off chance we generate a better error
                         -- message.
                         ScopeM [Declaration] -> TCMT IO ()
forall (f :: * -> *) a. Functor f => f a -> f ()
void ScopeM [Declaration]
importPrimitives
                         ScopeM [Declaration] -> TCMT IO ()
forall (f :: * -> *) a. Functor f => f a -> f ()
void (ScopeM [Declaration] -> TCMT IO ())
-> ScopeM [Declaration] -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ Declarations -> ScopeM (AbsOfCon Declarations)
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract ([Declaration] -> Declarations
Declarations [Declaration]
outsideDecls)
                         ScopeM [Declaration] -> TCMT IO ()
forall (f :: * -> *) a. Functor f => f a -> f ()
void (ScopeM [Declaration] -> TCMT IO ())
-> ScopeM [Declaration] -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ Declarations -> ScopeM (AbsOfCon Declarations)
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract ([Declaration] -> Declarations
Declarations [Declaration]
ds0)
                         -- Fail with a crude error otherwise
                         [Declaration]
-> TCMT IO (QName, TopLevelModuleName' Range)
-> TCMT IO (QName, TopLevelModuleName' Range)
forall (m :: * -> *) x a.
(MonadTrace m, HasRange x) =>
x -> m a -> m a
setCurrentRange [Declaration]
ds0 (TCMT IO (QName, TopLevelModuleName' Range)
 -> TCMT IO (QName, TopLevelModuleName' Range))
-> TCMT IO (QName, TopLevelModuleName' Range)
-> TCMT IO (QName, TopLevelModuleName' Range)
forall a b. (a -> b) -> a -> b
$ TypeError -> TCMT IO (QName, TopLevelModuleName' Range)
forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
TypeError -> m a
typeError TypeError
IllegalDeclarationBeforeTopLevelModule

                    -- Otherwise, reconstruct the top-level module name
                    ([Declaration], [Declaration])
_ -> do
                      let m :: QName
m = Name -> QName
C.QName (Name -> QName) -> Name -> QName
forall a b. (a -> b) -> a -> b
$ Range -> Name -> Name
forall a. SetRange a => Range -> a -> a
setRange (QName -> Range
forall a. HasRange a => a -> Range
getRange QName
m0) (Name -> Name) -> Name -> Name
forall a b. (a -> b) -> a -> b
$
                              [Char] -> Name
C.simpleName ([Char] -> Name) -> [Char] -> Name
forall a b. (a -> b) -> a -> b
$ [Char] -> [Char]
stringToRawName ([Char] -> [Char]) -> [Char] -> [Char]
forall a b. (a -> b) -> a -> b
$
                              AbsolutePath -> [Char]
rootNameModule AbsolutePath
file
                      top <- RawTopLevelModuleName -> TCM (TopLevelModuleName' Range)
S.topLevelModuleName
                               (QName -> RawTopLevelModuleName
rawTopLevelModuleNameForQName QName
m)
                      return (m, top)
                -- Andreas, 2017-05-17, issue #2574, keep name as jump target!
                -- Andreas, 2016-07-12, ALTERNATIVE:
                -- -- We assign an anonymous file module the name expected from
                -- -- its import.  For flat file structures, this is the same.
                -- -- For hierarchical file structures, this reverses the behavior:
                -- -- Loading the file by itself will fail, but it can be imported.
                -- -- The previous behavior is: it can be loaded by itself, but not
                -- -- be imported
                -- then return $ C.fromTopLevelModuleName expectedMName
                else do
                -- Andreas, 2014-03-28  Issue 1078
                -- We need to check the module name against the file name here.
                -- Otherwise one could sneak in a lie and confuse the scope
                -- checker.
                  top <- RawTopLevelModuleName -> TCM (TopLevelModuleName' Range)
S.topLevelModuleName
                           (QName -> RawTopLevelModuleName
rawTopLevelModuleNameForQName QName
m0)
                  checkModuleName top (SourceFile file) (Just expectedMName)
                  return (m0, top)
          setTopLevelModule top
          am <- toAbstract (NewModuleQName m)
          primitiveImport <- importPrimitives
          -- Scope check the declarations outside
          outsideDecls <- toAbstract (Declarations outsideDecls)
          (insideScope, insideDecl) <- scopeCheckModule r e m am tel $
             toAbstract (Declarations insideDecls)
          -- Andreas, 2020-05-13, issue #1804, #4647
          -- Do not eagerly remove private definitions, only when serializing
          -- let scope = over scopeModules (fmap $ restrictLocalPrivate am) insideScope
          let scope = ScopeInfo
insideScope
          setScope scope

          -- While scope-checking the top-level module we might have
          -- encountered several (possibly nested) opaque blocks. We
          -- must now ensure that these have transitively-closed
          -- unfolding sets.
          saturateOpaqueBlocks

          return $ TopLevelInfo (primitiveImport ++ outsideDecls ++ [ insideDecl ]) scope

        -- We already inserted the missing top-level module, see
        -- 'Agda.Syntax.Parser.Parser.figureOutTopLevelModule',
        -- thus, this case is impossible:
        ([Declaration], [Declaration])
_ -> TCMT IO TopLevelInfo
ScopeM (AbsOfCon (TopLevel [Declaration]))
forall a. HasCallStack => a
__IMPOSSIBLE__

-- | Declaration @open import Agda.Primitive using (Set)@ when 'optImportSorts'.
--   @Prop@ is added when 'optProp', and @SSet@ when 'optTwoLevel'.
importPrimitives :: ScopeM [A.Declaration]
importPrimitives :: ScopeM [Declaration]
importPrimitives = do
  TCMT IO Bool
-> ScopeM [Declaration]
-> ScopeM [Declaration]
-> ScopeM [Declaration]
forall (m :: * -> *) a. Monad m => m Bool -> m a -> m a -> m a
ifNotM (PragmaOptions -> Bool
optImportSorts (PragmaOptions -> Bool) -> TCMT IO PragmaOptions -> TCMT IO Bool
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> TCMT IO PragmaOptions
forall (m :: * -> *). HasOptions m => m PragmaOptions
pragmaOptions) ([Declaration] -> ScopeM [Declaration]
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return []) {- else -} do
    prop     <- PragmaOptions -> Bool
optProp     (PragmaOptions -> Bool) -> TCMT IO PragmaOptions -> TCMT IO Bool
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> TCMT IO PragmaOptions
forall (m :: * -> *). HasOptions m => m PragmaOptions
pragmaOptions
    twoLevel <- optTwoLevel <$> pragmaOptions
    -- Add implicit `open import Agda.Primitive using (Prop; Set; SSet)`
    let agdaPrimitiveName   = Name -> QName -> QName
Qual ([Char] -> Name
C.simpleName [Char]
"Agda") (QName -> QName) -> QName -> QName
forall a b. (a -> b) -> a -> b
$ Name -> QName
C.QName (Name -> QName) -> Name -> QName
forall a b. (a -> b) -> a -> b
$ [Char] -> Name
C.simpleName [Char]
"Primitive"
        usingDirective      = ([Char] -> ImportedName' Name Name)
-> [[Char]] -> [ImportedName' Name Name]
forall a b. (a -> b) -> [a] -> [b]
map (Name -> ImportedName' Name Name
forall n m. n -> ImportedName' n m
ImportedName (Name -> ImportedName' Name Name)
-> ([Char] -> Name) -> [Char] -> ImportedName' Name Name
forall b c a. (b -> c) -> (a -> b) -> a -> c
. [Char] -> Name
C.simpleName) ([[Char]] -> [ImportedName' Name Name])
-> [[Char]] -> [ImportedName' Name Name]
forall a b. (a -> b) -> a -> b
$ [[[Char]]] -> [[Char]]
forall (t :: * -> *) a. Foldable t => t [a] -> [a]
concat
          [ [ [Char]
"Prop" | Bool
prop     ]
          , [ [Char]
"Set"  | Bool
True     ]
          , [ [Char]
"SSet" | Bool
twoLevel ]
          ]
        directives          = Range
-> Using' Name Name
-> [ImportedName' Name Name]
-> [Renaming' Name Name]
-> Maybe KwRange
-> ImportDirective
forall n m.
Range
-> Using' n m
-> HidingDirective' n m
-> RenamingDirective' n m
-> Maybe KwRange
-> ImportDirective' n m
ImportDirective Range
forall a. Range' a
noRange ([ImportedName' Name Name] -> Using' Name Name
forall n m. [ImportedName' n m] -> Using' n m
Using [ImportedName' Name Name]
usingDirective) [] [] Maybe KwRange
forall a. Maybe a
Nothing
        importAgdaPrimitive = [Range
-> QName
-> Maybe AsName
-> OpenShortHand
-> ImportDirective
-> Declaration
C.Import Range
forall a. Range' a
noRange QName
agdaPrimitiveName Maybe AsName
forall a. Maybe a
Nothing OpenShortHand
C.DoOpen ImportDirective
directives]
    toAbstract (Declarations importAgdaPrimitive)

-- | runs Syntax.Concrete.Definitions.niceDeclarations on main module
niceDecls :: DoWarn -> [C.Declaration] -> ([NiceDeclaration] -> ScopeM a) -> ScopeM a
niceDecls :: forall a.
DoWarn
-> [Declaration] -> ([NiceDeclaration] -> ScopeM a) -> ScopeM a
niceDecls DoWarn
warn [Declaration]
ds [NiceDeclaration] -> ScopeM a
ret = [Declaration] -> ScopeM a -> ScopeM a
forall (m :: * -> *) x a.
(MonadTrace m, HasRange x) =>
x -> m a -> m a
setCurrentRange [Declaration]
ds (ScopeM a -> ScopeM a) -> ScopeM a -> ScopeM a
forall a b. (a -> b) -> a -> b
$ DoWarn -> [Declaration] -> ScopeM a -> ScopeM a
forall a. DoWarn -> [Declaration] -> ScopeM a -> ScopeM a
computeFixitiesAndPolarities DoWarn
warn [Declaration]
ds (ScopeM a -> ScopeM a) -> ScopeM a -> ScopeM a
forall a b. (a -> b) -> a -> b
$ do

  -- Some pragmas are not allowed in safe mode unless we are in a builtin module.
  -- So we need to tell the nicifier whether it should yell about unsafe pragmas.
  isSafe <- PragmaOptions -> Bool
forall a. LensSafeMode a => a -> Bool
Lens.getSafeMode (PragmaOptions -> Bool) -> TCMT IO PragmaOptions -> TCMT IO Bool
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> TCMT IO PragmaOptions
forall (m :: * -> *). HasOptions m => m PragmaOptions
pragmaOptions
  safeButNotBuiltin <- and2M
    -- NB: BlockArguments allow bullet-point style argument lists using @do@, hehe!
    do pure isSafe
    do not <$> do Lens.isBuiltinModuleWithSafePostulates . filePath =<< getCurrentPath

  -- We need to pass the fixities to the nicifier for clause grouping.
  fixs <- useScope scopeFixities

  -- Run nicifier.
  let (result, warns) = runNice (NiceEnv safeButNotBuiltin) $ niceDeclarations fixs ds

  -- Respect the @DoWarn@ directive. For this to be sound, we need to know for
  -- sure that each @Declaration@ is checked at least once with @DoWarn@.
  unless (warn == NoWarn || null warns) $ do
    -- If there are some warnings and the --safe flag is set,
    -- we check that none of the NiceWarnings are fatal
    when isSafe $ do
      let (errs, ws) = List.partition unsafeDeclarationWarning warns
      -- If some of them are, we fail
      List1.unlessNull errs \ List1 DeclarationWarning
errs -> do
        NiceWarnings
-> (List1 DeclarationWarning -> TCMT IO ()) -> TCMT IO ()
forall (m :: * -> *) a.
Applicative m =>
[a] -> (List1 a -> m ()) -> m ()
List1.unlessNull NiceWarnings
ws \ List1 DeclarationWarning
ws -> List1 Warning -> TCMT IO ()
forall (m :: * -> *).
(HasCallStack, MonadWarning m) =>
List1 Warning -> m ()
warnings (List1 Warning -> TCMT IO ()) -> List1 Warning -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ (DeclarationWarning -> Warning)
-> List1 DeclarationWarning -> List1 Warning
forall a b. (a -> b) -> NonEmpty a -> NonEmpty b
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
fmap DeclarationWarning -> Warning
NicifierIssue List1 DeclarationWarning
ws
        tcerrs <- (DeclarationWarning -> TCMT IO TCWarning)
-> List1 DeclarationWarning -> TCMT IO (NonEmpty TCWarning)
forall (t :: * -> *) (m :: * -> *) a b.
(Traversable t, Monad m) =>
(a -> m b) -> t a -> m (t b)
forall (m :: * -> *) a b.
Monad m =>
(a -> m b) -> NonEmpty a -> m (NonEmpty b)
mapM (Warning -> TCMT IO TCWarning
forall (m :: * -> *).
(HasCallStack, MonadWarning m) =>
Warning -> m TCWarning
warning_ (Warning -> TCMT IO TCWarning)
-> (DeclarationWarning -> Warning)
-> DeclarationWarning
-> TCMT IO TCWarning
forall b c a. (b -> c) -> (a -> b) -> a -> c
. DeclarationWarning -> Warning
NicifierIssue) List1 DeclarationWarning
errs
        setCurrentRange errs $ typeError $ NonFatalErrors $ Set1.fromList tcerrs
    -- Otherwise we simply record the warnings
    mapM_ (\ DeclarationWarning
w -> CallStack -> Warning -> TCMT IO ()
forall (m :: * -> *).
MonadWarning m =>
CallStack -> Warning -> m ()
warning' (DeclarationWarning -> CallStack
dwLocation DeclarationWarning
w) (Warning -> TCMT IO ()) -> Warning -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ DeclarationWarning -> Warning
NicifierIssue DeclarationWarning
w) warns
  case result of
    Left (DeclarationException CallStack
loc DeclarationException'
e) -> do
      [Char] -> Int -> [Char] -> TCMT IO ()
forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> [Char] -> m ()
reportSLn [Char]
"error" Int
2 ([Char] -> TCMT IO ()) -> [Char] -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ [Char]
"Error raised at " [Char] -> [Char] -> [Char]
forall a. [a] -> [a] -> [a]
++ CallStack -> [Char]
forall a. Pretty a => a -> [Char]
prettyShow CallStack
loc
      DeclarationException' -> ScopeM a -> ScopeM a
forall (m :: * -> *) x a.
(MonadTrace m, HasRange x) =>
x -> m a -> m a
setCurrentRange DeclarationException'
e (ScopeM a -> ScopeM a) -> ScopeM a -> ScopeM a
forall a b. (a -> b) -> a -> b
$ TypeError -> ScopeM a
forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
TypeError -> m a
typeError (TypeError -> ScopeM a) -> TypeError -> ScopeM a
forall a b. (a -> b) -> a -> b
$ DeclarationException' -> TypeError
NicifierError DeclarationException'
e
    Right [NiceDeclaration]
ds -> [NiceDeclaration] -> ScopeM a
ret [NiceDeclaration]
ds

-- | Wrapper to avoid instance conflict with generic list instance.
newtype Declarations = Declarations [C.Declaration]

instance ToAbstract Declarations where
  type AbsOfCon Declarations = [A.Declaration]

  toAbstract :: Declarations -> ScopeM (AbsOfCon Declarations)
toAbstract (Declarations [Declaration]
ds) = DoWarn
-> [Declaration]
-> ([NiceDeclaration] -> ScopeM [Declaration])
-> ScopeM [Declaration]
forall a.
DoWarn
-> [Declaration] -> ([NiceDeclaration] -> ScopeM a) -> ScopeM a
niceDecls DoWarn
DoWarn [Declaration]
ds [NiceDeclaration] -> ScopeM [Declaration]
[NiceDeclaration] -> ScopeM (AbsOfCon [NiceDeclaration])
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract

-- | Where did these 'LetDef's come from?
data LetDefOrigin
  = ExprLetDef
  -- ^ A let expression or do statement
  | RecordWhereLetDef
  -- ^ A @record where@ expression
  | RecordLetDef
  -- ^ Definitions in a record declaration, before the last field
  deriving (LetDefOrigin -> LetDefOrigin -> Bool
(LetDefOrigin -> LetDefOrigin -> Bool)
-> (LetDefOrigin -> LetDefOrigin -> Bool) -> Eq LetDefOrigin
forall a. (a -> a -> Bool) -> (a -> a -> Bool) -> Eq a
$c== :: LetDefOrigin -> LetDefOrigin -> Bool
== :: LetDefOrigin -> LetDefOrigin -> Bool
$c/= :: LetDefOrigin -> LetDefOrigin -> Bool
/= :: LetDefOrigin -> LetDefOrigin -> Bool
Eq, Int -> LetDefOrigin -> [Char] -> [Char]
[LetDefOrigin] -> [Char] -> [Char]
LetDefOrigin -> [Char]
(Int -> LetDefOrigin -> [Char] -> [Char])
-> (LetDefOrigin -> [Char])
-> ([LetDefOrigin] -> [Char] -> [Char])
-> Show LetDefOrigin
forall a.
(Int -> a -> [Char] -> [Char])
-> (a -> [Char]) -> ([a] -> [Char] -> [Char]) -> Show a
$cshowsPrec :: Int -> LetDefOrigin -> [Char] -> [Char]
showsPrec :: Int -> LetDefOrigin -> [Char] -> [Char]
$cshow :: LetDefOrigin -> [Char]
show :: LetDefOrigin -> [Char]
$cshowList :: [LetDefOrigin] -> [Char] -> [Char]
showList :: [LetDefOrigin] -> [Char] -> [Char]
Show)

data LetDefs = LetDefs LetDefOrigin (List1 C.Declaration)
data LetDef = LetDef LetDefOrigin NiceDeclaration

instance ToAbstract LetDefs where
  type AbsOfCon LetDefs = [A.LetBinding]

  toAbstract :: LetDefs -> ScopeM (AbsOfCon LetDefs)
  toAbstract :: LetDefs -> ScopeM (AbsOfCon LetDefs)
toAbstract (LetDefs LetDefOrigin
wh List1 Declaration
ds) =
    [List1 LetBinding] -> [LetBinding]
forall a. [List1 a] -> [a]
List1.concat ([List1 LetBinding] -> [LetBinding])
-> TCMT IO [List1 LetBinding] -> TCMT IO [LetBinding]
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> DoWarn
-> [Declaration]
-> ([NiceDeclaration] -> TCMT IO [List1 LetBinding])
-> TCMT IO [List1 LetBinding]
forall a.
DoWarn
-> [Declaration] -> ([NiceDeclaration] -> ScopeM a) -> ScopeM a
niceDecls DoWarn
DoWarn (List1 Declaration -> [Item (List1 Declaration)]
forall l. IsList l => l -> [Item l]
List1.toList List1 Declaration
ds) ([LetDef] -> TCMT IO [List1 LetBinding]
[LetDef] -> ScopeM (AbsOfCon [LetDef])
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract ([LetDef] -> TCMT IO [List1 LetBinding])
-> ([NiceDeclaration] -> [LetDef])
-> [NiceDeclaration]
-> TCMT IO [List1 LetBinding]
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (NiceDeclaration -> LetDef) -> [NiceDeclaration] -> [LetDef]
forall a b. (a -> b) -> [a] -> [b]
map (LetDefOrigin -> NiceDeclaration -> LetDef
LetDef LetDefOrigin
wh))

-- | Raise appropriate (error-)warnings for if a declaration with
-- illegal access, macro flag, or abstractness appear in a let
-- expression.
checkLetDefInfo :: LetDefOrigin -> Access -> IsMacro -> IsAbstract -> ScopeM ()
checkLetDefInfo :: LetDefOrigin -> Access -> IsMacro -> IsAbstract -> TCMT IO ()
checkLetDefInfo LetDefOrigin
wh Access
access IsMacro
macro IsAbstract
abstract = do
  Bool -> TCMT IO () -> TCMT IO ()
forall b (m :: * -> *). (IsBool b, Monad m) => b -> m () -> m ()
when (IsAbstract
abstract IsAbstract -> IsAbstract -> Bool
forall a. Eq a => a -> a -> Bool
== IsAbstract
AbstractDef) (TCMT IO () -> TCMT IO ()) -> TCMT IO () -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ Warning -> TCMT IO ()
forall (m :: * -> *).
(HasCallStack, MonadWarning m) =>
Warning -> m ()
warning Warning
AbstractInLetBindings

  Bool -> TCMT IO () -> TCMT IO ()
forall b (m :: * -> *). (IsBool b, Monad m) => b -> m () -> m ()
when (IsMacro
macro IsMacro -> IsMacro -> Bool
forall a. Eq a => a -> a -> Bool
== IsMacro
MacroDef) (TCMT IO () -> TCMT IO ()) -> TCMT IO () -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ Warning -> TCMT IO ()
forall (m :: * -> *).
(HasCallStack, MonadWarning m) =>
Warning -> m ()
warning Warning
MacroInLetBindings

  case Access
access of
    -- Marking a let declaration as private should only raise a warning
    -- in explicit, user-written expressions.
    --
    -- It should not raise a warning when scope-checking the type of a
    -- record constructor (it has an effect there), or when elaborating
    -- the lets generated by a 'record where' expression.
    PrivateAccess KwRange
rng Origin
_
      | LetDefOrigin
wh LetDefOrigin -> LetDefOrigin -> Bool
forall a. Eq a => a -> a -> Bool
== LetDefOrigin
ExprLetDef -> HasCallStack => DeclarationWarning' -> TCMT IO ()
DeclarationWarning' -> TCMT IO ()
scopeWarning (KwRange -> DeclarationWarning'
UselessPrivate KwRange
rng)
    Access
_ -> () -> TCMT IO ()
forall a. a -> TCMT IO a
forall (f :: * -> *) a. Applicative f => a -> f a
pure ()

instance ToAbstract LetDef where
  type AbsOfCon LetDef = List1 A.LetBinding
  toAbstract :: LetDef -> ScopeM (AbsOfCon LetDef)
  toAbstract :: LetDef -> ScopeM (AbsOfCon LetDef)
toAbstract (LetDef LetDefOrigin
wh NiceDeclaration
d) = NiceDeclaration
-> TCMT IO (List1 LetBinding) -> TCMT IO (List1 LetBinding)
forall (m :: * -> *) x a.
(MonadTrace m, HasRange x) =>
x -> m a -> m a
setCurrentRange NiceDeclaration
d case NiceDeclaration
d of
    NiceMutual KwRange
_ TerminationCheck
_ CoverageCheck
_ PositivityCheck
_ d :: [NiceDeclaration]
d@[C.FunSig Range
_ Access
access IsAbstract
_ IsInstance
instanc IsMacro
macro ArgInfo
info TerminationCheck
_ CoverageCheck
_ Name
x Expr
t, C.FunDef Range
_ [Declaration]
_ IsAbstract
abstract IsInstance
_ TerminationCheck
_ CoverageCheck
_ Name
_ [Clause
cl]] -> do
      LetDefOrigin -> Access -> IsMacro -> IsAbstract -> TCMT IO ()
checkLetDefInfo LetDefOrigin
wh Access
access IsMacro
macro IsAbstract
abstract

      t <- Expr -> ScopeM (AbsOfCon Expr)
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract Expr
t
      -- We bind the name here to make sure it's in scope for the LHS (#917).
      -- It's unbound for the RHS in letToAbstract.
      fx <- getConcreteFixity x
      x  <- A.unBind <$> toAbstract (NewName LetBound $ mkBoundName x fx)
      (x', e) <- letToAbstract cl

      -- If InstanceDef set info to Instance
      let info' = case IsInstance
instanc of
            InstanceDef KwRange
_  -> ArgInfo -> ArgInfo
forall a. LensHiding a => a -> a
makeInstance ArgInfo
info
            IsInstance
NotInstanceDef -> ArgInfo
info

      -- There are sometimes two instances of the let-bound variable,
      -- one declaration and one definition. The first list element
      -- below is used to highlight the declared instance in the right
      -- way (see Issue 1618).
      return $ A.LetDeclaredVariable (A.mkBindName (setRange (getRange x') x)) :|
        [ A.LetBind (LetRange $ getRange d) info' (A.mkBindName x) t e
        ]

    -- Function signature without a body
    C.Axiom Range
_ Access
acc IsAbstract
abs IsInstance
instanc ArgInfo
info Name
x Expr
t -> do
      LetDefOrigin -> Access -> IsMacro -> IsAbstract -> TCMT IO ()
checkLetDefInfo LetDefOrigin
wh Access
acc IsMacro
NotMacroDef IsAbstract
abs

      t <- Expr -> ScopeM (AbsOfCon Expr)
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract Expr
t
      fx <- getConcreteFixity x
      x  <- toAbstract (NewName LetBound $ mkBoundName x fx)

      let
        info' = case IsInstance
instanc of
          InstanceDef KwRange
_  -> ArgInfo -> ArgInfo
forall a. LensHiding a => a -> a
makeInstance ArgInfo
info
          IsInstance
NotInstanceDef -> ArgInfo
info

      pure $ A.LetAxiom (LetRange $ getRange d) info' x t :| []

    -- irrefutable let binding, like  (x , y) = rhs
    NiceFunClause Range
r Access
PublicAccess IsAbstract
ConcreteDef TerminationCheck
tc CoverageCheck
cc Bool
catchall d :: Declaration
d@(C.FunClause lhs :: LHS
lhs@(C.LHS Pattern
p0 [] []) RHS' Expr
rhs0 WhereClause' [Declaration]
whcl Bool
ca) -> do
      WhereClause' [Declaration] -> TCMT IO ()
noWhereInLetBinding WhereClause' [Declaration]
whcl
      rhs <- RHS' Expr -> TCMT IO Expr
letBindingMustHaveRHS RHS' Expr
rhs0
      -- Expand puns if optHiddenArgumentPuns is True.
      p0   <- expandPunsOpt p0
      mp   <- setCurrentRange p0 $
                (Right <$> parsePattern p0)
                  `catchError`
                (return . Left)
      case mp of
        Right Pattern
p -> do
          rhs <- Expr -> ScopeM (AbsOfCon Expr)
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract Expr
rhs
          setCurrentRange p0 $ do
            p   <- toAbstract p
            checkValidLetPattern p
            checkPatternLinearity p $ \List1 Name
ys ->
              TypeError -> TCMT IO ()
forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
TypeError -> m a
typeError (TypeError -> TCMT IO ()) -> TypeError -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ List1 Name -> TypeError
RepeatedVariablesInPattern List1 Name
ys
            bindVarsToBind
            p   <- toAbstract p
            return $ singleton $ A.LetPatBind (LetRange r) p rhs
        -- It's not a record pattern, so it should be a prefix left-hand side
        Left TCErr
err ->
          case Pattern -> Maybe Name
definedName Pattern
p0 of
            Maybe Name
Nothing -> TCErr -> TCMT IO (List1 LetBinding)
forall a. TCErr -> TCMT IO a
forall e (m :: * -> *) a. MonadError e m => e -> m a
throwError TCErr
err
            Just Name
x  -> LetDef -> ScopeM (AbsOfCon LetDef)
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract (LetDef -> ScopeM (AbsOfCon LetDef))
-> LetDef -> ScopeM (AbsOfCon LetDef)
forall a b. (a -> b) -> a -> b
$ LetDefOrigin -> NiceDeclaration -> LetDef
LetDef LetDefOrigin
wh (NiceDeclaration -> LetDef) -> NiceDeclaration -> LetDef
forall a b. (a -> b) -> a -> b
$ KwRange
-> TerminationCheck
-> CoverageCheck
-> PositivityCheck
-> [NiceDeclaration]
-> NiceDeclaration
NiceMutual KwRange
forall a. Null a => a
empty TerminationCheck
tc CoverageCheck
cc PositivityCheck
YesPositivityCheck
              [ Range
-> Access
-> IsAbstract
-> IsInstance
-> IsMacro
-> ArgInfo
-> TerminationCheck
-> CoverageCheck
-> Name
-> Expr
-> NiceDeclaration
C.FunSig Range
r Access
PublicAccess IsAbstract
ConcreteDef IsInstance
NotInstanceDef IsMacro
NotMacroDef
                  (Origin -> ArgInfo -> ArgInfo
forall a. LensOrigin a => Origin -> a -> a
setOrigin Origin
Inserted ArgInfo
defaultArgInfo) TerminationCheck
tc CoverageCheck
cc Name
x (Range -> Maybe [Char] -> Expr
C.Underscore (Name -> Range
forall a. HasRange a => a -> Range
getRange Name
x) Maybe [Char]
forall a. Maybe a
Nothing)
              , Range
-> [Declaration]
-> IsAbstract
-> IsInstance
-> TerminationCheck
-> CoverageCheck
-> Name
-> [Clause]
-> NiceDeclaration
C.FunDef Range
r [Declaration]
forall a. HasCallStack => a
__IMPOSSIBLE__ IsAbstract
ConcreteDef IsInstance
NotInstanceDef TerminationCheck
forall a. HasCallStack => a
__IMPOSSIBLE__ CoverageCheck
forall a. HasCallStack => a
__IMPOSSIBLE__ Name
forall a. HasCallStack => a
__IMPOSSIBLE__
                [Name
-> Bool
-> LHS
-> RHS' Expr
-> WhereClause' [Declaration]
-> [Clause]
-> Clause
C.Clause Name
x (Bool
ca Bool -> Bool -> Bool
forall a. Boolean a => a -> a -> a
|| Bool
catchall) LHS
lhs (Expr -> RHS' Expr
forall e. e -> RHS' e
C.RHS Expr
rhs) WhereClause' [Declaration]
forall decls. WhereClause' decls
NoWhere []]
              ]
          where
            definedName :: Pattern -> Maybe Name
definedName (C.IdentP Bool
_ (C.QName Name
x)) = Name -> Maybe Name
forall a. a -> Maybe a
Just Name
x
            definedName C.IdentP{}             = Maybe Name
forall a. Maybe a
Nothing
            definedName (C.RawAppP Range
_ (List2 Pattern
p Pattern
_ [Pattern]
_)) = Pattern -> Maybe Name
definedName Pattern
p
            definedName (C.ParenP Range
_ Pattern
p)         = Pattern -> Maybe Name
definedName Pattern
p
            definedName C.WildP{}              = Maybe Name
forall a. Maybe a
Nothing   -- for instance let _ + x = x in ... (not allowed)
            definedName C.AbsurdP{}            = Maybe Name
forall a. Maybe a
Nothing
            definedName C.AsP{}                = Maybe Name
forall a. Maybe a
Nothing
            definedName C.DotP{}               = Maybe Name
forall a. Maybe a
Nothing
            definedName C.EqualP{}             = Maybe Name
forall a. Maybe a
Nothing
            definedName C.LitP{}               = Maybe Name
forall a. Maybe a
Nothing
            definedName C.RecP{}               = Maybe Name
forall a. Maybe a
Nothing
            definedName C.QuoteP{}             = Maybe Name
forall a. Maybe a
Nothing
            definedName C.HiddenP{}            = Maybe Name
forall a. Maybe a
Nothing -- Not impossible, see issue #2291
            definedName C.InstanceP{}          = Maybe Name
forall a. Maybe a
Nothing
            definedName C.WithP{}              = Maybe Name
forall a. Maybe a
Nothing
            definedName C.AppP{}               = Maybe Name
forall a. Maybe a
Nothing -- Not impossible, see issue #4586
            definedName C.OpAppP{}             = Maybe Name
forall a. HasCallStack => a
__IMPOSSIBLE__
            definedName C.EllipsisP{}          = Maybe Name
forall a. Maybe a
Nothing -- Not impossible, see issue #3937

    -- You can't open public in a let
    NiceOpen Range
r QName
x ImportDirective
dirs -> do
      Maybe KwRange -> (KwRange -> TCMT IO ()) -> TCMT IO ()
forall (m :: * -> *) a. Monad m => Maybe a -> (a -> m ()) -> m ()
whenJust (ImportDirective -> Maybe KwRange
forall n m. ImportDirective' n m -> Maybe KwRange
publicOpen ImportDirective
dirs) ((KwRange -> TCMT IO ()) -> TCMT IO ())
-> (KwRange -> TCMT IO ()) -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ \KwRange
r -> KwRange -> TCMT IO () -> TCMT IO ()
forall (m :: * -> *) x a.
(MonadTrace m, HasRange x) =>
x -> m a -> m a
setCurrentRange KwRange
r (TCMT IO () -> TCMT IO ()) -> TCMT IO () -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ Warning -> TCMT IO ()
forall (m :: * -> *).
(HasCallStack, MonadWarning m) =>
Warning -> m ()
warning Warning
UselessPublic
      m    <- OldModuleName -> ScopeM (AbsOfCon OldModuleName)
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract (QName -> OldModuleName
OldModuleName QName
x)
      adir <- openModule_ LetOpenModule x dirs
      let minfo = ModuleInfo
            { minfoRange :: Range
minfoRange  = Range
r
            , minfoAsName :: Maybe Name
minfoAsName = Maybe Name
forall a. Maybe a
Nothing
            , minfoAsTo :: Range
minfoAsTo   = ImportDirective -> Range
renamingRange ImportDirective
dirs
            , minfoOpenShort :: Maybe OpenShortHand
minfoOpenShort = Maybe OpenShortHand
forall a. Maybe a
Nothing
            , minfoDirective :: Maybe ImportDirective
minfoDirective = ImportDirective -> Maybe ImportDirective
forall a. a -> Maybe a
Just ImportDirective
dirs
            }
      return $ singleton $ A.LetOpen minfo m adir

    NiceModuleMacro Range
r Access
p Erased
erased Name
x ModuleApplication
modapp OpenShortHand
open ImportDirective
dir -> do
      Maybe KwRange -> (KwRange -> TCMT IO ()) -> TCMT IO ()
forall (m :: * -> *) a. Monad m => Maybe a -> (a -> m ()) -> m ()
whenJust (ImportDirective -> Maybe KwRange
forall n m. ImportDirective' n m -> Maybe KwRange
publicOpen ImportDirective
dir) ((KwRange -> TCMT IO ()) -> TCMT IO ())
-> (KwRange -> TCMT IO ()) -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ \ KwRange
r -> KwRange -> TCMT IO () -> TCMT IO ()
forall (m :: * -> *) x a.
(MonadTrace m, HasRange x) =>
x -> m a -> m a
setCurrentRange KwRange
r (TCMT IO () -> TCMT IO ()) -> TCMT IO () -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ Warning -> TCMT IO ()
forall (m :: * -> *).
(HasCallStack, MonadWarning m) =>
Warning -> m ()
warning Warning
UselessPublic
      -- Andreas, 2014-10-09, Issue 1299: module macros in lets need
      -- to be private
      LetBinding -> List1 LetBinding
forall el coll. Singleton el coll => el -> coll
singleton (LetBinding -> List1 LetBinding)
-> ScopeM LetBinding -> TCMT IO (List1 LetBinding)
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> (ModuleInfo
 -> Erased
 -> ModuleName
 -> ModuleApplication
 -> ScopeCopyInfo
 -> ImportDirective
 -> LetBinding)
-> OpenKind
-> Range
-> Access
-> Erased
-> Name
-> ModuleApplication
-> OpenShortHand
-> ImportDirective
-> ScopeM LetBinding
forall a.
(ToConcrete a, Pretty (ConOfAbs a)) =>
(ModuleInfo
 -> Erased
 -> ModuleName
 -> ModuleApplication
 -> ScopeCopyInfo
 -> ImportDirective
 -> a)
-> OpenKind
-> Range
-> Access
-> Erased
-> Name
-> ModuleApplication
-> OpenShortHand
-> ImportDirective
-> ScopeM a
checkModuleMacro ModuleInfo
-> Erased
-> ModuleName
-> ModuleApplication
-> ScopeCopyInfo
-> ImportDirective
-> LetBinding
LetApply OpenKind
LetOpenModule Range
r
                      Access
privateAccessInserted Erased
erased Name
x ModuleApplication
modapp OpenShortHand
open ImportDirective
dir

    NiceDeclaration
_ -> Maybe NotAValidLetBinding -> TCMT IO (List1 LetBinding)
forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
Maybe NotAValidLetBinding -> m a
notAValidLetBinding Maybe NotAValidLetBinding
forall a. Maybe a
Nothing

    where
      letToAbstract :: Clause -> TCMT IO (QName, Expr)
letToAbstract (C.Clause Name
top Bool
_catchall (C.LHS Pattern
p [] []) RHS' Expr
rhs0 WhereClause' [Declaration]
wh []) = do
        WhereClause' [Declaration] -> TCMT IO ()
noWhereInLetBinding WhereClause' [Declaration]
wh
        rhs <- RHS' Expr -> TCMT IO Expr
letBindingMustHaveRHS RHS' Expr
rhs0
        (x, args) <- do
          res <- setCurrentRange p $ parseLHS NoDisplayLHS (C.QName top) p
          case res of
            C.LHSHead QName
x [NamedArg Pattern]
args -> (QName, [NamedArg Pattern]) -> TCMT IO (QName, [NamedArg Pattern])
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return (QName
x, [NamedArg Pattern]
args)
            C.LHSProj{}      -> TCMT IO (QName, [NamedArg Pattern])
forall a. HasCallStack => a
__IMPOSSIBLE__  -- notAValidLetBinding $ Just CopatternsNotAllowed
            C.LHSWith{}      -> TCMT IO (QName, [NamedArg Pattern])
forall a. HasCallStack => a
__IMPOSSIBLE__  -- notAValidLetBinding $ Just WithPatternsNotAllowed
            C.LHSEllipsis{}  -> TCMT IO (QName, [NamedArg Pattern])
forall a. HasCallStack => a
__IMPOSSIBLE__  -- notAValidLetBinding $ Just EllipsisNotAllowed

        e <- localToAbstract args $ \AbsOfCon [NamedArg Pattern]
args -> do
          TCMT IO ()
bindVarsToBind
          -- Make sure to unbind the function name in the RHS, since lets are non-recursive.
          rhs <- Name -> ScopeM Expr -> ScopeM Expr
forall a. Name -> ScopeM a -> ScopeM a
unbindVariable Name
top (ScopeM Expr -> ScopeM Expr) -> ScopeM Expr -> ScopeM Expr
forall a b. (a -> b) -> a -> b
$ Expr -> ScopeM (AbsOfCon Expr)
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract Expr
rhs
          foldM lambda rhs (reverse args)  -- just reverse because these are DomainFree
        return (x, e)

      letToAbstract Clause
_ = Maybe NotAValidLetBinding -> TCMT IO (QName, Expr)
forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
Maybe NotAValidLetBinding -> m a
notAValidLetBinding Maybe NotAValidLetBinding
forall a. Maybe a
Nothing

      -- These patterns all have a chance of being accepted in a lambda:
      allowedPat :: Pattern' e -> Bool
allowedPat A.VarP{}      = Bool
True
      allowedPat A.ConP{}      = Bool
True
      allowedPat A.WildP{}     = Bool
True
      allowedPat (A.AsP PatInfo
_ BindName
_ Pattern' e
x) = Pattern' e -> Bool
allowedPat Pattern' e
x
      allowedPat (A.RecP ConPatInfo
_ [FieldAssignment' (Pattern' e)]
as) = (FieldAssignment' (Pattern' e) -> Bool)
-> [FieldAssignment' (Pattern' e)] -> Bool
forall (t :: * -> *) a. Foldable t => (a -> Bool) -> t a -> Bool
all (Pattern' e -> Bool
allowedPat (Pattern' e -> Bool)
-> (FieldAssignment' (Pattern' e) -> Pattern' e)
-> FieldAssignment' (Pattern' e)
-> Bool
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Lens' (FieldAssignment' (Pattern' e)) (Pattern' e)
-> FieldAssignment' (Pattern' e) -> Pattern' e
forall o (m :: * -> *) i. MonadReader o m => Lens' o i -> m i
view (Pattern' e -> f (Pattern' e))
-> FieldAssignment' (Pattern' e)
-> f (FieldAssignment' (Pattern' e))
forall a (f :: * -> *).
Functor f =>
(a -> f a) -> FieldAssignment' a -> f (FieldAssignment' a)
Lens' (FieldAssignment' (Pattern' e)) (Pattern' e)
exprFieldA) [FieldAssignment' (Pattern' e)]
as
      allowedPat (A.PatternSynP PatInfo
_ AmbiguousQName
_ NAPs e
as) = (NamedArg (Pattern' e) -> Bool) -> NAPs e -> Bool
forall (t :: * -> *) a. Foldable t => (a -> Bool) -> t a -> Bool
all (Pattern' e -> Bool
allowedPat (Pattern' e -> Bool)
-> (NamedArg (Pattern' e) -> Pattern' e)
-> NamedArg (Pattern' e)
-> Bool
forall b c a. (b -> c) -> (a -> b) -> a -> c
. NamedArg (Pattern' e) -> Pattern' e
forall a. NamedArg a -> a
namedArg) NAPs e
as

      -- These have no chance:
      allowedPat A.AbsurdP{}     = Bool
False
      allowedPat A.ProjP{}       = Bool
False
      allowedPat A.DefP{}        = Bool
False
      allowedPat A.EqualP{}      = Bool
False
      allowedPat A.WithP{}       = Bool
False
      allowedPat A.DotP{}        = Bool
False
      allowedPat A.LitP{}        = Bool
False

      patternName :: Pattern' e -> Maybe BindName
patternName (A.VarP BindName
bn)    = BindName -> Maybe BindName
forall a. a -> Maybe a
Just BindName
bn
      patternName (A.AsP PatInfo
_ BindName
bn Pattern' e
_) = BindName -> Maybe BindName
forall a. a -> Maybe a
Just BindName
bn
      patternName Pattern' e
_ = Maybe BindName
forall a. Maybe a
Nothing

      -- Named patterns not allowed in let definitions
      lambda :: A.Expr -> A.NamedArg (A.Pattern' C.Expr) -> TCM A.Expr
      lambda :: Expr -> NamedArg (Pattern' Expr) -> ScopeM Expr
lambda Expr
e ai :: NamedArg (Pattern' Expr)
ai@(Arg ArgInfo
info (Named Maybe NamedName
thing Pattern' Expr
pat)) | Pattern' Expr -> Bool
forall {e}. Pattern' e -> Bool
allowedPat Pattern' Expr
pat = do
        let
          i :: ExprInfo
i = Range -> ExprInfo
ExprRange (Pattern' Expr -> Expr -> Range
forall u t. (HasRange u, HasRange t) => u -> t -> Range
fuseRange Pattern' Expr
pat Expr
e)

        pat <- Pattern' Expr -> ScopeM (AbsOfCon (Pattern' Expr))
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract Pattern' Expr
pat

        bn <- case pat of
          A.VarP BindName
bn    -> BindName -> TCMT IO BindName
forall a. a -> TCMT IO a
forall (f :: * -> *) a. Applicative f => a -> f a
pure BindName
bn
          A.AsP PatInfo
_ BindName
bn Pattern
_ -> BindName -> TCMT IO BindName
forall a. a -> TCMT IO a
forall (f :: * -> *) a. Applicative f => a -> f a
pure BindName
bn
          Pattern
_ -> (Name -> BindName) -> ScopeM Name -> TCMT IO BindName
forall a b. (a -> b) -> TCMT IO a -> TCMT IO b
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
fmap Name -> BindName
mkBindName (ScopeM Name -> TCMT IO BindName)
-> (Name -> ScopeM Name) -> Name -> TCMT IO BindName
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Name -> ScopeM Name
freshAbstractName_ (Name -> TCMT IO BindName) -> TCMT IO Name -> TCMT IO BindName
forall (m :: * -> *) a b. Monad m => (a -> m b) -> m a -> m b
=<< Range -> Int -> [Char] -> TCMT IO Name
freshConcreteName (Pattern -> Range
forall a. HasRange a => a -> Range
getRange Pattern
pat) Int
0 [Char]
patternInTeleName

        -- Annoyingly, for the lambdas to be elaborated properly, we
        -- have to generate domainful binders. Domain-free binders can
        -- not be named (or have pattern matching!).
        --
        -- Moreover, we need to avoid generating named patterns that are
        -- like {B = B @ B}.

        let
          pat' = case Pattern
pat of
            A.VarP{} -> Maybe Pattern
forall a. Maybe a
Nothing
            Pattern
pat -> Pattern -> Maybe Pattern
forall a. a -> Maybe a
Just Pattern
pat
          binder = ArgInfo
-> Named NamedName (Binder' BindName)
-> NamedArg (Binder' BindName)
forall e. ArgInfo -> e -> Arg e
Arg ArgInfo
info (Maybe NamedName
-> Binder' BindName -> Named NamedName (Binder' BindName)
forall name a. Maybe name -> a -> Named name a
Named Maybe NamedName
thing (Maybe Pattern -> BinderNameOrigin -> BindName -> Binder' BindName
forall a. Maybe Pattern -> BinderNameOrigin -> a -> Binder' a
A.Binder Maybe Pattern
pat' BinderNameOrigin
InsertedBinderName BindName
bn)) NamedArg (Binder' BindName)
-> [NamedArg (Binder' BindName)]
-> List1 (NamedArg (Binder' BindName))
forall a. a -> [a] -> NonEmpty a
:| []

        pure $ A.Lam i (A.DomainFull (A.TBind (getRange ai) empty binder (A.Underscore empty))) e

      lambda Expr
_ NamedArg (Pattern' Expr)
_ = Maybe NotAValidLetBinding -> ScopeM Expr
forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
Maybe NotAValidLetBinding -> m a
notAValidLetBinding Maybe NotAValidLetBinding
forall a. Maybe a
Nothing

      noWhereInLetBinding :: C.WhereClause -> ScopeM ()
      noWhereInLetBinding :: WhereClause' [Declaration] -> TCMT IO ()
noWhereInLetBinding = \case
        WhereClause' [Declaration]
NoWhere -> () -> TCMT IO ()
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return ()
        WhereClause' [Declaration]
wh -> WhereClause' [Declaration] -> TCMT IO () -> TCMT IO ()
forall (m :: * -> *) x a.
(MonadTrace m, HasRange x) =>
x -> m a -> m a
setCurrentRange WhereClause' [Declaration]
wh (TCMT IO () -> TCMT IO ()) -> TCMT IO () -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ Maybe NotAValidLetBinding -> TCMT IO ()
forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
Maybe NotAValidLetBinding -> m a
notAValidLetBinding (Maybe NotAValidLetBinding -> TCMT IO ())
-> Maybe NotAValidLetBinding -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ NotAValidLetBinding -> Maybe NotAValidLetBinding
forall a. a -> Maybe a
Just NotAValidLetBinding
WhereClausesNotAllowed
      letBindingMustHaveRHS :: C.RHS -> ScopeM C.Expr
      letBindingMustHaveRHS :: RHS' Expr -> TCMT IO Expr
letBindingMustHaveRHS = \case
        C.RHS Expr
e -> Expr -> TCMT IO Expr
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return Expr
e
        RHS' Expr
C.AbsurdRHS -> Maybe NotAValidLetBinding -> TCMT IO Expr
forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
Maybe NotAValidLetBinding -> m a
notAValidLetBinding (Maybe NotAValidLetBinding -> TCMT IO Expr)
-> Maybe NotAValidLetBinding -> TCMT IO Expr
forall a b. (a -> b) -> a -> b
$ NotAValidLetBinding -> Maybe NotAValidLetBinding
forall a. a -> Maybe a
Just NotAValidLetBinding
MissingRHS

      -- Only record patterns allowed, but we do not exclude data constructors here.
      -- They will fail in the type checker.
      checkValidLetPattern :: A.Pattern' e -> ScopeM ()
      checkValidLetPattern :: forall e. Pattern' e -> TCMT IO ()
checkValidLetPattern Pattern' e
a = Bool -> TCMT IO () -> TCMT IO ()
forall b (m :: * -> *). (IsBool b, Monad m) => b -> m () -> m ()
unless (Pattern' e -> Bool
forall {e}. Pattern' e -> Bool
allowedPat Pattern' e
a) do
        Maybe NotAValidLetBinding -> TCMT IO ()
forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
Maybe NotAValidLetBinding -> m a
notAValidLetBinding (Maybe NotAValidLetBinding -> TCMT IO ())
-> Maybe NotAValidLetBinding -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ NotAValidLetBinding -> Maybe NotAValidLetBinding
forall a. a -> Maybe a
Just NotAValidLetBinding
NotAValidLetPattern


instance ToAbstract NiceDeclaration where
  type AbsOfCon NiceDeclaration = A.Declaration

  toAbstract :: NiceDeclaration -> ScopeM (AbsOfCon NiceDeclaration)
toAbstract NiceDeclaration
d = ScopeM [Declaration] -> TCMT IO Declaration
forall (m :: * -> *).
ReadTCState m =>
m [Declaration] -> m Declaration
annotateDecls (ScopeM [Declaration] -> TCMT IO Declaration)
-> ScopeM [Declaration] -> TCMT IO Declaration
forall a b. (a -> b) -> a -> b
$
    [Char]
-> Int -> [[Char]] -> ScopeM [Declaration] -> ScopeM [Declaration]
forall a (m :: * -> *) c.
(TraceS a, MonadDebug m) =>
[Char] -> Int -> a -> m c -> m c
forall (m :: * -> *) c.
MonadDebug m =>
[Char] -> Int -> [[Char]] -> m c -> m c
traceS [Char]
"scope.decl.trace" Int
50
      [ [Char]
"scope checking declaration"
      , [Char]
"  " [Char] -> [Char] -> [Char]
forall a. [a] -> [a] -> [a]
++  NiceDeclaration -> [Char]
forall a. Pretty a => a -> [Char]
prettyShow NiceDeclaration
d
      ] (ScopeM [Declaration] -> ScopeM [Declaration])
-> ScopeM [Declaration] -> ScopeM [Declaration]
forall a b. (a -> b) -> a -> b
$
    [Char]
-> Int -> [[Char]] -> ScopeM [Declaration] -> ScopeM [Declaration]
forall a (m :: * -> *) c.
(TraceS a, MonadDebug m) =>
[Char] -> Int -> a -> m c -> m c
forall (m :: * -> *) c.
MonadDebug m =>
[Char] -> Int -> [[Char]] -> m c -> m c
traceS [Char]
"scope.decl.trace" Int
80  -- keep this debug message for testing issue #4016
      [ [Char]
"scope checking declaration (raw)"
      , [Char]
"  " [Char] -> [Char] -> [Char]
forall a. [a] -> [a] -> [a]
++  NiceDeclaration -> [Char]
forall a. Show a => a -> [Char]
show NiceDeclaration
d
      ] (ScopeM [Declaration] -> ScopeM [Declaration])
-> ScopeM [Declaration] -> ScopeM [Declaration]
forall a b. (a -> b) -> a -> b
$
    Call -> ScopeM [Declaration] -> ScopeM [Declaration]
forall a. Call -> TCMT IO a -> TCMT IO a
forall (m :: * -> *) a. MonadTrace m => Call -> m a -> m a
traceCall (NiceDeclaration -> Call
ScopeCheckDeclaration NiceDeclaration
d) (ScopeM [Declaration] -> ScopeM [Declaration])
-> ScopeM [Declaration] -> ScopeM [Declaration]
forall a b. (a -> b) -> a -> b
$
    -- Andreas, 2015-10-05, Issue 1677:
    -- We record in the environment whether we are scope checking an
    -- abstract definition.  This way, we can propagate this attribute
    -- the extended lambdas.
    Maybe IsAbstract
-> (ScopeM [Declaration] -> ScopeM [Declaration])
-> (IsAbstract -> ScopeM [Declaration] -> ScopeM [Declaration])
-> ScopeM [Declaration]
-> ScopeM [Declaration]
forall a b. Maybe a -> b -> (a -> b) -> b
caseMaybe (NiceDeclaration -> Maybe IsAbstract
niceHasAbstract NiceDeclaration
d) ScopeM [Declaration] -> ScopeM [Declaration]
forall a. a -> a
id (\ IsAbstract
a -> (TCEnv -> TCEnv) -> ScopeM [Declaration] -> ScopeM [Declaration]
forall a. (TCEnv -> TCEnv) -> TCMT IO a -> TCMT IO a
forall (m :: * -> *) a.
MonadTCEnv m =>
(TCEnv -> TCEnv) -> m a -> m a
localTC ((TCEnv -> TCEnv) -> ScopeM [Declaration] -> ScopeM [Declaration])
-> (TCEnv -> TCEnv) -> ScopeM [Declaration] -> ScopeM [Declaration]
forall a b. (a -> b) -> a -> b
$ \ TCEnv
e -> TCEnv
e { envAbstractMode = aDefToMode a }) (ScopeM [Declaration] -> ScopeM [Declaration])
-> ScopeM [Declaration] -> ScopeM [Declaration]
forall a b. (a -> b) -> a -> b
$
    case NiceDeclaration
d of

  -- Axiom (actual postulate)
    C.Axiom Range
r Access
p IsAbstract
a IsInstance
i ArgInfo
rel Name
x Expr
t -> do
      -- check that we do not postulate in --safe mode, unless it is a
      -- builtin module with safe postulates
      TCMT IO Bool -> TCMT IO () -> TCMT IO ()
forall (m :: * -> *). Monad m => m Bool -> m () -> m ()
whenM ((CommandLineOptions -> Bool
forall a. LensSafeMode a => a -> Bool
Lens.getSafeMode (CommandLineOptions -> Bool)
-> TCMT IO CommandLineOptions -> TCMT IO Bool
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> TCMT IO CommandLineOptions
forall (m :: * -> *). HasOptions m => m CommandLineOptions
commandLineOptions) TCMT IO Bool -> TCMT IO Bool -> TCMT IO Bool
forall (m :: * -> *). Monad m => m Bool -> m Bool -> m Bool
`and2M`
             (Bool -> Bool
not (Bool -> Bool) -> TCMT IO Bool -> TCMT IO Bool
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> ([Char] -> TCMT IO Bool
forall (m :: * -> *). MonadIO m => [Char] -> m Bool
Lens.isBuiltinModuleWithSafePostulates ([Char] -> TCMT IO Bool)
-> (AbsolutePath -> [Char]) -> AbsolutePath -> TCMT IO Bool
forall b c a. (b -> c) -> (a -> b) -> a -> c
. AbsolutePath -> [Char]
filePath (AbsolutePath -> TCMT IO Bool)
-> TCMT IO AbsolutePath -> TCMT IO Bool
forall (m :: * -> *) a b. Monad m => (a -> m b) -> m a -> m b
=<< TCMT IO AbsolutePath
forall (m :: * -> *). MonadTCEnv m => m AbsolutePath
getCurrentPath)))
            (Warning -> TCMT IO ()
forall (m :: * -> *).
(HasCallStack, MonadWarning m) =>
Warning -> m ()
warning (Warning -> TCMT IO ()) -> Warning -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ Name -> Warning
SafeFlagPostulate Name
x)
      -- check the postulate
      Declaration -> [Declaration]
forall el coll. Singleton el coll => el -> coll
singleton (Declaration -> [Declaration])
-> TCMT IO Declaration -> ScopeM [Declaration]
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> KindOfName -> NiceDeclaration -> TCMT IO Declaration
toAbstractNiceAxiom KindOfName
AxiomName NiceDeclaration
d

    C.NiceGeneralize Range
r Access
p ArgInfo
i TacticAttribute
tac Name
x Expr
t -> do
      [Char] -> Int -> [Char] -> TCMT IO ()
forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> [Char] -> m ()
reportSLn [Char]
"scope.decl" Int
30 ([Char] -> TCMT IO ()) -> [Char] -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ [Char]
"found nice generalize: " [Char] -> [Char] -> [Char]
forall a. [a] -> [a] -> [a]
++ Name -> [Char]
forall a. Pretty a => a -> [Char]
prettyShow Name
x
      tac <- (Expr -> ScopeM Expr)
-> TacticAttribute -> TCMT IO (TacticAttribute' Expr)
forall (t :: * -> *) (f :: * -> *) a b.
(Traversable t, Applicative f) =>
(a -> f b) -> t a -> f (t b)
forall (f :: * -> *) a b.
Applicative f =>
(a -> f b) -> TacticAttribute' a -> f (TacticAttribute' b)
traverse (Precedence -> Expr -> ScopeM (AbsOfCon Expr)
forall c. ToAbstract c => Precedence -> c -> ScopeM (AbsOfCon c)
toAbstractCtx Precedence
TopCtx) TacticAttribute
tac
      t_ <- toAbstractCtx TopCtx t
      let (s, t) = unGeneralized t_
      reportSLn "scope.decl" 50 $ "generalizations: " ++ show (Set.toList s, t)
      f <- getConcreteFixity x
      y <- freshAbstractQName f x
      bindName p GeneralizeName x y
      let info = (Name -> Fixity' -> Access -> IsAbstract -> Range -> DefInfo' Any
forall t.
Name -> Fixity' -> Access -> IsAbstract -> Range -> DefInfo' t
mkDefInfo Name
x Fixity'
f Access
p IsAbstract
ConcreteDef Range
r) { defTactic = tac }
      return [A.Generalize s info i y t]

  -- Fields
    C.NiceField Range
r Access
p IsAbstract
a IsInstance
i TacticAttribute
tac Name
x Arg Expr
t -> do
      Bool -> TCMT IO () -> TCMT IO ()
forall b (m :: * -> *). (IsBool b, Monad m) => b -> m () -> m ()
unless (Access
p Access -> Access -> Bool
forall a. Eq a => a -> a -> Bool
== Access
PublicAccess) (TCMT IO () -> TCMT IO ()) -> TCMT IO () -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ TypeError -> TCMT IO ()
forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
TypeError -> m a
typeError TypeError
PrivateRecordField
      -- Interaction points for record fields have already been introduced
      -- when checking the type of the record constructor.
      -- To avoid introducing interaction points (IP) twice, we turn
      -- all question marks to underscores.  (See issue 1138.)
      let maskIP :: Expr -> Expr
maskIP (C.QuestionMark Range
r Maybe Int
_) = Range -> Maybe [Char] -> Expr
C.Underscore Range
r Maybe [Char]
forall a. Maybe a
Nothing
          maskIP Expr
e                     = Expr
e
      tac <- (Expr -> ScopeM Expr)
-> TacticAttribute -> TCMT IO (TacticAttribute' Expr)
forall (t :: * -> *) (f :: * -> *) a b.
(Traversable t, Applicative f) =>
(a -> f b) -> t a -> f (t b)
forall (f :: * -> *) a b.
Applicative f =>
(a -> f b) -> TacticAttribute' a -> f (TacticAttribute' b)
traverse (Precedence -> Expr -> ScopeM (AbsOfCon Expr)
forall c. ToAbstract c => Precedence -> c -> ScopeM (AbsOfCon c)
toAbstractCtx Precedence
TopCtx) TacticAttribute
tac
      t' <- toAbstractCtx TopCtx $ mapExpr maskIP t
      f  <- getConcreteFixity x
      y  <- freshAbstractQName f x
      -- Andreas, 2018-06-09 issue #2170
      -- We want dependent irrelevance without irrelevant projections,
      -- thus, do not disable irrelevant projections via the scope checker.
      -- irrProj <- optIrrelevantProjections <$> pragmaOptions
      -- unless (isIrrelevant t && not irrProj) $
      --   -- Andreas, 2010-09-24: irrelevant fields are not in scope
      --   -- this ensures that projections out of irrelevant fields cannot occur
      --   -- Ulf: unless you turn on --irrelevant-projections
      bindName p FldName x y
      let info = (Name
-> Fixity'
-> Access
-> IsAbstract
-> IsInstance
-> IsMacro
-> Range
-> DefInfo' Any
forall t.
Name
-> Fixity'
-> Access
-> IsAbstract
-> IsInstance
-> IsMacro
-> Range
-> DefInfo' t
mkDefInfoInstance Name
x Fixity'
f Access
p IsAbstract
a IsInstance
i IsMacro
NotMacroDef Range
r) { defTactic = tac }
      return [ A.Field info y t' ]

  -- Primitive function
    PrimitiveFunction Range
r Access
p IsAbstract
a Name
x Arg Expr
t -> ScopeM [Declaration] -> ScopeM [Declaration]
forall a. ScopeM a -> ScopeM a
notAffectedByOpaque (ScopeM [Declaration] -> ScopeM [Declaration])
-> ScopeM [Declaration] -> ScopeM [Declaration]
forall a b. (a -> b) -> a -> b
$ do
      t' <- (Expr -> ScopeM Expr) -> Arg Expr -> TCMT IO (Arg Expr)
forall (t :: * -> *) (f :: * -> *) a b.
(Traversable t, Applicative f) =>
(a -> f b) -> t a -> f (t b)
forall (f :: * -> *) a b.
Applicative f =>
(a -> f b) -> Arg a -> f (Arg b)
traverse (Precedence -> Expr -> ScopeM (AbsOfCon Expr)
forall c. ToAbstract c => Precedence -> c -> ScopeM (AbsOfCon c)
toAbstractCtx Precedence
TopCtx) Arg Expr
t
      f  <- getConcreteFixity x
      y  <- freshAbstractQName f x
      bindName p PrimName x y
      unfoldFunction y
      let di = Name -> Fixity' -> Access -> IsAbstract -> Range -> DefInfo
forall t.
Name -> Fixity' -> Access -> IsAbstract -> Range -> DefInfo' t
mkDefInfo Name
x Fixity'
f Access
p IsAbstract
a Range
r
      return [ A.Primitive di y t' ]

  -- Definitions (possibly mutual)
    NiceMutual KwRange
kwr TerminationCheck
tc CoverageCheck
cc PositivityCheck
pc [NiceDeclaration]
ds -> do
      [Char] -> Int -> [Char] -> TCMT IO ()
forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> [Char] -> m ()
reportSLn [Char]
"scope.mutual" Int
40 ([Char]
"starting checking mutual definitions: " [Char] -> [Char] -> [Char]
forall a. [a] -> [a] -> [a]
++ [NiceDeclaration] -> [Char]
forall a. Pretty a => a -> [Char]
prettyShow [NiceDeclaration]
ds)
      ds' <- [NiceDeclaration] -> ScopeM (AbsOfCon [NiceDeclaration])
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract [NiceDeclaration]
ds
      reportSLn "scope.mutual" 40 ("finishing checking mutual definitions")
      -- We only termination check blocks that do not have a measure.
      return [ A.Mutual (MutualInfo tc cc pc (fuseRange kwr ds)) ds' ]

    C.NiceRecSig Range
r Erased
er Access
p IsAbstract
a PositivityCheck
_pc UniverseCheck
_uc Name
x [LamBinding' TypedBinding]
ls Expr
t -> do
      [LamBinding' TypedBinding] -> TCMT IO ()
forall a. EnsureNoLetStms a => a -> TCMT IO ()
ensureNoLetStms [LamBinding' TypedBinding]
ls
      ScopeM [Declaration] -> ScopeM [Declaration]
forall a. ScopeM a -> ScopeM a
withLocalVars (ScopeM [Declaration] -> ScopeM [Declaration])
-> ScopeM [Declaration] -> ScopeM [Declaration]
forall a b. (a -> b) -> a -> b
$ do
        (ls', _) <- ScopeM (GeneralizeTelescope, Expr)
-> ScopeM (GeneralizeTelescope, Expr)
forall a. ScopeM a -> ScopeM a
withCheckNoShadowing (ScopeM (GeneralizeTelescope, Expr)
 -> ScopeM (GeneralizeTelescope, Expr))
-> ScopeM (GeneralizeTelescope, Expr)
-> ScopeM (GeneralizeTelescope, Expr)
forall a b. (a -> b) -> a -> b
$
          -- Minor hack: record types don't have indices so we include t when
          -- computing generalised parameters, but in the type checker any named
          -- generalizable arguments in the sort should be bound variables.
          GenTelAndType -> ScopeM (AbsOfCon GenTelAndType)
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract (Telescope -> Expr -> GenTelAndType
GenTelAndType ((LamBinding' TypedBinding -> TypedBinding)
-> [LamBinding' TypedBinding] -> Telescope
forall a b. (a -> b) -> [a] -> [b]
map LamBinding' TypedBinding -> TypedBinding
makeDomainFull [LamBinding' TypedBinding]
ls) Expr
t)
        t' <- toAbstract t
        f  <- getConcreteFixity x
        x' <- freshAbstractQName f x
        bindName' p RecName (GeneralizedVarsMetadata $ generalizeTelVars ls') x x'
        return [ A.RecSig (mkDefInfo x f p a r) er x' ls' t' ]

    C.NiceDataSig Range
r Erased
er Access
p IsAbstract
a PositivityCheck
pc UniverseCheck
uc Name
x [LamBinding' TypedBinding]
ls Expr
t -> do
        [Char] -> Int -> [Char] -> TCMT IO ()
forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> [Char] -> m ()
reportSLn [Char]
"scope.data.sig" Int
40 ([Char]
"checking DataSig for " [Char] -> [Char] -> [Char]
forall a. [a] -> [a] -> [a]
++ Name -> [Char]
forall a. Pretty a => a -> [Char]
prettyShow Name
x)
        [LamBinding' TypedBinding] -> TCMT IO ()
forall a. EnsureNoLetStms a => a -> TCMT IO ()
ensureNoLetStms [LamBinding' TypedBinding]
ls
        ScopeM [Declaration] -> ScopeM [Declaration]
forall a. ScopeM a -> ScopeM a
withLocalVars (ScopeM [Declaration] -> ScopeM [Declaration])
-> ScopeM [Declaration] -> ScopeM [Declaration]
forall a b. (a -> b) -> a -> b
$ do
          ls' <- TCMT IO GeneralizeTelescope -> TCMT IO GeneralizeTelescope
forall a. ScopeM a -> ScopeM a
withCheckNoShadowing (TCMT IO GeneralizeTelescope -> TCMT IO GeneralizeTelescope)
-> TCMT IO GeneralizeTelescope -> TCMT IO GeneralizeTelescope
forall a b. (a -> b) -> a -> b
$
            GenTel -> ScopeM (AbsOfCon GenTel)
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract (GenTel -> ScopeM (AbsOfCon GenTel))
-> GenTel -> ScopeM (AbsOfCon GenTel)
forall a b. (a -> b) -> a -> b
$ Telescope -> GenTel
GenTel (Telescope -> GenTel) -> Telescope -> GenTel
forall a b. (a -> b) -> a -> b
$ (LamBinding' TypedBinding -> TypedBinding)
-> [LamBinding' TypedBinding] -> Telescope
forall a b. (a -> b) -> [a] -> [b]
map LamBinding' TypedBinding -> TypedBinding
makeDomainFull [LamBinding' TypedBinding]
ls
          t'  <- toAbstract $ C.Generalized t
          f  <- getConcreteFixity x
          x' <- freshAbstractQName f x
          mErr <- bindName'' p DataName (GeneralizedVarsMetadata $ generalizeTelVars ls') x x'
          whenJust mErr $ \case
            err :: TypeError
err@(ClashingDefinition QName
cn QName
an Maybe NiceDeclaration
_) -> do
              QName -> ScopeM ResolvedName
resolveName (Name -> QName
C.QName Name
x) ScopeM ResolvedName -> (ResolvedName -> TCMT IO ()) -> TCMT IO ()
forall a b. TCMT IO a -> (a -> TCMT IO b) -> TCMT IO b
forall (m :: * -> *) a b. Monad m => m a -> (a -> m b) -> m b
>>= \case
                -- #4435: if a data type signature causes a ClashingDefinition error, and if
                -- the data type name is bound to an Axiom, then the error may be caused by
                -- the illegal type signature. Convert the NiceDataSig into a NiceDataDef
                -- (which removes the type signature) and suggest it as a possible fix.
                DefinedName Access
p AbstractName
ax Suffix
NoSuffix | AbstractName -> KindOfName
anameKind AbstractName
ax KindOfName -> KindOfName -> Bool
forall a. Eq a => a -> a -> Bool
== KindOfName
AxiomName -> do
                  let suggestion :: NiceDeclaration
suggestion = Range
-> Origin
-> IsAbstract
-> PositivityCheck
-> UniverseCheck
-> Name
-> [LamBinding' TypedBinding]
-> [NiceDeclaration]
-> NiceDeclaration
NiceDataDef Range
r Origin
Inserted IsAbstract
a PositivityCheck
pc UniverseCheck
uc Name
x [LamBinding' TypedBinding]
ls []
                  TypeError -> TCMT IO ()
forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
TypeError -> m a
typeError (TypeError -> TCMT IO ()) -> TypeError -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ QName -> QName -> Maybe NiceDeclaration -> TypeError
ClashingDefinition QName
cn QName
an (NiceDeclaration -> Maybe NiceDeclaration
forall a. a -> Maybe a
Just NiceDeclaration
suggestion)
                ResolvedName
_ -> TypeError -> TCMT IO ()
forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
TypeError -> m a
typeError TypeError
err
            TypeError
otherErr -> TypeError -> TCMT IO ()
forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
TypeError -> m a
typeError TypeError
otherErr
          return [ A.DataSig (mkDefInfo x f p a r) er x' ls' t' ]

  -- Type signatures
    C.FunSig Range
r Access
p IsAbstract
a IsInstance
i IsMacro
m ArgInfo
rel TerminationCheck
_ CoverageCheck
_ Name
x Expr
t -> do
        let kind :: KindOfName
kind = if IsMacro
m IsMacro -> IsMacro -> Bool
forall a. Eq a => a -> a -> Bool
== IsMacro
MacroDef then KindOfName
MacroName else KindOfName
FunName
        Declaration -> [Declaration]
forall el coll. Singleton el coll => el -> coll
singleton (Declaration -> [Declaration])
-> TCMT IO Declaration -> ScopeM [Declaration]
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> KindOfName -> NiceDeclaration -> TCMT IO Declaration
toAbstractNiceAxiom KindOfName
kind (Range
-> Access
-> IsAbstract
-> IsInstance
-> ArgInfo
-> Name
-> Expr
-> NiceDeclaration
C.Axiom Range
r Access
p IsAbstract
a IsInstance
i ArgInfo
rel Name
x Expr
t)

  -- Function definitions
    C.FunDef Range
r [Declaration]
ds IsAbstract
a IsInstance
i TerminationCheck
_ CoverageCheck
_ Name
x [Clause]
cs -> do
        Int -> [Char] -> TCMT IO ()
printLocals Int
30 ([Char] -> TCMT IO ()) -> [Char] -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ [Char]
"checking def " [Char] -> [Char] -> [Char]
forall a. [a] -> [a] -> [a]
++ Name -> [Char]
forall a. Pretty a => a -> [Char]
prettyShow Name
x
        (x',cs) <- (OldName Name, [Clause])
-> ScopeM (AbsOfCon (OldName Name, [Clause]))
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract (Name -> OldName Name
forall a. a -> OldName a
OldName Name
x,[Clause]
cs)
        -- Andreas, 2017-12-04 the name must reside in the current module
        unlessM ((A.qnameModule x' ==) <$> getCurrentModule) $
          __IMPOSSIBLE__
        f <- getConcreteFixity x

        unfoldFunction x'
        di <- updateDefInfoOpacity (mkDefInfoInstance x f PublicAccess a i NotMacroDef r)
        return [ A.FunDef di x' cs ]

  -- Uncategorized function clauses
    C.NiceFunClause Range
_ Access
_ IsAbstract
_ TerminationCheck
_ CoverageCheck
_ Bool
_ (C.FunClause LHS
lhs RHS' Expr
_ WhereClause' [Declaration]
_ Bool
_) ->
      TypeError -> ScopeM [Declaration]
forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
TypeError -> m a
typeError (TypeError -> ScopeM [Declaration])
-> TypeError -> ScopeM [Declaration]
forall a b. (a -> b) -> a -> b
$ MissingTypeSignatureInfo -> TypeError
MissingTypeSignature (MissingTypeSignatureInfo -> TypeError)
-> MissingTypeSignatureInfo -> TypeError
forall a b. (a -> b) -> a -> b
$ LHS -> MissingTypeSignatureInfo
MissingFunctionSignature LHS
lhs
    C.NiceFunClause{} -> ScopeM [Declaration]
forall a. HasCallStack => a
__IMPOSSIBLE__

  -- Data definitions
    C.NiceDataDef Range
r Origin
o IsAbstract
a PositivityCheck
_ UniverseCheck
uc Name
x [LamBinding' TypedBinding]
pars [NiceDeclaration]
cons -> ScopeM [Declaration] -> ScopeM [Declaration]
forall a. ScopeM a -> ScopeM a
notAffectedByOpaque (ScopeM [Declaration] -> ScopeM [Declaration])
-> ScopeM [Declaration] -> ScopeM [Declaration]
forall a b. (a -> b) -> a -> b
$ do
        [Char] -> Int -> [Char] -> TCMT IO ()
forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> [Char] -> m ()
reportSLn [Char]
"scope.data.def" Int
40 ([Char]
"checking " [Char] -> [Char] -> [Char]
forall a. [a] -> [a] -> [a]
++ Origin -> [Char]
forall a. Show a => a -> [Char]
show Origin
o [Char] -> [Char] -> [Char]
forall a. [a] -> [a] -> [a]
++ [Char]
" DataDef for " [Char] -> [Char] -> [Char]
forall a. [a] -> [a] -> [a]
++ Name -> [Char]
forall a. Pretty a => a -> [Char]
prettyShow Name
x)
        (p, ax) <- QName -> ScopeM ResolvedName
resolveName (Name -> QName
C.QName Name
x) ScopeM ResolvedName
-> (ResolvedName -> TCMT IO (Access, AbstractName))
-> TCMT IO (Access, AbstractName)
forall a b. TCMT IO a -> (a -> TCMT IO b) -> TCMT IO b
forall (m :: * -> *) a b. Monad m => m a -> (a -> m b) -> m b
>>= \case
          DefinedName Access
p AbstractName
ax Suffix
NoSuffix -> do
            Name -> KindOfName -> AbstractName -> TCMT IO ()
clashUnless Name
x KindOfName
DataName AbstractName
ax  -- Andreas 2019-07-07, issue #3892
            AbstractName -> TCMT IO ()
forall a. LivesInCurrentModule a => a -> TCMT IO ()
livesInCurrentModule AbstractName
ax  -- Andreas, 2017-12-04, issue #2862
            Name -> AbstractName -> TCMT IO ()
clashIfModuleAlreadyDefinedInCurrentModule Name
x AbstractName
ax
            (Access, AbstractName) -> TCMT IO (Access, AbstractName)
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return (Access
p, AbstractName
ax)
          ResolvedName
_ -> TypeError -> TCMT IO (Access, AbstractName)
forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
TypeError -> m a
typeError (TypeError -> TCMT IO (Access, AbstractName))
-> TypeError -> TCMT IO (Access, AbstractName)
forall a b. (a -> b) -> a -> b
$ MissingTypeSignatureInfo -> TypeError
MissingTypeSignature (MissingTypeSignatureInfo -> TypeError)
-> MissingTypeSignatureInfo -> TypeError
forall a b. (a -> b) -> a -> b
$ Name -> MissingTypeSignatureInfo
MissingDataSignature Name
x
        ensureNoLetStms pars
        withLocalVars $ do
          gvars <- bindGeneralizablesIfInserted o ax
          -- Check for duplicate constructors
          do cs <- mapM conName cons
             List1.unlessNull (duplicates cs) $ \ List1 Name
dups -> do
               let bad :: [Name]
bad = (Name -> Bool) -> [Name] -> [Name]
forall a. (a -> Bool) -> [a] -> [a]
filter (Name -> List1 Name -> Bool
forall a. Eq a => a -> NonEmpty a -> Bool
forall (t :: * -> *) a. (Foldable t, Eq a) => a -> t a -> Bool
`elem` List1 Name
dups) [Name]
cs
               [Name] -> TCMT IO () -> TCMT IO ()
forall (m :: * -> *) x a.
(MonadTrace m, HasRange x) =>
x -> m a -> m a
setCurrentRange [Name]
bad (TCMT IO () -> TCMT IO ()) -> TCMT IO () -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$
                 TypeError -> TCMT IO ()
forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
TypeError -> m a
typeError (TypeError -> TCMT IO ()) -> TypeError -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ List1 Name -> TypeError
DuplicateConstructors List1 Name
dups

          pars <- catMaybes <$> toAbstract pars
          let x' = AbstractName -> QName
anameName AbstractName
ax
          -- Create the module for the qualified constructors
          checkForModuleClash x -- disallow shadowing previously defined modules
          let m = QName -> ModuleName
qnameToMName QName
x'
          createModule (Just IsDataModule) m
          bindModule p x m  -- make it a proper module
          cons <- toAbstract (map (DataConstrDecl m a p) cons)
          printScope "data" 40 $ "Checked data " ++ prettyShow x
          f <- getConcreteFixity x
          return [ A.DataDef (mkDefInfo x f PublicAccess a r) x' uc (DataDefParams gvars pars) cons ]
      where
        conName :: NiceDeclaration -> TCMT IO Name
conName (C.Axiom Range
_ Access
_ IsAbstract
_ IsInstance
_ ArgInfo
_ Name
c Expr
_) = Name -> TCMT IO Name
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return Name
c
        conName NiceDeclaration
d = NiceDeclaration -> TCMT IO Name
forall a. NiceDeclaration -> ScopeM a
errorNotConstrDecl NiceDeclaration
d

  -- Record definitions (mucho interesting)
    C.NiceRecDef Range
r Origin
o IsAbstract
a PositivityCheck
_ UniverseCheck
uc Name
x [RecordDirective]
directives [LamBinding' TypedBinding]
pars [Declaration]
fields -> ScopeM [Declaration] -> ScopeM [Declaration]
forall a. ScopeM a -> ScopeM a
notAffectedByOpaque (ScopeM [Declaration] -> ScopeM [Declaration])
-> ScopeM [Declaration] -> ScopeM [Declaration]
forall a b. (a -> b) -> a -> b
$ do
      [Char] -> Int -> [Char] -> TCMT IO ()
forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> [Char] -> m ()
reportSLn [Char]
"scope.rec.def" Int
40 ([Char]
"checking " [Char] -> [Char] -> [Char]
forall a. [a] -> [a] -> [a]
++ Origin -> [Char]
forall a. Show a => a -> [Char]
show Origin
o [Char] -> [Char] -> [Char]
forall a. [a] -> [a] -> [a]
++ [Char]
" RecDef for " [Char] -> [Char] -> [Char]
forall a. [a] -> [a] -> [a]
++ Name -> [Char]
forall a. Pretty a => a -> [Char]
prettyShow Name
x)
      -- #3008: Termination pragmas are ignored in records
      WhereOrRecord -> [Declaration] -> TCMT IO ()
forall a. FoldDecl a => WhereOrRecord -> a -> TCMT IO ()
checkNoTerminationPragma WhereOrRecord
InRecordDef [Declaration]
fields
      RecordDirectives ind eta pat cm <- [RecordDirective] -> ScopeM RecordDirectives
gatherRecordDirectives [RecordDirective]
directives
      -- Andreas, 2020-04-19, issue #4560
      -- 'pattern' declaration is incompatible with 'coinductive' or 'eta-equality'.
      pat <- case pat of
        Just Range
r
          | Just (Ranged Range
_ Induction
CoInductive) <- Maybe (Ranged Induction)
ind -> Maybe Range
forall a. Maybe a
Nothing Maybe Range -> TCMT IO () -> TCMT IO (Maybe Range)
forall a b. a -> TCMT IO b -> TCMT IO a
forall (f :: * -> *) a b. Functor f => a -> f b -> f a
<$ [Char] -> TCMT IO ()
warn [Char]
"coinductive"
          | Just (Ranged Range
_ HasEta0
YesEta)      <- Maybe (Ranged HasEta0)
eta -> Maybe Range
forall a. Maybe a
Nothing Maybe Range -> TCMT IO () -> TCMT IO (Maybe Range)
forall a b. a -> TCMT IO b -> TCMT IO a
forall (f :: * -> *) a b. Functor f => a -> f b -> f a
<$ [Char] -> TCMT IO ()
warn [Char]
"eta"
          | Bool
otherwise -> Maybe Range -> TCMT IO (Maybe Range)
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return Maybe Range
pat
          where warn :: [Char] -> TCMT IO ()
warn = Range -> TCMT IO () -> TCMT IO ()
forall (m :: * -> *) x a.
(MonadTrace m, HasRange x) =>
x -> m a -> m a
setCurrentRange Range
r (TCMT IO () -> TCMT IO ())
-> ([Char] -> TCMT IO ()) -> [Char] -> TCMT IO ()
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Warning -> TCMT IO ()
forall (m :: * -> *).
(HasCallStack, MonadWarning m) =>
Warning -> m ()
warning (Warning -> TCMT IO ())
-> ([Char] -> Warning) -> [Char] -> TCMT IO ()
forall b c a. (b -> c) -> (a -> b) -> a -> c
. [Char] -> Warning
UselessPatternDeclarationForRecord
        Maybe Range
Nothing -> Maybe Range -> TCMT IO (Maybe Range)
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return Maybe Range
pat

      (p, ax) <- resolveName (C.QName x) >>= \case
        DefinedName Access
p AbstractName
ax Suffix
NoSuffix -> do
          Name -> KindOfName -> AbstractName -> TCMT IO ()
clashUnless Name
x KindOfName
RecName AbstractName
ax  -- Andreas 2019-07-07, issue #3892
          AbstractName -> TCMT IO ()
forall a. LivesInCurrentModule a => a -> TCMT IO ()
livesInCurrentModule AbstractName
ax  -- Andreas, 2017-12-04, issue #2862
          Name -> AbstractName -> TCMT IO ()
clashIfModuleAlreadyDefinedInCurrentModule Name
x AbstractName
ax
          (Access, AbstractName) -> TCMT IO (Access, AbstractName)
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return (Access
p, AbstractName
ax)
        ResolvedName
_ -> TypeError -> TCMT IO (Access, AbstractName)
forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
TypeError -> m a
typeError (TypeError -> TCMT IO (Access, AbstractName))
-> TypeError -> TCMT IO (Access, AbstractName)
forall a b. (a -> b) -> a -> b
$ MissingTypeSignatureInfo -> TypeError
MissingTypeSignature (MissingTypeSignatureInfo -> TypeError)
-> MissingTypeSignatureInfo -> TypeError
forall a b. (a -> b) -> a -> b
$ Name -> MissingTypeSignatureInfo
MissingRecordSignature Name
x
      ensureNoLetStms pars
      withLocalVars $ do
        gvars <- bindGeneralizablesIfInserted o ax
        -- Check that the generated module doesn't clash with a previously
        -- defined module
        checkForModuleClash x
        pars   <- catMaybes <$> toAbstract pars
        let x' = AbstractName -> QName
anameName AbstractName
ax
        -- We scope check the fields a first time when putting together
        -- the type of the constructor.
        contel <- localToAbstract (RecordConstructorType fields) return
        m0     <- getCurrentModule
        let m = ModuleName -> ModuleName -> ModuleName
A.qualifyM ModuleName
m0 (ModuleName -> ModuleName) -> ModuleName -> ModuleName
forall a b. (a -> b) -> a -> b
$ List1 Name -> ModuleName
mnameFromList1 (List1 Name -> ModuleName) -> List1 Name -> ModuleName
forall a b. (a -> b) -> a -> b
$ Name -> List1 Name
forall el coll. Singleton el coll => el -> coll
singleton (Name -> List1 Name) -> Name -> List1 Name
forall a b. (a -> b) -> a -> b
$ List1 Name -> Name
forall a. NonEmpty a -> a
List1.last (List1 Name -> Name) -> List1 Name -> Name
forall a b. (a -> b) -> a -> b
$ QName -> List1 Name
qnameToList QName
x'
        printScope "rec" 25 "before record"
        createModule (Just IsRecordModule) m
        -- We scope check the fields a second time, as actual fields.
        afields <- withCurrentModule m $ do
          afields <- toAbstract (Declarations fields)
          printScope "rec" 25 "checked fields"
          return afields
        -- Andreas, 2017-07-13 issue #2642 disallow duplicate fields
        -- Check for duplicate fields. (See "Check for duplicate constructors")
        do let fs :: [C.Name]
               fs = [[Name]] -> [Name]
forall (t :: * -> *) a. Foldable t => t [a] -> [a]
concat ([[Name]] -> [Name]) -> [[Name]] -> [Name]
forall a b. (a -> b) -> a -> b
$ [Declaration] -> (Declaration -> Maybe [Name]) -> [[Name]]
forall a b. [a] -> (a -> Maybe b) -> [b]
forMaybe [Declaration]
fields ((Declaration -> Maybe [Name]) -> [[Name]])
-> (Declaration -> Maybe [Name]) -> [[Name]]
forall a b. (a -> b) -> a -> b
$ \case
                 C.Field KwRange
_ [Declaration]
fs -> [Name] -> Maybe [Name]
forall a. a -> Maybe a
Just ([Name] -> Maybe [Name]) -> [Name] -> Maybe [Name]
forall a b. (a -> b) -> a -> b
$ [Declaration]
fs [Declaration] -> (Declaration -> Name) -> [Name]
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \case
                   -- a Field block only contains field signatures
                   C.FieldSig IsInstance
_ TacticAttribute
_ Name
f Arg Expr
_ -> Name
f
                   Declaration
_ -> Name
forall a. HasCallStack => a
__IMPOSSIBLE__
                 Declaration
_ -> Maybe [Name]
forall a. Maybe a
Nothing
           List1.unlessNull (duplicates fs) $ \ List1 Name
dups -> do
             let bad :: [Name]
bad = (Name -> Bool) -> [Name] -> [Name]
forall a. (a -> Bool) -> [a] -> [a]
filter (Name -> List1 Name -> Bool
forall a. Eq a => a -> NonEmpty a -> Bool
forall (t :: * -> *) a. (Foldable t, Eq a) => a -> t a -> Bool
`elem` List1 Name
dups) [Name]
fs
             [Name] -> TCMT IO () -> TCMT IO ()
forall (m :: * -> *) x a.
(MonadTrace m, HasRange x) =>
x -> m a -> m a
setCurrentRange [Name]
bad (TCMT IO () -> TCMT IO ()) -> TCMT IO () -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$
               TypeError -> TCMT IO ()
forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
TypeError -> m a
typeError (TypeError -> TCMT IO ()) -> TypeError -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ List1 Name -> TypeError
DuplicateFields List1 Name
dups

        bindModule p x m
        let kind = KindOfName
-> (Ranged Induction -> KindOfName)
-> Maybe (Ranged Induction)
-> KindOfName
forall b a. b -> (a -> b) -> Maybe a -> b
maybe KindOfName
ConName (Induction -> KindOfName
conKindOfName (Induction -> KindOfName)
-> (Ranged Induction -> Induction)
-> Ranged Induction
-> KindOfName
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Ranged Induction -> Induction
forall a. Ranged a -> a
rangedThing) Maybe (Ranged Induction)
ind

        cm' <- case cm of
          -- Andreas, 2019-11-11, issue #4189, no longer add record constructor to record module.
          Just (Name
c, IsInstance
_) -> QName -> RecordConName
NamedRecCon (QName -> RecordConName) -> TCMT IO QName -> TCMT IO RecordConName
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> Name -> KindOfName -> IsAbstract -> Access -> TCMT IO QName
bindRecordConstructorName Name
c KindOfName
kind IsAbstract
a Access
p

          -- Amy, 2024-09-25: if the record does not have a named
          -- constructor, then generate the QName here, and record it in
          -- the TC state so that 'Record.constructor' can be resolved.
          Maybe (Name, IsInstance)
Nothing -> do
            -- Technically it doesn't matter with what this name is
            -- qualified since record constructor names have a special
            -- printing rule in lookupQName.
            constr <- ModuleName -> TCMT IO QName -> TCMT IO QName
forall (m :: * -> *) a.
(ReadTCState m, MonadTCState m) =>
ModuleName -> m a -> m a
withCurrentModule ModuleName
m (TCMT IO QName -> TCMT IO QName) -> TCMT IO QName -> TCMT IO QName
forall a b. (a -> b) -> a -> b
$
              Fixity' -> Name -> TCMT IO QName
freshAbstractQName Fixity'
noFixity' (Name -> TCMT IO QName) -> Name -> TCMT IO QName
forall a b. (a -> b) -> a -> b
$ [Char] -> Name
simpleName [Char]
"constructor"
            pure $ FreshRecCon constr

        setRecordConstructor x' (recordConName cm', fmap rangedThing ind)

        let inst = Maybe (Name, IsInstance)
-> IsInstance -> ((Name, IsInstance) -> IsInstance) -> IsInstance
forall a b. Maybe a -> b -> (a -> b) -> b
caseMaybe Maybe (Name, IsInstance)
cm IsInstance
NotInstanceDef (Name, IsInstance) -> IsInstance
forall a b. (a, b) -> b
snd
        printScope "rec" 25 "record complete"
        f <- getConcreteFixity x
        let params = Set Name -> [LamBinding] -> DataDefParams
DataDefParams Set Name
gvars [LamBinding]
pars
        let dir' = Maybe (Ranged Induction)
-> Maybe (Ranged HasEta0)
-> Maybe Range
-> RecordConName
-> RecordDirectives' RecordConName
forall a.
Maybe (Ranged Induction)
-> Maybe (Ranged HasEta0)
-> Maybe Range
-> a
-> RecordDirectives' a
RecordDirectives Maybe (Ranged Induction)
ind Maybe (Ranged HasEta0)
eta Maybe Range
pat RecordConName
cm'
        return [ A.RecDef (mkDefInfoInstance x f PublicAccess a inst NotMacroDef r) x' uc dir' params contel afields ]

    NiceModule Range
r Access
p IsAbstract
a Erased
e x :: QName
x@(C.QName Name
name) Telescope
tel [Declaration]
ds -> ScopeM [Declaration] -> ScopeM [Declaration]
forall a. ScopeM a -> ScopeM a
notAffectedByOpaque (ScopeM [Declaration] -> ScopeM [Declaration])
-> ScopeM [Declaration] -> ScopeM [Declaration]
forall a b. (a -> b) -> a -> b
$ do
      [Char] -> Int -> TCMT IO Doc -> TCMT IO ()
forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> TCMT IO Doc -> m ()
reportSDoc [Char]
"scope.decl" Int
70 (TCMT IO Doc -> TCMT IO ()) -> TCMT IO Doc -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ [TCMT IO Doc] -> TCMT IO Doc
forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
vcat ([TCMT IO Doc] -> TCMT IO Doc) -> [TCMT IO Doc] -> TCMT IO Doc
forall a b. (a -> b) -> a -> b
$
        [ [Char] -> TCMT IO Doc
forall (m :: * -> *). Applicative m => [Char] -> m Doc
text ([Char] -> TCMT IO Doc) -> [Char] -> TCMT IO Doc
forall a b. (a -> b) -> a -> b
$ [Char]
"scope checking NiceModule " [Char] -> [Char] -> [Char]
forall a. [a] -> [a] -> [a]
++ QName -> [Char]
forall a. Pretty a => a -> [Char]
prettyShow QName
x
        ]

      adecl <- Call -> TCMT IO Declaration -> TCMT IO Declaration
forall a. Call -> TCMT IO a -> TCMT IO a
forall (m :: * -> *) a. MonadTrace m => Call -> m a -> m a
traceCall (NiceDeclaration -> Call
ScopeCheckDeclaration (NiceDeclaration -> Call) -> NiceDeclaration -> Call
forall a b. (a -> b) -> a -> b
$
                          Range
-> Access
-> IsAbstract
-> Erased
-> QName
-> Telescope
-> [Declaration]
-> NiceDeclaration
NiceModule Range
r Access
p IsAbstract
a Erased
e QName
x Telescope
tel []) (TCMT IO Declaration -> TCMT IO Declaration)
-> TCMT IO Declaration -> TCMT IO Declaration
forall a b. (a -> b) -> a -> b
$ do
        Range
-> Access
-> Erased
-> Name
-> Telescope
-> ScopeM [Declaration]
-> TCMT IO Declaration
scopeCheckNiceModule Range
r Access
p Erased
e Name
name Telescope
tel (ScopeM [Declaration] -> TCMT IO Declaration)
-> ScopeM [Declaration] -> TCMT IO Declaration
forall a b. (a -> b) -> a -> b
$
          Declarations -> ScopeM (AbsOfCon Declarations)
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract ([Declaration] -> Declarations
Declarations [Declaration]
ds)

      reportSDoc "scope.decl" 70 $ vcat $
        [ text $ "scope checked NiceModule " ++ prettyShow x
        , nest 2 $ prettyA adecl
        ]
      return [ adecl ]

    NiceModule Range
_ Access
_ IsAbstract
_ Erased
_ m :: QName
m@C.Qual{} Telescope
_ [Declaration]
_ -> TypeError -> ScopeM [Declaration]
forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
TypeError -> m a
typeError TypeError
QualifiedLocalModule

    NiceModuleMacro Range
r Access
p Erased
e Name
x ModuleApplication
modapp OpenShortHand
open ImportDirective
dir -> do
      [Char] -> Int -> TCMT IO Doc -> TCMT IO ()
forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> TCMT IO Doc -> m ()
reportSDoc [Char]
"scope.decl" Int
70 (TCMT IO Doc -> TCMT IO ()) -> TCMT IO Doc -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ [TCMT IO Doc] -> TCMT IO Doc
forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
vcat ([TCMT IO Doc] -> TCMT IO Doc) -> [TCMT IO Doc] -> TCMT IO Doc
forall a b. (a -> b) -> a -> b
$
        [ [Char] -> TCMT IO Doc
forall (m :: * -> *). Applicative m => [Char] -> m Doc
text ([Char] -> TCMT IO Doc) -> [Char] -> TCMT IO Doc
forall a b. (a -> b) -> a -> b
$ [Char]
"scope checking NiceModuleMacro " [Char] -> [Char] -> [Char]
forall a. [a] -> [a] -> [a]
++ Name -> [Char]
forall a. Pretty a => a -> [Char]
prettyShow Name
x
        ]

      adecl <- (ModuleInfo
 -> Erased
 -> ModuleName
 -> ModuleApplication
 -> ScopeCopyInfo
 -> ImportDirective
 -> Declaration)
-> OpenKind
-> Range
-> Access
-> Erased
-> Name
-> ModuleApplication
-> OpenShortHand
-> ImportDirective
-> TCMT IO Declaration
forall a.
(ToConcrete a, Pretty (ConOfAbs a)) =>
(ModuleInfo
 -> Erased
 -> ModuleName
 -> ModuleApplication
 -> ScopeCopyInfo
 -> ImportDirective
 -> a)
-> OpenKind
-> Range
-> Access
-> Erased
-> Name
-> ModuleApplication
-> OpenShortHand
-> ImportDirective
-> ScopeM a
checkModuleMacro ModuleInfo
-> Erased
-> ModuleName
-> ModuleApplication
-> ScopeCopyInfo
-> ImportDirective
-> Declaration
Apply OpenKind
TopOpenModule
                 Range
r Access
p Erased
e Name
x ModuleApplication
modapp OpenShortHand
open ImportDirective
dir

      reportSDoc "scope.decl" 70 $ vcat $
        [ text $ "scope checked NiceModuleMacro " ++ prettyShow x
        , nest 2 $ prettyA adecl
        ]
      return [ adecl ]

    NiceOpen Range
r QName
x ImportDirective
dir -> do
      (minfo, m, adir) <- Range
-> Maybe ModuleName
-> QName
-> ImportDirective
-> ScopeM (ModuleInfo, ModuleName, ImportDirective)
checkOpen Range
r Maybe ModuleName
forall a. Maybe a
Nothing QName
x ImportDirective
dir
      return [A.Open minfo m adir]

    NicePragma Range
r Pragma
p -> do
      ps <- Pragma -> ScopeM (AbsOfCon Pragma)
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract Pragma
p  -- could result in empty list of pragmas
      return $ map (A.Pragma r) ps

    NiceImport Range
r QName
x Maybe AsName
as OpenShortHand
open ImportDirective
dir -> Range -> ScopeM [Declaration] -> ScopeM [Declaration]
forall (m :: * -> *) x a.
(MonadTrace m, HasRange x) =>
x -> m a -> m a
setCurrentRange Range
r (ScopeM [Declaration] -> ScopeM [Declaration])
-> ScopeM [Declaration] -> ScopeM [Declaration]
forall a b. (a -> b) -> a -> b
$ do
      dir <- OpenShortHand -> ImportDirective -> ScopeM ImportDirective
notPublicWithoutOpen OpenShortHand
open ImportDirective
dir

      -- Andreas, 2018-11-03, issue #3364, parse expression in as-clause as Name.
      let illformedAs [Char]
s = Maybe AsName
-> TCMT IO (Maybe (AsName' Name)) -> TCMT IO (Maybe (AsName' Name))
forall (m :: * -> *) x a.
(MonadTrace m, HasRange x) =>
x -> m a -> m a
setCurrentRange Maybe AsName
as (TCMT IO (Maybe (AsName' Name)) -> TCMT IO (Maybe (AsName' Name)))
-> TCMT IO (Maybe (AsName' Name)) -> TCMT IO (Maybe (AsName' Name))
forall a b. (a -> b) -> a -> b
$ do
            -- If @as@ is followed by something that is not a simple name,
            -- throw a warning and discard the as-clause.
            Maybe (AsName' Name)
forall a. Maybe a
Nothing Maybe (AsName' Name)
-> TCMT IO () -> TCMT IO (Maybe (AsName' Name))
forall a b. a -> TCMT IO b -> TCMT IO a
forall (f :: * -> *) a b. Functor f => a -> f b -> f a
<$ Warning -> TCMT IO ()
forall (m :: * -> *).
(HasCallStack, MonadWarning m) =>
Warning -> m ()
warning ([Char] -> Warning
IllformedAsClause [Char]
s)
      as <- case as of
        -- Ok if no as-clause or it (already) contains a Name.
        Maybe AsName
Nothing -> Maybe (AsName' Name) -> TCMT IO (Maybe (AsName' Name))
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return Maybe (AsName' Name)
forall a. Maybe a
Nothing
        Just (AsName (Right Name
asName) Range
r)                    -> Maybe (AsName' Name) -> TCMT IO (Maybe (AsName' Name))
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return (Maybe (AsName' Name) -> TCMT IO (Maybe (AsName' Name)))
-> Maybe (AsName' Name) -> TCMT IO (Maybe (AsName' Name))
forall a b. (a -> b) -> a -> b
$ AsName' Name -> Maybe (AsName' Name)
forall a. a -> Maybe a
Just (AsName' Name -> Maybe (AsName' Name))
-> AsName' Name -> Maybe (AsName' Name)
forall a b. (a -> b) -> a -> b
$ Name -> Range -> AsName' Name
forall a. a -> Range -> AsName' a
AsName Name
asName Range
r
        Just (AsName (Left (C.Ident (C.QName Name
asName))) Range
r) -> Maybe (AsName' Name) -> TCMT IO (Maybe (AsName' Name))
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return (Maybe (AsName' Name) -> TCMT IO (Maybe (AsName' Name)))
-> Maybe (AsName' Name) -> TCMT IO (Maybe (AsName' Name))
forall a b. (a -> b) -> a -> b
$ AsName' Name -> Maybe (AsName' Name)
forall a. a -> Maybe a
Just (AsName' Name -> Maybe (AsName' Name))
-> AsName' Name -> Maybe (AsName' Name)
forall a b. (a -> b) -> a -> b
$ Name -> Range -> AsName' Name
forall a. a -> Range -> AsName' a
AsName Name
asName Range
r
        Just (AsName (Left C.Underscore{})     Range
r)         -> Maybe (AsName' Name) -> TCMT IO (Maybe (AsName' Name))
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return (Maybe (AsName' Name) -> TCMT IO (Maybe (AsName' Name)))
-> Maybe (AsName' Name) -> TCMT IO (Maybe (AsName' Name))
forall a b. (a -> b) -> a -> b
$ AsName' Name -> Maybe (AsName' Name)
forall a. a -> Maybe a
Just (AsName' Name -> Maybe (AsName' Name))
-> AsName' Name -> Maybe (AsName' Name)
forall a b. (a -> b) -> a -> b
$ Name -> Range -> AsName' Name
forall a. a -> Range -> AsName' a
AsName Name
forall a. Underscore a => a
underscore Range
r
        Just (AsName (Left (C.Ident C.Qual{})) Range
r) -> [Char] -> TCMT IO (Maybe (AsName' Name))
illformedAs [Char]
"; a qualified name is not allowed here"
        Just (AsName (Left Expr
e)                  Range
r) -> [Char] -> TCMT IO (Maybe (AsName' Name))
illformedAs [Char]
""

      top <- S.topLevelModuleName (rawTopLevelModuleNameForQName x)
      -- First scope check the imported module and return its name and
      -- interface. This is done with that module as the top-level module.
      -- This is quite subtle. We rely on the fact that when setting the
      -- top-level module and generating a fresh module name, the generated
      -- name will be exactly the same as the name generated when checking
      -- the imported module.
      (m, i) <- withCurrentModule noModuleName $
                withTopLevelModule top $ do
        m <- toAbstract $ NewModuleQName x  -- (No longer erases the contents of @m@.)
        printScope "import" 30 "before import:"
        (m, i) <- scopeCheckImport top m
        printScope "import" 30 $ "scope checked import: " ++ prettyShow i
        -- We don't want the top scope of the imported module (things happening
        -- before the module declaration)
        return (m, Map.delete noModuleName i)

      -- Bind the desired module name to the right abstract name.
      (name, theAsSymbol, theAsName) <- case as of

         Just AsName' Name
a | let y :: Name
y = AsName' Name -> Name
forall a. AsName' a -> a
asName AsName' Name
a, Bool -> Bool
not (Name -> Bool
forall a. IsNoName a => a -> Bool
isNoName Name
y) -> do
           Access -> Name -> ModuleName -> TCMT IO ()
bindModule Access
privateAccessInserted Name
y ModuleName
m
           (QName, Range, Maybe Name) -> TCMT IO (QName, Range, Maybe Name)
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return (Name -> QName
C.QName Name
y, AsName' Name -> Range
forall a. AsName' a -> Range
asRange AsName' Name
a, Name -> Maybe Name
forall a. a -> Maybe a
Just Name
y)

         Maybe (AsName' Name)
_ -> do
           -- Don't bind if @import ... as _@ with "no name"
           Maybe (AsName' Name) -> TCMT IO () -> TCMT IO ()
forall (m :: * -> *) a. Monad m => Maybe a -> m () -> m ()
whenNothing Maybe (AsName' Name)
as (TCMT IO () -> TCMT IO ()) -> TCMT IO () -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ Access -> QName -> ModuleName -> TCMT IO ()
bindQModule (Access
privateAccessInserted) QName
x ModuleName
m
           (QName, Range, Maybe Name) -> TCMT IO (QName, Range, Maybe Name)
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return (QName
x, Range
forall a. Range' a
noRange, Maybe Name
forall a. Maybe a
Nothing)

      -- Open if specified, otherwise apply import directives
      adir <- case open of

        -- With @open@ import directives apply to the opening.
        -- The module is thus present in its qualified form without restrictions.
        OpenShortHand
DoOpen   -> do

          -- Merge the imported scopes with the current scopes.
          -- This might override a previous import of @m@, but monotonously (add stuff).
          (Map ModuleName Scope -> Map ModuleName Scope) -> TCMT IO ()
modifyScopes ((Map ModuleName Scope -> Map ModuleName Scope) -> TCMT IO ())
-> (Map ModuleName Scope -> Map ModuleName Scope) -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ \ Map ModuleName Scope
ms -> (Scope -> Scope -> Scope)
-> Map ModuleName Scope
-> Map ModuleName Scope
-> Map ModuleName Scope
forall k a. Ord k => (a -> a -> a) -> Map k a -> Map k a -> Map k a
Map.unionWith Scope -> Scope -> Scope
mergeScope (ModuleName -> Map ModuleName Scope -> Map ModuleName Scope
forall k a. Ord k => k -> Map k a -> Map k a
Map.delete ModuleName
m Map ModuleName Scope
ms) Map ModuleName Scope
i

          -- Andreas, 2019-05-29, issue #3818.
          -- Pass the resolved name to open instead triggering another resolution.
          -- This helps in situations like
          -- @
          --    module Top where
          --    module M where
          --    open import M
          -- @
          -- It is clear than in @open import M@, name @M@ must refer to a file
          -- rather than the above defined local module @M@.
          -- This already worked in the situation
          -- @
          --    module Top where
          --    module M where
          --    import M
          -- @
          -- Note that the manual desugaring of @open import@ as
          -- @
          --    module Top where
          --    module M where
          --    import M
          --    open M
          -- @
          -- will not work, as @M@ is now ambiguous in @open M@;
          -- the information that @M@ is external is lost here.
          (_minfo, _m, adir) <- Range
-> Maybe ModuleName
-> QName
-> ImportDirective
-> ScopeM (ModuleInfo, ModuleName, ImportDirective)
checkOpen Range
r (ModuleName -> Maybe ModuleName
forall a. a -> Maybe a
Just ModuleName
m) QName
name ImportDirective
dir
          return adir

        -- If not opening, import directives are applied to the original scope.
        OpenShortHand
DontOpen -> do
          (adir, i') <- (Scope -> ScopeM (ImportDirective, Scope))
-> ModuleName
-> Map ModuleName Scope
-> TCMT IO (ImportDirective, Map ModuleName Scope)
forall (f :: * -> *) k v a.
(Functor f, Ord k) =>
(v -> f (a, v)) -> k -> Map k v -> f (a, Map k v)
Map.adjustM' (QName
-> ImportDirective -> Scope -> ScopeM (ImportDirective, Scope)
applyImportDirectiveM QName
x ImportDirective
dir) ModuleName
m Map ModuleName Scope
i
          -- Andreas, 2020-05-18, issue #3933
          -- We merge the new imports without deleting old imports, to be monotone.
          modifyScopes $ \ Map ModuleName Scope
ms -> (Scope -> Scope -> Scope)
-> Map ModuleName Scope
-> Map ModuleName Scope
-> Map ModuleName Scope
forall k a. Ord k => (a -> a -> a) -> Map k a -> Map k a -> Map k a
Map.unionWith Scope -> Scope -> Scope
mergeScope Map ModuleName Scope
ms Map ModuleName Scope
i'
          return adir

      printScope "import" 30 "merged imported sig:"
      let minfo = ModuleInfo
            { minfoRange :: Range
minfoRange     = Range
r
            , minfoAsName :: Maybe Name
minfoAsName    = Maybe Name
theAsName
            , minfoAsTo :: Range
minfoAsTo      = (Range, Range) -> Range
forall a. HasRange a => a -> Range
getRange (Range
theAsSymbol, ImportDirective -> Range
renamingRange ImportDirective
dir)
            , minfoOpenShort :: Maybe OpenShortHand
minfoOpenShort = OpenShortHand -> Maybe OpenShortHand
forall a. a -> Maybe a
Just OpenShortHand
open
            , minfoDirective :: Maybe ImportDirective
minfoDirective = ImportDirective -> Maybe ImportDirective
forall a. a -> Maybe a
Just ImportDirective
dir
            }
      return [ A.Import minfo m adir ]

    NiceUnquoteDecl Range
r Access
p IsAbstract
a IsInstance
i TerminationCheck
tc CoverageCheck
cc [Name]
xs Expr
e -> do
      fxs <- (Name -> ScopeM Fixity') -> [Name] -> TCMT IO [Fixity']
forall (t :: * -> *) (m :: * -> *) a b.
(Traversable t, Monad m) =>
(a -> m b) -> t a -> m (t b)
forall (m :: * -> *) a b. Monad m => (a -> m b) -> [a] -> m [b]
mapM Name -> ScopeM Fixity'
getConcreteFixity [Name]
xs
      ys <- zipWithM freshAbstractQName fxs xs
      zipWithM_ (bindName p QuotableName) xs ys
      e <- toAbstract e
      zipWithM_ (rebindName p OtherDefName) xs ys
      let mi = TerminationCheck
-> CoverageCheck -> PositivityCheck -> Range -> MutualInfo
MutualInfo TerminationCheck
tc CoverageCheck
cc PositivityCheck
YesPositivityCheck Range
r
      mapM_ unfoldFunction ys
      opaque <- contextIsOpaque
      return [ A.Mutual mi
        [ A.UnquoteDecl mi
            [ (mkDefInfoInstance x fx p a i NotMacroDef r) { Info.defOpaque = opaque } | (fx, x) <- zip fxs xs ]
          ys e
        ] ]

    NiceUnquoteDef Range
r Access
p IsAbstract
a TerminationCheck
_ CoverageCheck
_ [Name]
xs Expr
e -> do
      fxs <- (Name -> ScopeM Fixity') -> [Name] -> TCMT IO [Fixity']
forall (t :: * -> *) (m :: * -> *) a b.
(Traversable t, Monad m) =>
(a -> m b) -> t a -> m (t b)
forall (m :: * -> *) a b. Monad m => (a -> m b) -> [a] -> m [b]
mapM Name -> ScopeM Fixity'
getConcreteFixity [Name]
xs
      ys <- mapM (toAbstract . OldName) xs
      zipWithM_ (rebindName p QuotableName) xs ys
      e <- toAbstract e
      zipWithM_ (rebindName p OtherDefName) xs ys
      mapM_ unfoldFunction ys
      opaque <- contextIsOpaque
      return [ A.UnquoteDef [ (mkDefInfo x fx PublicAccess a r) { Info.defOpaque = opaque } | (fx, x) <- zip fxs xs ] ys e ]

    NiceUnquoteData Range
r Access
p IsAbstract
a PositivityCheck
pc UniverseCheck
uc Name
x [Name]
cs Expr
e -> ScopeM [Declaration] -> ScopeM [Declaration]
forall a. ScopeM a -> ScopeM a
notAffectedByOpaque (ScopeM [Declaration] -> ScopeM [Declaration])
-> ScopeM [Declaration] -> ScopeM [Declaration]
forall a b. (a -> b) -> a -> b
$ do
      fx <- Name -> ScopeM Fixity'
getConcreteFixity Name
x
      x' <- freshAbstractQName fx x
      bindName p QuotableName x x'

      -- Create the module for the qualified constructors
      checkForModuleClash x
      let m = QName -> ModuleName
qnameToMName QName
x'
      createModule (Just IsDataModule) m
      bindModule p x m  -- make it a proper module

      cs' <- mapM (bindUnquoteConstructorName m p) cs

      e <- withCurrentModule m $ toAbstract e

      rebindName p DataName x x'
      zipWithM_ (rebindName p ConName) cs cs'
      withCurrentModule m $ zipWithM_ (rebindName p ConName) cs cs'

      fcs <- mapM getConcreteFixity cs
      let mi = TerminationCheck
-> CoverageCheck -> PositivityCheck -> Range -> MutualInfo
MutualInfo TerminationCheck
forall m. TerminationCheck m
TerminationCheck CoverageCheck
YesCoverageCheck PositivityCheck
pc Range
r
      return
        [ A.Mutual
          mi [A.UnquoteData
            [ mkDefInfo x fx p a r ] x' uc
            [ mkDefInfo c fc p a r | (fc, c) <- zip fcs cs] cs' e ]
        ]

    NicePatternSyn Range
r Access
a Name
n [WithHiding Name]
as Pattern
p -> do
      [Char] -> Int -> [Char] -> TCMT IO ()
forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> [Char] -> m ()
reportSLn [Char]
"scope.pat" Int
30 ([Char] -> TCMT IO ()) -> [Char] -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ [Char]
"found nice pattern syn: " [Char] -> [Char] -> [Char]
forall a. [a] -> [a] -> [a]
++ Name -> [Char]
forall a. Pretty a => a -> [Char]
prettyShow Name
n
      (as, p) <- ScopeM ([WithHiding Name], Pattern' Void)
-> ScopeM ([WithHiding Name], Pattern' Void)
forall a. ScopeM a -> ScopeM a
withLocalVars (ScopeM ([WithHiding Name], Pattern' Void)
 -> ScopeM ([WithHiding Name], Pattern' Void))
-> ScopeM ([WithHiding Name], Pattern' Void)
-> ScopeM ([WithHiding Name], Pattern' Void)
forall a b. (a -> b) -> a -> b
$ do
         -- Expand puns if optHiddenArgumentPuns is True.
         p <- Pattern -> TCMT IO Pattern
parsePatternSyn (Pattern -> TCMT IO Pattern) -> TCMT IO Pattern -> TCMT IO Pattern
forall (m :: * -> *) a b. Monad m => (a -> m b) -> m a -> m b
=<< Pattern -> TCMT IO Pattern
expandPunsOpt Pattern
p
         p <- toAbstract p
         when (containsAsPattern p) $
           typeError AsPatternInPatternSynonym
         checkPatternLinearity p $ \List1 Name
ys ->
           TypeError -> TCMT IO ()
forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
TypeError -> m a
typeError (TypeError -> TCMT IO ()) -> TypeError -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ List1 Name -> TypeError
RepeatedVariablesInPattern List1 Name
ys
         -- Bind the pattern variables accumulated by @ToAbstract Pattern@ applied to the rhs.
         bindVarsToBind
         p <- A.noDotOrEqPattern (typeError DotPatternInPatternSynonym) p
         as <- mapM checkPatSynParam as
         List1.unlessNull (patternVars p List.\\ map whThing as) $ \ List1 Name
xs -> do
           TypeError -> TCMT IO ()
forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
TypeError -> m a
typeError (TypeError -> TCMT IO ()) -> TypeError -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ List1 Name -> TypeError
UnboundVariablesInPatternSynonym List1 Name
xs
         return (as, p)
      y <- freshAbstractQName' n
      bindName a PatternSynName n y
      -- Expanding pattern synonyms already at definition makes it easier to
      -- fold them back when printing (issue #2762).
      ep <- expandPatternSynonyms p
      modifyPatternSyns (Map.insert y (as, ep))
      return [A.PatternSynDef y (map (fmap BindName) as) p]   -- only for highlighting, so use unexpanded version
      where
        checkPatSynParam :: WithHiding C.Name -> ScopeM (WithHiding A.Name)
        checkPatSynParam :: WithHiding Name -> TCMT IO (WithHiding Name)
checkPatSynParam (WithHiding Hiding
h Name
x) = do
          let err :: TypeError -> TCMT IO (WithHiding Name)
err = Name -> TCMT IO (WithHiding Name) -> TCMT IO (WithHiding Name)
forall (m :: * -> *) x a.
(MonadTrace m, HasRange x) =>
x -> m a -> m a
setCurrentRange Name
x (TCMT IO (WithHiding Name) -> TCMT IO (WithHiding Name))
-> (TypeError -> TCMT IO (WithHiding Name))
-> TypeError
-> TCMT IO (WithHiding Name)
forall b c a. (b -> c) -> (a -> b) -> a -> c
. TypeError -> TCMT IO (WithHiding Name)
forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
TypeError -> m a
typeError
          QName -> ScopeM ResolvedName
resolveName (Name -> QName
C.QName Name
x) ScopeM ResolvedName
-> (ResolvedName -> TCMT IO (WithHiding Name))
-> TCMT IO (WithHiding Name)
forall a b. TCMT IO a -> (a -> TCMT IO b) -> TCMT IO b
forall (m :: * -> *) a b. Monad m => m a -> (a -> m b) -> m b
>>= \case
            VarName Name
a (PatternBound Hiding
h')
              | Hiding -> Bool
forall a. LensHiding a => a -> Bool
isInstance Hiding
h, Bool -> Bool
not (Hiding -> Bool
forall a. LensHiding a => a -> Bool
isInstance Hiding
h') -> TypeError -> TCMT IO (WithHiding Name)
err (TypeError -> TCMT IO (WithHiding Name))
-> TypeError -> TCMT IO (WithHiding Name)
forall a b. (a -> b) -> a -> b
$ Name -> TypeError
IllegalInstanceVariableInPatternSynonym Name
x
              | Bool
otherwise -> WithHiding Name -> TCMT IO (WithHiding Name)
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return (WithHiding Name -> TCMT IO (WithHiding Name))
-> WithHiding Name -> TCMT IO (WithHiding Name)
forall a b. (a -> b) -> a -> b
$ Hiding -> Name -> WithHiding Name
forall a. Hiding -> a -> WithHiding a
WithHiding Hiding
h Name
a
            ConstructorName Set1 Induction
_ List1 AbstractName
ys -> TypeError -> TCMT IO (WithHiding Name)
err (TypeError -> TCMT IO (WithHiding Name))
-> TypeError -> TCMT IO (WithHiding Name)
forall a b. (a -> b) -> a -> b
$ ConstructorOrPatternSynonym
-> Name -> List1 AbstractName -> TypeError
PatternSynonymArgumentShadows ConstructorOrPatternSynonym
IsConstructor Name
x List1 AbstractName
ys
            PatternSynResName List1 AbstractName
ys -> TypeError -> TCMT IO (WithHiding Name)
err (TypeError -> TCMT IO (WithHiding Name))
-> TypeError -> TCMT IO (WithHiding Name)
forall a b. (a -> b) -> a -> b
$ ConstructorOrPatternSynonym
-> Name -> List1 AbstractName -> TypeError
PatternSynonymArgumentShadows ConstructorOrPatternSynonym
IsPatternSynonym Name
x List1 AbstractName
ys
            ResolvedName
UnknownName -> TypeError -> TCMT IO (WithHiding Name)
err (TypeError -> TCMT IO (WithHiding Name))
-> TypeError -> TCMT IO (WithHiding Name)
forall a b. (a -> b) -> a -> b
$ Name -> TypeError
UnusedVariableInPatternSynonym Name
x
            -- Other cases are impossible because parsing the pattern syn rhs would have failed.
            ResolvedName
_ -> TCMT IO (WithHiding Name)
forall a. HasCallStack => a
__IMPOSSIBLE__

    d :: NiceDeclaration
d@NiceLoneConstructor{} -> (CallStack -> ScopeM [Declaration]) -> ScopeM [Declaration]
forall b. HasCallStack => (CallStack -> b) -> b
withCurrentCallStack ((CallStack -> ScopeM [Declaration]) -> ScopeM [Declaration])
-> (CallStack -> ScopeM [Declaration]) -> ScopeM [Declaration]
forall a b. (a -> b) -> a -> b
$ \ CallStack
stk -> do
      Warning -> TCMT IO ()
forall (m :: * -> *).
(HasCallStack, MonadWarning m) =>
Warning -> m ()
warning (Warning -> TCMT IO ()) -> Warning -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ DeclarationWarning -> Warning
NicifierIssue (CallStack -> DeclarationWarning' -> DeclarationWarning
DeclarationWarning CallStack
stk (Range -> DeclarationWarning'
InvalidConstructorBlock (NiceDeclaration -> Range
forall a. HasRange a => a -> Range
getRange NiceDeclaration
d)))
      [Declaration] -> ScopeM [Declaration]
forall a. a -> TCMT IO a
forall (f :: * -> *) a. Applicative f => a -> f a
pure []

    d :: NiceDeclaration
d@(NiceOpaque KwRange
kwr [QName]
xs [NiceDeclaration]
decls) -> do
      -- The names in an 'unfolding' clause must be unambiguous names of definitions:
      -- Resolve all the names, and use them as an initial unfolding set:
      names  <- [Maybe QName] -> [QName]
forall a. [Maybe a] -> [a]
catMaybes ([Maybe QName] -> [QName])
-> TCMT IO [Maybe QName] -> TCMT IO [QName]
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> [QName]
-> (QName -> TCMT IO (Maybe QName)) -> TCMT IO [Maybe QName]
forall (t :: * -> *) (m :: * -> *) a b.
(Traversable t, Monad m) =>
t a -> (a -> m b) -> m (t b)
forM [QName]
xs \ QName
x -> do
        QName -> TCMT IO (Maybe QName) -> TCMT IO (Maybe QName)
forall (m :: * -> *) x a.
(MonadTrace m, HasRange x) =>
x -> m a -> m a
setCurrentRange QName
x (TCMT IO (Maybe QName) -> TCMT IO (Maybe QName))
-> TCMT IO (Maybe QName) -> TCMT IO (Maybe QName)
forall a b. (a -> b) -> a -> b
$ (QName -> IsAmbiguous -> Warning) -> QName -> TCMT IO (Maybe QName)
unambiguousConOrDef (Warning -> IsAmbiguous -> Warning
forall a b. a -> b -> a
const (Warning -> IsAmbiguous -> Warning)
-> (QName -> Warning) -> QName -> IsAmbiguous -> Warning
forall b c a. (b -> c) -> (a -> b) -> a -> c
. QName -> Warning
UnfoldingWrongName) QName
x
      -- Generate the identifier for this block:
      oid    <- fresh
      -- Record the parent unfolding block, if any:
      parent <- asksTC envCurrentOpaqueId

      let r = NiceDeclaration -> Range
forall a. HasRange a => a -> Range
getRange NiceDeclaration
d
      stOpaqueBlocks `modifyTCLens` Map.insert oid OpaqueBlock
        { opaqueId        = oid
        , opaqueUnfolding = HashSet.fromList names
        , opaqueDecls     = mempty
        , opaqueParent    = parent
        , opaqueRange     = r
        }

      -- Keep going!
      localTC (\TCEnv
e -> TCEnv
e { envCurrentOpaqueId = Just oid }) $ do
        out <- traverse toAbstract decls
        unless (any interestingOpaqueDecl out) $ setCurrentRange kwr $ warning UselessOpaque
        pure $ UnfoldingDecl r names : out
    where
      -- checking postulate or type sig. without checking safe flag
      toAbstractNiceAxiom :: KindOfName -> C.NiceDeclaration -> ScopeM A.Declaration
      toAbstractNiceAxiom :: KindOfName -> NiceDeclaration -> TCMT IO Declaration
toAbstractNiceAxiom KindOfName
kind (C.Axiom Range
r Access
p IsAbstract
a IsInstance
i ArgInfo
info Name
x Expr
t) = do
        t' <- Precedence -> Expr -> ScopeM (AbsOfCon Expr)
forall c. ToAbstract c => Precedence -> c -> ScopeM (AbsOfCon c)
toAbstractCtx Precedence
TopCtx Expr
t
        f  <- getConcreteFixity x
        mp <- getConcretePolarity x
        y  <- freshAbstractQName f x
        let isMacro | KindOfName
kind KindOfName -> KindOfName -> Bool
forall a. Eq a => a -> a -> Bool
== KindOfName
MacroName = IsMacro
MacroDef
                    | Bool
otherwise         = IsMacro
NotMacroDef
        bindName p kind x y
        definfo <- updateDefInfoOpacity $ mkDefInfoInstance x f p a i isMacro r
        return $ A.Axiom kind definfo info mp y t'
      toAbstractNiceAxiom KindOfName
_ NiceDeclaration
_ = TCMT IO Declaration
forall a. HasCallStack => a
__IMPOSSIBLE__

      interestingOpaqueDecl :: A.Declaration -> Bool
      interestingOpaqueDecl :: Declaration -> Bool
interestingOpaqueDecl (A.Mutual MutualInfo
_ [Declaration]
ds)     = (Declaration -> Bool) -> [Declaration] -> Bool
forall (t :: * -> *) a. Foldable t => (a -> Bool) -> t a -> Bool
any Declaration -> Bool
interestingOpaqueDecl [Declaration]
ds
      interestingOpaqueDecl (A.ScopedDecl ScopeInfo
_ [Declaration]
ds) = (Declaration -> Bool) -> [Declaration] -> Bool
forall (t :: * -> *) a. Foldable t => (a -> Bool) -> t a -> Bool
any Declaration -> Bool
interestingOpaqueDecl [Declaration]
ds

      interestingOpaqueDecl A.FunDef{}      = Bool
True
      interestingOpaqueDecl A.UnquoteDecl{} = Bool
True
      interestingOpaqueDecl A.UnquoteDef{}  = Bool
True

      interestingOpaqueDecl Declaration
_ = Bool
False

-- ** Helper functions for @opaque@
------------------------------------------------------------------------

-- | Add a 'QName' to the set of declarations /contained in/ the current
-- opaque block.
unfoldFunction :: A.QName -> ScopeM ()
unfoldFunction :: QName -> TCMT IO ()
unfoldFunction QName
qn = (TCEnv -> Maybe OpaqueId) -> TCMT IO (Maybe OpaqueId)
forall (m :: * -> *) a. MonadTCEnv m => (TCEnv -> a) -> m a
asksTC TCEnv -> Maybe OpaqueId
envCurrentOpaqueId TCMT IO (Maybe OpaqueId)
-> (Maybe OpaqueId -> TCMT IO ()) -> TCMT IO ()
forall a b. TCMT IO a -> (a -> TCMT IO b) -> TCMT IO b
forall (m :: * -> *) a b. Monad m => m a -> (a -> m b) -> m b
>>= \case
  Just OpaqueId
id -> do
    let go :: Maybe OpaqueBlock -> Maybe OpaqueBlock
go Maybe OpaqueBlock
Nothing   = Maybe OpaqueBlock
forall a. HasCallStack => a
__IMPOSSIBLE__
        go (Just OpaqueBlock
ob) = OpaqueBlock -> Maybe OpaqueBlock
forall a. a -> Maybe a
Just OpaqueBlock
ob{ opaqueDecls = qn `HashSet.insert` opaqueDecls ob }
    (Map OpaqueId OpaqueBlock -> f (Map OpaqueId OpaqueBlock))
-> TCState -> f TCState
Lens' TCState (Map OpaqueId OpaqueBlock)
stOpaqueBlocks Lens' TCState (Map OpaqueId OpaqueBlock)
-> (Map OpaqueId OpaqueBlock -> Map OpaqueId OpaqueBlock)
-> TCMT IO ()
forall (m :: * -> *) a.
MonadTCState m =>
Lens' TCState a -> (a -> a) -> m ()
`modifyTCLens` (Maybe OpaqueBlock -> Maybe OpaqueBlock)
-> OpaqueId -> Map OpaqueId OpaqueBlock -> Map OpaqueId OpaqueBlock
forall k a.
Ord k =>
(Maybe a -> Maybe a) -> k -> Map k a -> Map k a
Map.alter Maybe OpaqueBlock -> Maybe OpaqueBlock
go OpaqueId
id
  Maybe OpaqueId
Nothing -> () -> TCMT IO ()
forall a. a -> TCMT IO a
forall (f :: * -> *) a. Applicative f => a -> f a
pure ()

-- | Look up the current opaque identifier as a value in 'IsOpaque'.
contextIsOpaque :: ScopeM IsOpaque
contextIsOpaque :: ScopeM IsOpaque
contextIsOpaque =  IsOpaque -> (OpaqueId -> IsOpaque) -> Maybe OpaqueId -> IsOpaque
forall b a. b -> (a -> b) -> Maybe a -> b
maybe IsOpaque
TransparentDef OpaqueId -> IsOpaque
OpaqueDef (Maybe OpaqueId -> IsOpaque)
-> TCMT IO (Maybe OpaqueId) -> ScopeM IsOpaque
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> (TCEnv -> Maybe OpaqueId) -> TCMT IO (Maybe OpaqueId)
forall (m :: * -> *) a. MonadTCEnv m => (TCEnv -> a) -> m a
asksTC TCEnv -> Maybe OpaqueId
envCurrentOpaqueId

updateDefInfoOpacity :: DefInfo -> ScopeM DefInfo
updateDefInfoOpacity :: DefInfo -> ScopeM DefInfo
updateDefInfoOpacity DefInfo
di = (\IsOpaque
a -> DefInfo
di { Info.defOpaque = a }) (IsOpaque -> DefInfo) -> ScopeM IsOpaque -> ScopeM DefInfo
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> ScopeM IsOpaque
contextIsOpaque

-- | Raise a warning indicating that the current Declaration is not
-- affected by opacity, but only if we are actually in an Opaque block.
notAffectedByOpaque :: ScopeM a -> ScopeM a
notAffectedByOpaque :: forall a. ScopeM a -> ScopeM a
notAffectedByOpaque ScopeM a
k = do
  t <- (TCEnv -> Bool) -> TCMT IO Bool
forall (m :: * -> *) a. MonadTCEnv m => (TCEnv -> a) -> m a
asksTC TCEnv -> Bool
envCheckingWhere
  unless t $
    maybe (pure ()) (const (warning NotAffectedByOpaque)) =<< asksTC envCurrentOpaqueId
  notUnderOpaque k

-- * Helper functions for @variable@ generalization
------------------------------------------------------------------------

unGeneralized :: A.Expr -> (Set A.QName, A.Expr)
unGeneralized :: Expr -> (Set QName, Expr)
unGeneralized (A.Generalized Set1 QName
s Expr
t) = (Set1 QName -> Set QName
forall a. NESet a -> Set a
Set1.toSet Set1 QName
s, Expr
t)
unGeneralized (A.ScopedExpr ScopeInfo
si Expr
e) = ScopeInfo -> Expr -> Expr
A.ScopedExpr ScopeInfo
si (Expr -> Expr) -> (Set QName, Expr) -> (Set QName, Expr)
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> Expr -> (Set QName, Expr)
unGeneralized Expr
e
unGeneralized Expr
t = (Set QName
forall a. Monoid a => a
mempty, Expr
t)

alreadyGeneralizing :: ScopeM Bool
alreadyGeneralizing :: TCMT IO Bool
alreadyGeneralizing = Maybe (Set QName) -> Bool
forall a. Maybe a -> Bool
isJust (Maybe (Set QName) -> Bool)
-> TCMT IO (Maybe (Set QName)) -> TCMT IO Bool
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> Lens' TCState (Maybe (Set QName)) -> TCMT IO (Maybe (Set QName))
forall (m :: * -> *) a. ReadTCState m => Lens' TCState a -> m a
useTC (Maybe (Set QName) -> f (Maybe (Set QName)))
-> TCState -> f TCState
Lens' TCState (Maybe (Set QName))
stGeneralizedVars

collectGeneralizables :: ScopeM a -> ScopeM (Set A.QName, a)
collectGeneralizables :: forall a. ScopeM a -> ScopeM (Set QName, a)
collectGeneralizables ScopeM a
m =
  -- #5683: No nested generalization
  TCMT IO Bool
-> TCMT IO (Set QName, a)
-> TCMT IO (Set QName, a)
-> TCMT IO (Set QName, a)
forall (m :: * -> *) a. Monad m => m Bool -> m a -> m a -> m a
ifM TCMT IO Bool
alreadyGeneralizing ((Set QName
forall a. Set a
Set.empty,) (a -> (Set QName, a)) -> ScopeM a -> TCMT IO (Set QName, a)
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> ScopeM a
m) (TCMT IO (Set QName, a) -> TCMT IO (Set QName, a))
-> TCMT IO (Set QName, a) -> TCMT IO (Set QName, a)
forall a b. (a -> b) -> a -> b
$
  {-else-} TCMT IO (Maybe (Set QName))
-> (Maybe (Set QName) -> TCMT IO ())
-> TCMT IO (Set QName, a)
-> TCMT IO (Set QName, a)
forall (m :: * -> *) a b.
Monad m =>
m a -> (a -> m ()) -> m b -> m b
bracket_ TCMT IO (Maybe (Set QName))
open Maybe (Set QName) -> TCMT IO ()
close (TCMT IO (Set QName, a) -> TCMT IO (Set QName, a))
-> TCMT IO (Set QName, a) -> TCMT IO (Set QName, a)
forall a b. (a -> b) -> a -> b
$ do
      a <- ScopeM a
m
      s <- useTC stGeneralizedVars
      case s of
          Maybe (Set QName)
Nothing -> TCMT IO (Set QName, a)
forall a. HasCallStack => a
__IMPOSSIBLE__
          Just Set QName
s -> (Set QName, a) -> TCMT IO (Set QName, a)
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return (Set QName
s, a
a)
  where
    open :: TCMT IO (Maybe (Set QName))
open = do
        gvs <- Lens' TCState (Maybe (Set QName)) -> TCMT IO (Maybe (Set QName))
forall (m :: * -> *) a. ReadTCState m => Lens' TCState a -> m a
useTC (Maybe (Set QName) -> f (Maybe (Set QName)))
-> TCState -> f TCState
Lens' TCState (Maybe (Set QName))
stGeneralizedVars
        stGeneralizedVars `setTCLens` Just mempty
        pure gvs
    close :: Maybe (Set QName) -> TCMT IO ()
close = ((Maybe (Set QName) -> f (Maybe (Set QName)))
-> TCState -> f TCState
Lens' TCState (Maybe (Set QName))
stGeneralizedVars Lens' TCState (Maybe (Set QName))
-> Maybe (Set QName) -> TCMT IO ()
forall (m :: * -> *) a.
MonadTCState m =>
Lens' TCState a -> a -> m ()
`setTCLens`)

createBoundNamesForGeneralizables :: Set A.QName -> ScopeM (Map A.QName A.Name)
createBoundNamesForGeneralizables :: Set QName -> ScopeM (Map QName Name)
createBoundNamesForGeneralizables Set QName
vs =
  ((QName -> () -> ScopeM Name)
 -> Map QName () -> ScopeM (Map QName Name))
-> Map QName ()
-> (QName -> () -> ScopeM Name)
-> ScopeM (Map QName Name)
forall a b c. (a -> b -> c) -> b -> a -> c
flip (QName -> () -> ScopeM Name)
-> Map QName () -> ScopeM (Map QName Name)
forall (t :: * -> *) k a b.
Applicative t =>
(k -> a -> t b) -> Map k a -> t (Map k b)
Map.traverseWithKey ((QName -> ()) -> Set QName -> Map QName ()
forall k a. (k -> a) -> Set k -> Map k a
Map.fromSet (() -> QName -> ()
forall a b. a -> b -> a
const ()) Set QName
vs) ((QName -> () -> ScopeM Name) -> ScopeM (Map QName Name))
-> (QName -> () -> ScopeM Name) -> ScopeM (Map QName Name)
forall a b. (a -> b) -> a -> b
$ \ QName
q ()
_ -> do
    let x :: Name
x  = Name -> Name
nameConcrete (Name -> Name) -> Name -> Name
forall a b. (a -> b) -> a -> b
$ QName -> Name
qnameName QName
q
        fx :: Fixity'
fx = Name -> Fixity'
nameFixity   (Name -> Fixity') -> Name -> Fixity'
forall a b. (a -> b) -> a -> b
$ QName -> Name
qnameName QName
q
    Fixity' -> Name -> ScopeM Name
freshAbstractName Fixity'
fx Name
x

collectAndBindGeneralizables :: ScopeM a -> ScopeM (Map A.QName A.Name, a)
collectAndBindGeneralizables :: forall a. ScopeM a -> ScopeM (Map QName Name, a)
collectAndBindGeneralizables ScopeM a
m = do
  fvBefore <- [(Name, LocalVar)] -> Int
forall a. [a] -> Int
forall (t :: * -> *) a. Foldable t => t a -> Int
length ([(Name, LocalVar)] -> Int)
-> TCMT IO [(Name, LocalVar)] -> TCMT IO Int
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> TCMT IO [(Name, LocalVar)]
forall (m :: * -> *). ReadTCState m => m [(Name, LocalVar)]
getLocalVars
  (s, res) <- collectGeneralizables m
  fvAfter  <- length <$> getLocalVars
  -- We should bind the named generalizable variables as fresh variables
  binds <- createBoundNamesForGeneralizables s
  -- Issue #3735: We need to bind the generalizable variables outside any variables bound by `m`.
  outsideLocalVars (fvAfter - fvBefore) $ bindGeneralizables binds
  return (binds, res)

bindGeneralizables :: Map A.QName A.Name -> ScopeM ()
bindGeneralizables :: Map QName Name -> TCMT IO ()
bindGeneralizables Map QName Name
vars =
  [(QName, Name)] -> ((QName, Name) -> TCMT IO ()) -> TCMT IO ()
forall (t :: * -> *) (m :: * -> *) a b.
(Foldable t, Monad m) =>
t a -> (a -> m b) -> m ()
forM_ (Map QName Name -> [(QName, Name)]
forall k a. Map k a -> [(k, a)]
Map.toList Map QName Name
vars) (((QName, Name) -> TCMT IO ()) -> TCMT IO ())
-> ((QName, Name) -> TCMT IO ()) -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ \ (QName
q, Name
y) ->
    BindingSource -> Name -> Name -> TCMT IO ()
bindVariable BindingSource
LambdaBound (Name -> Name
nameConcrete (Name -> Name) -> Name -> Name
forall a b. (a -> b) -> a -> b
$ QName -> Name
qnameName QName
q) Name
y

-- | Bind generalizable variables if data or record decl was split by the system
--   (origin == Inserted)
bindGeneralizablesIfInserted :: Origin -> AbstractName -> ScopeM (Set A.Name)
bindGeneralizablesIfInserted :: Origin -> AbstractName -> ScopeM (Set Name)
bindGeneralizablesIfInserted Origin
Inserted AbstractName
y = Set Name
bound Set Name -> TCMT IO () -> ScopeM (Set Name)
forall a b. a -> TCMT IO b -> TCMT IO a
forall (f :: * -> *) a b. Functor f => a -> f b -> f a
<$ Map QName Name -> TCMT IO ()
bindGeneralizables Map QName Name
gvars
  where gvars :: Map QName Name
gvars = case AbstractName -> NameMetadata
anameMetadata AbstractName
y of
          GeneralizedVarsMetadata Map QName Name
gvars -> Map QName Name
gvars
          NameMetadata
NoMetadata                    -> Map QName Name
forall k a. Map k a
Map.empty
        bound :: Set Name
bound = [Name] -> Set Name
forall a. Ord a => [a] -> Set a
Set.fromList (Map QName Name -> [Name]
forall k a. Map k a -> [a]
Map.elems Map QName Name
gvars)
bindGeneralizablesIfInserted Origin
UserWritten AbstractName
_ = Set Name -> ScopeM (Set Name)
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return Set Name
forall a. Set a
Set.empty
bindGeneralizablesIfInserted Origin
_ AbstractName
_           = ScopeM (Set Name)
forall a. HasCallStack => a
__IMPOSSIBLE__

newtype GenTel = GenTel C.Telescope
data GenTelAndType = GenTelAndType C.Telescope C.Expr

instance ToAbstract GenTel where
  type AbsOfCon GenTel = A.GeneralizeTelescope
  toAbstract :: GenTel -> ScopeM (AbsOfCon GenTel)
toAbstract (GenTel Telescope
tel) =
    (Map QName Name -> [TypedBinding] -> GeneralizeTelescope)
-> (Map QName Name, [TypedBinding]) -> GeneralizeTelescope
forall a b c. (a -> b -> c) -> (a, b) -> c
uncurry Map QName Name -> [TypedBinding] -> GeneralizeTelescope
A.GeneralizeTel ((Map QName Name, [TypedBinding]) -> GeneralizeTelescope)
-> TCMT IO (Map QName Name, [TypedBinding])
-> TCMT IO GeneralizeTelescope
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> TCMT IO [TypedBinding] -> TCMT IO (Map QName Name, [TypedBinding])
forall a. ScopeM a -> ScopeM (Map QName Name, a)
collectAndBindGeneralizables ([Maybe TypedBinding] -> [TypedBinding]
forall a. [Maybe a] -> [a]
catMaybes ([Maybe TypedBinding] -> [TypedBinding])
-> TCMT IO [Maybe TypedBinding] -> TCMT IO [TypedBinding]
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> Telescope -> ScopeM (AbsOfCon Telescope)
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract Telescope
tel)

instance ToAbstract GenTelAndType where
  type AbsOfCon GenTelAndType = (A.GeneralizeTelescope, A.Expr)

  toAbstract :: GenTelAndType -> ScopeM (AbsOfCon GenTelAndType)
toAbstract (GenTelAndType Telescope
tel Expr
t) = do
    (binds, (tel, t)) <- ScopeM ([Maybe TypedBinding], Expr)
-> ScopeM (Map QName Name, ([Maybe TypedBinding], Expr))
forall a. ScopeM a -> ScopeM (Map QName Name, a)
collectAndBindGeneralizables (ScopeM ([Maybe TypedBinding], Expr)
 -> ScopeM (Map QName Name, ([Maybe TypedBinding], Expr)))
-> ScopeM ([Maybe TypedBinding], Expr)
-> ScopeM (Map QName Name, ([Maybe TypedBinding], Expr))
forall a b. (a -> b) -> a -> b
$
                          (,) ([Maybe TypedBinding] -> Expr -> ([Maybe TypedBinding], Expr))
-> TCMT IO [Maybe TypedBinding]
-> TCMT IO (Expr -> ([Maybe TypedBinding], Expr))
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> Telescope -> ScopeM (AbsOfCon Telescope)
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract Telescope
tel TCMT IO (Expr -> ([Maybe TypedBinding], Expr))
-> ScopeM Expr -> ScopeM ([Maybe TypedBinding], Expr)
forall a b. TCMT IO (a -> b) -> TCMT IO a -> TCMT IO b
forall (f :: * -> *) a b. Applicative f => f (a -> b) -> f a -> f b
<*> Expr -> ScopeM (AbsOfCon Expr)
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract Expr
t
    return (A.GeneralizeTel binds (catMaybes tel), t)

-- ** Record directives
------------------------------------------------------------------------

-- | Check for duplicate record directives.
gatherRecordDirectives :: [C.RecordDirective] -> ScopeM C.RecordDirectives
gatherRecordDirectives :: [RecordDirective] -> ScopeM RecordDirectives
gatherRecordDirectives [RecordDirective]
ds = (RecordDirective -> StateT RecordDirectives (TCMT IO) ())
-> [RecordDirective] -> StateT RecordDirectives (TCMT IO) ()
forall (t :: * -> *) (m :: * -> *) a b.
(Foldable t, Monad m) =>
(a -> m b) -> t a -> m ()
mapM_ RecordDirective -> StateT RecordDirectives (TCMT IO) ()
gatherRecordDirective [RecordDirective]
ds StateT RecordDirectives (TCMT IO) ()
-> RecordDirectives -> ScopeM RecordDirectives
forall (m :: * -> *) s a. Monad m => StateT s m a -> s -> m s
`execStateT` RecordDirectives
forall a. Null a => a
empty

-- | Fill the respective field of 'C.RecordDirectives' by the given 'C.RecordDirective'.
--
-- Ignore it with a dead-code warning if the field is already filled.
--
gatherRecordDirective :: C.RecordDirective -> StateT C.RecordDirectives ScopeM ()
gatherRecordDirective :: RecordDirective -> StateT RecordDirectives (TCMT IO) ()
gatherRecordDirective RecordDirective
d = do
  dir@RecordDirectives{ recInductive = ind, recHasEta = eta, recPattern = pat, recConstructor = con } <- StateT RecordDirectives (TCMT IO) RecordDirectives
forall s (m :: * -> *). MonadState s m => m s
get
  case d of
    Induction Ranged Induction
ri         -> Maybe (Ranged Induction)
-> StateT RecordDirectives (TCMT IO) ()
-> StateT RecordDirectives (TCMT IO) ()
forall a.
Maybe a
-> StateT RecordDirectives (TCMT IO) ()
-> StateT RecordDirectives (TCMT IO) ()
assertNothing Maybe (Ranged Induction)
ind (StateT RecordDirectives (TCMT IO) ()
 -> StateT RecordDirectives (TCMT IO) ())
-> StateT RecordDirectives (TCMT IO) ()
-> StateT RecordDirectives (TCMT IO) ()
forall a b. (a -> b) -> a -> b
$ RecordDirectives -> StateT RecordDirectives (TCMT IO) ()
forall s (m :: * -> *). MonadState s m => s -> m ()
put RecordDirectives
dir{ recInductive   = Just ri }
    Eta Ranged HasEta0
re               -> Maybe (Ranged HasEta0)
-> StateT RecordDirectives (TCMT IO) ()
-> StateT RecordDirectives (TCMT IO) ()
forall a.
Maybe a
-> StateT RecordDirectives (TCMT IO) ()
-> StateT RecordDirectives (TCMT IO) ()
assertNothing Maybe (Ranged HasEta0)
eta (StateT RecordDirectives (TCMT IO) ()
 -> StateT RecordDirectives (TCMT IO) ())
-> StateT RecordDirectives (TCMT IO) ()
-> StateT RecordDirectives (TCMT IO) ()
forall a b. (a -> b) -> a -> b
$ RecordDirectives -> StateT RecordDirectives (TCMT IO) ()
forall s (m :: * -> *). MonadState s m => s -> m ()
put RecordDirectives
dir{ recHasEta      = Just re }
    PatternOrCopattern Range
r -> Maybe Range
-> StateT RecordDirectives (TCMT IO) ()
-> StateT RecordDirectives (TCMT IO) ()
forall a.
Maybe a
-> StateT RecordDirectives (TCMT IO) ()
-> StateT RecordDirectives (TCMT IO) ()
assertNothing Maybe Range
pat (StateT RecordDirectives (TCMT IO) ()
 -> StateT RecordDirectives (TCMT IO) ())
-> StateT RecordDirectives (TCMT IO) ()
-> StateT RecordDirectives (TCMT IO) ()
forall a b. (a -> b) -> a -> b
$ RecordDirectives -> StateT RecordDirectives (TCMT IO) ()
forall s (m :: * -> *). MonadState s m => s -> m ()
put RecordDirectives
dir{ recPattern     = Just r  }
    C.Constructor Name
x IsInstance
inst -> Maybe (Name, IsInstance)
-> StateT RecordDirectives (TCMT IO) ()
-> StateT RecordDirectives (TCMT IO) ()
forall a.
Maybe a
-> StateT RecordDirectives (TCMT IO) ()
-> StateT RecordDirectives (TCMT IO) ()
assertNothing Maybe (Name, IsInstance)
con (StateT RecordDirectives (TCMT IO) ()
 -> StateT RecordDirectives (TCMT IO) ())
-> StateT RecordDirectives (TCMT IO) ()
-> StateT RecordDirectives (TCMT IO) ()
forall a b. (a -> b) -> a -> b
$ RecordDirectives -> StateT RecordDirectives (TCMT IO) ()
forall s (m :: * -> *). MonadState s m => s -> m ()
put RecordDirectives
dir{ recConstructor = Just (x, inst) }
  where
    assertNothing :: Maybe a -> StateT C.RecordDirectives ScopeM () -> StateT C.RecordDirectives ScopeM ()
    assertNothing :: forall a.
Maybe a
-> StateT RecordDirectives (TCMT IO) ()
-> StateT RecordDirectives (TCMT IO) ()
assertNothing Maybe a
Nothing StateT RecordDirectives (TCMT IO) ()
cont = StateT RecordDirectives (TCMT IO) ()
cont
    assertNothing Just{}  StateT RecordDirectives (TCMT IO) ()
_    = TCMT IO () -> StateT RecordDirectives (TCMT IO) ()
forall (m :: * -> *) a.
Monad m =>
m a -> StateT RecordDirectives m a
forall (t :: (* -> *) -> * -> *) (m :: * -> *) a.
(MonadTrans t, Monad m) =>
m a -> t m a
lift (TCMT IO () -> StateT RecordDirectives (TCMT IO) ())
-> TCMT IO () -> StateT RecordDirectives (TCMT IO) ()
forall a b. (a -> b) -> a -> b
$ RecordDirective -> TCMT IO () -> TCMT IO ()
forall (m :: * -> *) x a.
(MonadTrace m, HasRange x) =>
x -> m a -> m a
setCurrentRange RecordDirective
d (TCMT IO () -> TCMT IO ()) -> TCMT IO () -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ Warning -> TCMT IO ()
forall (m :: * -> *).
(HasCallStack, MonadWarning m) =>
Warning -> m ()
warning (Warning -> TCMT IO ()) -> Warning -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ RecordDirective -> Warning
DuplicateRecordDirective RecordDirective
d

-- ** Helper functions for name clashes
------------------------------------------------------------------------

-- | Make sure definition is in same module as signature.
class LivesInCurrentModule a where
  livesInCurrentModule :: a -> ScopeM ()

instance LivesInCurrentModule AbstractName where
  livesInCurrentModule :: AbstractName -> TCMT IO ()
livesInCurrentModule = QName -> TCMT IO ()
forall a. LivesInCurrentModule a => a -> TCMT IO ()
livesInCurrentModule (QName -> TCMT IO ())
-> (AbstractName -> QName) -> AbstractName -> TCMT IO ()
forall b c a. (b -> c) -> (a -> b) -> a -> c
. AbstractName -> QName
anameName

instance LivesInCurrentModule A.QName where
  livesInCurrentModule :: QName -> TCMT IO ()
livesInCurrentModule QName
x = do
    m <- TCMT IO ModuleName
forall (m :: * -> *). ReadTCState m => m ModuleName
getCurrentModule
    reportS "scope.data.def" 30
      [ "  A.QName of data type: " ++ prettyShow x
      , "  current module: " ++ prettyShow m
      ]
    unless (A.qnameModule x == m) $
      typeError $ DefinitionInDifferentModule x

-- | Unless the resolved 'AbstractName' has the given 'KindOfName',
--   report a 'ClashingDefinition' for the 'C.Name'.
clashUnless :: C.Name -> KindOfName -> AbstractName -> ScopeM ()
clashUnless :: Name -> KindOfName -> AbstractName -> TCMT IO ()
clashUnless Name
x KindOfName
k AbstractName
ax = Bool -> TCMT IO () -> TCMT IO ()
forall b (m :: * -> *). (IsBool b, Monad m) => b -> m () -> m ()
unless (AbstractName -> KindOfName
anameKind AbstractName
ax KindOfName -> KindOfName -> Bool
forall a. Eq a => a -> a -> Bool
== KindOfName
k) (TCMT IO () -> TCMT IO ()) -> TCMT IO () -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$
  TypeError -> TCMT IO ()
forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
TypeError -> m a
typeError (TypeError -> TCMT IO ()) -> TypeError -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ QName -> QName -> Maybe NiceDeclaration -> TypeError
ClashingDefinition (Name -> QName
C.QName Name
x) (AbstractName -> QName
anameName AbstractName
ax) Maybe NiceDeclaration
forall a. Maybe a
Nothing

-- | If a (data/record) module with the given name is already present in the current module,
--   we take this as evidence that a data/record with that name is already defined.
clashIfModuleAlreadyDefinedInCurrentModule :: C.Name -> AbstractName -> ScopeM ()
clashIfModuleAlreadyDefinedInCurrentModule :: Name -> AbstractName -> TCMT IO ()
clashIfModuleAlreadyDefinedInCurrentModule Name
x AbstractName
ax = do
  datRecMods <- [Maybe DataOrRecordModule] -> [DataOrRecordModule]
forall a. [Maybe a] -> [a]
catMaybes ([Maybe DataOrRecordModule] -> [DataOrRecordModule])
-> TCMT IO [Maybe DataOrRecordModule]
-> TCMT IO [DataOrRecordModule]
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> do
    (AbstractModule -> TCMT IO (Maybe DataOrRecordModule))
-> [AbstractModule] -> TCMT IO [Maybe DataOrRecordModule]
forall (t :: * -> *) (m :: * -> *) a b.
(Traversable t, Monad m) =>
(a -> m b) -> t a -> m (t b)
forall (m :: * -> *) a b. Monad m => (a -> m b) -> [a] -> m [b]
mapM (ModuleName -> TCMT IO (Maybe DataOrRecordModule)
forall (m :: * -> *).
ReadTCState m =>
ModuleName -> m (Maybe DataOrRecordModule)
isDatatypeModule (ModuleName -> TCMT IO (Maybe DataOrRecordModule))
-> (AbstractModule -> ModuleName)
-> AbstractModule
-> TCMT IO (Maybe DataOrRecordModule)
forall b c a. (b -> c) -> (a -> b) -> a -> c
. AbstractModule -> ModuleName
amodName) ([AbstractModule] -> TCMT IO [Maybe DataOrRecordModule])
-> TCMT IO [AbstractModule] -> TCMT IO [Maybe DataOrRecordModule]
forall (m :: * -> *) a b. Monad m => (a -> m b) -> m a -> m b
=<< Name -> TCMT IO [AbstractModule]
lookupModuleInCurrentModule Name
x
  unlessNull datRecMods $ const $
    typeError $ ClashingDefinition (C.QName x) (anameName ax) Nothing

lookupModuleInCurrentModule :: C.Name -> ScopeM [AbstractModule]
lookupModuleInCurrentModule :: Name -> TCMT IO [AbstractModule]
lookupModuleInCurrentModule Name
x =
  Maybe (List1 AbstractModule) -> [AbstractModule]
forall a. Maybe (List1 a) -> [a]
List1.toList' (Maybe (List1 AbstractModule) -> [AbstractModule])
-> (Scope -> Maybe (List1 AbstractModule))
-> Scope
-> [AbstractModule]
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Name
-> Map Name (List1 AbstractModule) -> Maybe (List1 AbstractModule)
forall k a. Ord k => k -> Map k a -> Maybe a
Map.lookup Name
x (Map Name (List1 AbstractModule) -> Maybe (List1 AbstractModule))
-> (Scope -> Map Name (List1 AbstractModule))
-> Scope
-> Maybe (List1 AbstractModule)
forall b c a. (b -> c) -> (a -> b) -> a -> c
. NameSpace -> Map Name (List1 AbstractModule)
nsModules (NameSpace -> Map Name (List1 AbstractModule))
-> (Scope -> NameSpace) -> Scope -> Map Name (List1 AbstractModule)
forall b c a. (b -> c) -> (a -> b) -> a -> c
. [NameSpaceId] -> Scope -> NameSpace
thingsInScope [NameSpaceId
PublicNS, NameSpaceId
PrivateNS] (Scope -> [AbstractModule])
-> TCMT IO Scope -> TCMT IO [AbstractModule]
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> TCMT IO Scope
getCurrentScope

-- ** Helper functions for constructor declarations
------------------------------------------------------------------------

data DataConstrDecl = DataConstrDecl A.ModuleName IsAbstract Access C.NiceDeclaration

-- | Bind a @data@ constructor.
bindConstructorName
  :: ModuleName      -- ^ Name of @data@/@record@ module.
  -> C.Name          -- ^ Constructor name.
  -> IsAbstract
  -> Access
  -> ScopeM A.QName
bindConstructorName :: ModuleName -> Name -> IsAbstract -> Access -> TCMT IO QName
bindConstructorName ModuleName
m Name
x IsAbstract
a Access
p = do
  f <- Name -> ScopeM Fixity'
getConcreteFixity Name
x
  -- The abstract name is the qualified one
  y <- withCurrentModule m $ freshAbstractQName f x
  -- Bind it twice, once unqualified and once qualified
  bindName p' ConName x y
  withCurrentModule m $ bindName p'' ConName x y
  return y
  where
    -- An abstract constructor is private (abstract constructor means
    -- abstract datatype, so the constructor should not be exported).
    p' :: Access
p' = case IsAbstract
a of
           IsAbstract
AbstractDef -> Access
privateAccessInserted
           IsAbstract
_           -> Access
p
    p'' :: Access
p'' = case IsAbstract
a of
            IsAbstract
AbstractDef -> Access
privateAccessInserted
            IsAbstract
_           -> Access
PublicAccess

-- | Record constructors do not live in the record module (as it is parameterized).
--   Abstract constructors are bound privately, so that they are not exported.
bindRecordConstructorName :: C.Name -> KindOfName -> IsAbstract -> Access -> ScopeM A.QName
bindRecordConstructorName :: Name -> KindOfName -> IsAbstract -> Access -> TCMT IO QName
bindRecordConstructorName Name
x KindOfName
kind IsAbstract
a Access
p = do
  y <- Name -> TCMT IO QName
freshAbstractQName' Name
x
  bindName p' kind x y
  return y
  where
    -- An abstract constructor is private (abstract constructor means
    -- abstract datatype, so the constructor should not be exported).
    p' :: Access
p' = case IsAbstract
a of
           IsAbstract
AbstractDef -> Access
privateAccessInserted
           IsAbstract
_           -> Access
p

bindUnquoteConstructorName :: ModuleName -> Access -> C.Name -> TCM A.QName
bindUnquoteConstructorName :: ModuleName -> Access -> Name -> TCMT IO QName
bindUnquoteConstructorName ModuleName
m Access
p Name
c = do

  r <- QName -> ScopeM ResolvedName
resolveName (Name -> QName
C.QName Name
c)
  fc <- getConcreteFixity c
  c' <- withCurrentModule m $ freshAbstractQName fc c
  let aname QName
qn = QName -> KindOfName -> WhyInScope -> NameMetadata -> AbstractName
AbsName QName
qn KindOfName
QuotableName WhyInScope
Defined NameMetadata
NoMetadata
      addName = (Scope -> Scope) -> TCMT IO ()
modifyCurrentScope ((Scope -> Scope) -> TCMT IO ()) -> (Scope -> Scope) -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ NameSpaceId -> Name -> AbstractName -> Scope -> Scope
addNameToScope (Access -> NameSpaceId
localNameSpace Access
p) Name
c (AbstractName -> Scope -> Scope) -> AbstractName -> Scope -> Scope
forall a b. (a -> b) -> a -> b
$ QName -> AbstractName
aname QName
c'
      success = TCMT IO ()
addName TCMT IO () -> TCMT IO () -> TCMT IO ()
forall a b. TCMT IO a -> TCMT IO b -> TCMT IO b
forall (m :: * -> *) a b. Monad m => m a -> m b -> m b
>> (ModuleName -> TCMT IO () -> TCMT IO ()
forall (m :: * -> *) a.
(ReadTCState m, MonadTCState m) =>
ModuleName -> m a -> m a
withCurrentModule ModuleName
m (TCMT IO () -> TCMT IO ()) -> TCMT IO () -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ TCMT IO ()
addName)
      failure QName
y = TypeError -> TCMT IO ()
forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
TypeError -> m a
typeError (TypeError -> TCMT IO ()) -> TypeError -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ QName -> QName -> Maybe NiceDeclaration -> TypeError
ClashingDefinition (Name -> QName
C.QName Name
c) QName
y Maybe NiceDeclaration
forall a. Maybe a
Nothing
  case r of
    ResolvedName
_ | Name -> Bool
forall a. IsNoName a => a -> Bool
isNoName Name
c       -> TCMT IO ()
success
    ResolvedName
UnknownName          -> TCMT IO ()
success
    ConstructorName Set1 Induction
i List1 AbstractName
ds -> if (AbstractName -> Bool) -> List1 AbstractName -> Bool
forall (t :: * -> *) a. Foldable t => (a -> Bool) -> t a -> Bool
all (Maybe Induction -> Bool
forall a. Maybe a -> Bool
isJust (Maybe Induction -> Bool)
-> (AbstractName -> Maybe Induction) -> AbstractName -> Bool
forall b c a. (b -> c) -> (a -> b) -> a -> c
. KindOfName -> Maybe Induction
isConName (KindOfName -> Maybe Induction)
-> (AbstractName -> KindOfName) -> AbstractName -> Maybe Induction
forall b c a. (b -> c) -> (a -> b) -> a -> c
. AbstractName -> KindOfName
anameKind) List1 AbstractName
ds
      then TCMT IO ()
success
      else QName -> TCMT IO ()
failure (QName -> TCMT IO ()) -> QName -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ AbstractName -> QName
anameName (AbstractName -> QName) -> AbstractName -> QName
forall a b. (a -> b) -> a -> b
$ List1 AbstractName -> AbstractName
forall a. NonEmpty a -> a
List1.head List1 AbstractName
ds
    DefinedName Access
_ AbstractName
d Suffix
_    -> QName -> TCMT IO ()
failure (QName -> TCMT IO ()) -> QName -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ AbstractName -> QName
anameName AbstractName
d
    FieldName List1 AbstractName
ds         -> QName -> TCMT IO ()
failure (QName -> TCMT IO ()) -> QName -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ AbstractName -> QName
anameName (AbstractName -> QName) -> AbstractName -> QName
forall a b. (a -> b) -> a -> b
$ List1 AbstractName -> AbstractName
forall a. NonEmpty a -> a
List1.head List1 AbstractName
ds
    PatternSynResName List1 AbstractName
ds -> QName -> TCMT IO ()
failure (QName -> TCMT IO ()) -> QName -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ AbstractName -> QName
anameName (AbstractName -> QName) -> AbstractName -> QName
forall a b. (a -> b) -> a -> b
$ List1 AbstractName -> AbstractName
forall a. NonEmpty a -> a
List1.head List1 AbstractName
ds
    VarName Name
y BindingSource
_          -> QName -> TCMT IO ()
failure (QName -> TCMT IO ()) -> QName -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ Name -> QName
qualify_ Name
y
  return c'

instance ToAbstract DataConstrDecl where
  type AbsOfCon DataConstrDecl = A.Declaration

  toAbstract :: DataConstrDecl -> ScopeM (AbsOfCon DataConstrDecl)
toAbstract (DataConstrDecl ModuleName
m IsAbstract
a Access
p NiceDeclaration
d) = Call -> TCMT IO Declaration -> TCMT IO Declaration
forall a. Call -> TCMT IO a -> TCMT IO a
forall (m :: * -> *) a. MonadTrace m => Call -> m a -> m a
traceCall (NiceDeclaration -> Call
ScopeCheckDeclaration NiceDeclaration
d) do
    case NiceDeclaration
d of
      C.Axiom Range
r Access
p1 IsAbstract
a1 IsInstance
i ArgInfo
info Name
x Expr
t -> do -- rel==Relevant
        -- unless (p1 == p) __IMPOSSIBLE__  -- This invariant is currently violated by test/Succeed/Issue282.agda
        Bool -> TCMT IO () -> TCMT IO ()
forall b (m :: * -> *). (IsBool b, Monad m) => b -> m () -> m ()
unless (IsAbstract
a1 IsAbstract -> IsAbstract -> Bool
forall a. Eq a => a -> a -> Bool
== IsAbstract
a) TCMT IO ()
forall a. HasCallStack => a
__IMPOSSIBLE__
        let msg :: [Char]
msg = [Char]
"of constructor"
        info <- [Char] -> ArgInfo -> ScopeM ArgInfo
forall a. LensQuantity a => [Char] -> a -> ScopeM a
ensureNotLinear [Char]
msg (ArgInfo -> ScopeM ArgInfo) -> ScopeM ArgInfo -> ScopeM ArgInfo
forall (m :: * -> *) a b. Monad m => (a -> m b) -> m a -> m b
=<< [Char] -> ArgInfo -> ScopeM ArgInfo
forall a. LensRelevance a => [Char] -> a -> ScopeM a
ensureRelevant [Char]
msg ArgInfo
info
        t' <- toAbstractCtx TopCtx t
        -- The abstract name is the qualified one
        -- Bind it twice, once unqualified and once qualified
        f <- getConcreteFixity x
        y <- bindConstructorName m x a p
        printScope "con" 25 "bound constructor"
        return $ A.Axiom ConName (mkDefInfoInstance x f p a i NotMacroDef r)
                         info Nothing y t'
      NiceDeclaration
_ -> NiceDeclaration -> TCMT IO Declaration
forall a. NiceDeclaration -> ScopeM a
errorNotConstrDecl NiceDeclaration
d

errorNotConstrDecl :: C.NiceDeclaration -> ScopeM a
errorNotConstrDecl :: forall a. NiceDeclaration -> ScopeM a
errorNotConstrDecl NiceDeclaration
d = NiceDeclaration -> TCMT IO a -> TCMT IO a
forall (m :: * -> *) x a.
(MonadTrace m, HasRange x) =>
x -> m a -> m a
setCurrentRange NiceDeclaration
d (TCMT IO a -> TCMT IO a) -> TCMT IO a -> TCMT IO a
forall a b. (a -> b) -> a -> b
$
  TypeError -> TCMT IO a
forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
TypeError -> m a
typeError (TypeError -> TCMT IO a) -> TypeError -> TCMT IO a
forall a b. (a -> b) -> a -> b
$ List1 Declaration -> TypeError
IllegalDeclarationInDataDefinition (List1 Declaration -> TypeError) -> List1 Declaration -> TypeError
forall a b. (a -> b) -> a -> b
$ NiceDeclaration -> List1 Declaration
notSoNiceDeclarations NiceDeclaration
d

ensureRelevant :: LensRelevance a => String -> a -> ScopeM a
ensureRelevant :: forall a. LensRelevance a => [Char] -> a -> ScopeM a
ensureRelevant [Char]
s a
info = do
  if a -> Bool
forall a. LensRelevance a => a -> Bool
isRelevant a
info then a -> TCMT IO a
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return a
info else do
    Warning -> TCMT IO ()
forall (m :: * -> *).
(HasCallStack, MonadWarning m) =>
Warning -> m ()
warning (Warning -> TCMT IO ()) -> Warning -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ [Char] -> Relevance -> Relevance -> Warning
FixingRelevance [Char]
s (a -> Relevance
forall a. LensRelevance a => a -> Relevance
getRelevance a
info) Relevance
relevant
    a -> TCMT IO a
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return (a -> TCMT IO a) -> a -> TCMT IO a
forall a b. (a -> b) -> a -> b
$ Relevance -> a -> a
forall a. LensRelevance a => Relevance -> a -> a
setRelevance Relevance
relevant a
info

ensureNotLinear :: LensQuantity a => String -> a -> ScopeM a
ensureNotLinear :: forall a. LensQuantity a => [Char] -> a -> ScopeM a
ensureNotLinear [Char]
s a
info = do
  case a -> Quantity
forall a. LensQuantity a => a -> Quantity
getQuantity a
info of
    Quantityω{} -> a -> ScopeM a
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return a
info
    Quantity0{} -> a -> ScopeM a
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return a
info
    q :: Quantity
q@Quantity1{} -> do
      -- Andreas, 2024-08-24, "@1" is still not parsed, so this is impossible.
      __IMPOSSIBLE__
      -- TODO: linearity
      -- let q' = Quantityω QωInferred
      -- warning $ FixingQuantity s q q'
      -- return $ setQuantity q' info

-- ** More scope checking
------------------------------------------------------------------------

instance ToAbstract C.Pragma where
  type AbsOfCon C.Pragma = [A.Pragma]

  toAbstract :: Pragma -> ScopeM (AbsOfCon Pragma)
toAbstract (C.ImpossiblePragma Range
_ [[Char]]
strs) =
    case [[Char]]
strs of
      [Char]
"ReduceM" : [[Char]]
_ -> [[Char]] -> TCMT IO [Pragma]
forall a. HasCallStack => [[Char]] -> TCM a
impossibleTestReduceM [[Char]]
strs
      [[Char]]
_ -> [[Char]] -> TCMT IO [Pragma]
forall (m :: * -> *) a.
(MonadDebug m, HasCallStack) =>
[[Char]] -> m a
impossibleTest [[Char]]
strs

  toAbstract (C.OptionsPragma Range
_ [[Char]]
opts) = [Pragma] -> TCMT IO [Pragma]
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return [ [[Char]] -> Pragma
A.OptionsPragma [[Char]]
opts ]

  toAbstract (C.RewritePragma Range
_ Range
_ []) = [] [Pragma] -> TCMT IO () -> TCMT IO [Pragma]
forall a b. a -> TCMT IO b -> TCMT IO a
forall (f :: * -> *) a b. Functor f => a -> f b -> f a
<$ Warning -> TCMT IO ()
forall (m :: * -> *).
(HasCallStack, MonadWarning m) =>
Warning -> m ()
warning Warning
EmptyRewritePragma
  toAbstract (C.RewritePragma Range
_ Range
r [QName]
xs) = Pragma -> [Pragma]
forall el coll. Singleton el coll => el -> coll
singleton (Pragma -> [Pragma])
-> ([Maybe QName] -> Pragma) -> [Maybe QName] -> [Pragma]
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Range -> [QName] -> Pragma
A.RewritePragma Range
r ([QName] -> Pragma)
-> ([Maybe QName] -> [QName]) -> [Maybe QName] -> Pragma
forall b c a. (b -> c) -> (a -> b) -> a -> c
. [Maybe QName] -> [QName]
forall a. [Maybe a] -> [a]
catMaybes ([Maybe QName] -> [Pragma])
-> TCMT IO [Maybe QName] -> TCMT IO [Pragma]
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> do
    [QName]
-> (QName -> TCMT IO (Maybe QName)) -> TCMT IO [Maybe QName]
forall (t :: * -> *) (m :: * -> *) a b.
(Traversable t, Monad m) =>
t a -> (a -> m b) -> m (t b)
forM [QName]
xs \ QName
x -> QName -> TCMT IO (Maybe QName) -> TCMT IO (Maybe QName)
forall (m :: * -> *) x a.
(MonadTrace m, HasRange x) =>
x -> m a -> m a
setCurrentRange QName
x (TCMT IO (Maybe QName) -> TCMT IO (Maybe QName))
-> TCMT IO (Maybe QName) -> TCMT IO (Maybe QName)
forall a b. (a -> b) -> a -> b
$ (QName -> IsAmbiguous -> Warning) -> QName -> TCMT IO (Maybe QName)
unambiguousConOrDef QName -> IsAmbiguous -> Warning
NotARewriteRule QName
x

  toAbstract (C.ForeignPragma Range
_ Ranged Text
rb [Char]
s) = [] [Pragma] -> TCMT IO () -> TCMT IO [Pragma]
forall a b. a -> TCMT IO b -> TCMT IO a
forall (f :: * -> *) a b. Functor f => a -> f b -> f a
<$ Text -> [Char] -> TCMT IO ()
addForeignCode (Ranged Text -> Text
forall a. Ranged a -> a
rangedThing Ranged Text
rb) [Char]
s

  toAbstract (C.CompilePragma Range
_ Ranged Text
rb QName
x [Char]
s) =
    [Pragma] -> (QName -> [Pragma]) -> Maybe QName -> [Pragma]
forall b a. b -> (a -> b) -> Maybe a -> b
maybe [] (\ QName
y -> [ Ranged Text -> QName -> [Char] -> Pragma
A.CompilePragma Ranged Text
rb QName
y [Char]
s ]) (Maybe QName -> [Pragma])
-> TCMT IO (Maybe QName) -> TCMT IO [Pragma]
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$>
      (QName -> IsAmbiguous -> Warning) -> QName -> TCMT IO (Maybe QName)
unambiguousConOrDef QName -> IsAmbiguous -> Warning
PragmaCompileWrongName QName
x

  toAbstract (C.StaticPragma Range
_ QName
x) = do
    (QName -> Pragma) -> [QName] -> [Pragma]
forall a b. (a -> b) -> [a] -> [b]
map QName -> Pragma
A.StaticPragma ([QName] -> [Pragma])
-> (Maybe QName -> [QName]) -> Maybe QName -> [Pragma]
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Maybe QName -> [QName]
forall a. Maybe a -> [a]
maybeToList (Maybe QName -> [Pragma])
-> TCMT IO (Maybe QName) -> TCMT IO [Pragma]
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> do
      (QName -> IsAmbiguous -> Warning) -> QName -> TCMT IO (Maybe QName)
unambiguousDef ([Char] -> QName -> IsAmbiguous -> Warning
PragmaExpectsUnambiguousProjectionOrFunction [Char]
"STATIC") QName
x

  toAbstract (C.InjectivePragma Range
_ QName
x) = do
    (QName -> Pragma) -> [QName] -> [Pragma]
forall a b. (a -> b) -> [a] -> [b]
map QName -> Pragma
A.InjectivePragma ([QName] -> [Pragma])
-> (Maybe QName -> [QName]) -> Maybe QName -> [Pragma]
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Maybe QName -> [QName]
forall a. Maybe a -> [a]
maybeToList (Maybe QName -> [Pragma])
-> TCMT IO (Maybe QName) -> TCMT IO [Pragma]
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> do
      (QName -> IsAmbiguous -> Warning) -> QName -> TCMT IO (Maybe QName)
unambiguousDef ([Char] -> QName -> IsAmbiguous -> Warning
PragmaExpectsUnambiguousProjectionOrFunction [Char]
"INJECTIVE") QName
x

  toAbstract (C.InjectiveForInferencePragma Range
_ QName
x) = do
    (QName -> Pragma) -> [QName] -> [Pragma]
forall a b. (a -> b) -> [a] -> [b]
map QName -> Pragma
A.InjectiveForInferencePragma ([QName] -> [Pragma])
-> (Maybe QName -> [QName]) -> Maybe QName -> [Pragma]
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Maybe QName -> [QName]
forall a. Maybe a -> [a]
maybeToList (Maybe QName -> [Pragma])
-> TCMT IO (Maybe QName) -> TCMT IO [Pragma]
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> do
      (QName -> Warning) -> QName -> TCMT IO (Maybe QName)
scopeCheckDef ([Char] -> QName -> Warning
PragmaExpectsDefinedSymbol [Char]
"INJECTIVE_FOR_INFERENCE") QName
x

  toAbstract pragma :: Pragma
pragma@(C.InlinePragma Range
_ Bool
b QName
x) = do
      TCMT IO (Maybe Expr)
-> TCMT IO [Pragma]
-> (Expr -> TCMT IO [Pragma])
-> TCMT IO [Pragma]
forall (m :: * -> *) a b.
Monad m =>
m (Maybe a) -> m b -> (a -> m b) -> m b
caseMaybeM (MaybeOldQName -> ScopeM (AbsOfCon MaybeOldQName)
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract (MaybeOldQName -> ScopeM (AbsOfCon MaybeOldQName))
-> MaybeOldQName -> ScopeM (AbsOfCon MaybeOldQName)
forall a b. (a -> b) -> a -> b
$ OldQName -> MaybeOldQName
MaybeOldQName (OldQName -> MaybeOldQName) -> OldQName -> MaybeOldQName
forall a b. (a -> b) -> a -> b
$ QName -> Maybe (Set1 Name) -> OldQName
OldQName QName
x Maybe (Set1 Name)
forall a. Maybe a
Nothing) TCMT IO [Pragma]
notInScope \case
        A.Con (AmbQ List1 QName
xs)                -> (QName -> TCMT IO [Pragma]) -> [QName] -> TCMT IO [Pragma]
forall (m :: * -> *) a b. Monad m => (a -> m [b]) -> [a] -> m [b]
concatMapM QName -> TCMT IO [Pragma]
ret ([QName] -> TCMT IO [Pragma]) -> [QName] -> TCMT IO [Pragma]
forall a b. (a -> b) -> a -> b
$ List1 QName -> [Item (List1 QName)]
forall l. IsList l => l -> [Item l]
List1.toList List1 QName
xs
        A.Def QName
x                        -> QName -> TCMT IO [Pragma]
ret QName
x
        A.Proj ProjOrigin
_ AmbiguousQName
p
          | Just QName
x <- AmbiguousQName -> Maybe QName
getUnambiguous AmbiguousQName
p -> QName -> TCMT IO [Pragma]
ret QName
x
          | Bool
otherwise                  -> Pragma -> [Char] -> TCMT IO [Pragma]
forall p a. HasRange p => p -> [Char] -> ScopeM [a]
uselessPragma Pragma
pragma ([Char] -> TCMT IO [Pragma]) -> [Char] -> TCMT IO [Pragma]
forall a b. (a -> b) -> a -> b
$ [Char]
sINLINE [Char] -> [Char] -> [Char]
forall a. [a] -> [a] -> [a]
++ [Char]
" used on ambiguous name " [Char] -> [Char] -> [Char]
forall a. [a] -> [a] -> [a]
++ QName -> [Char]
forall a. Pretty a => a -> [Char]
prettyShow QName
x
        Expr
_ -> Pragma -> [Char] -> TCMT IO [Pragma]
forall p a. HasRange p => p -> [Char] -> ScopeM [a]
uselessPragma Pragma
pragma ([Char] -> TCMT IO [Pragma]) -> [Char] -> TCMT IO [Pragma]
forall a b. (a -> b) -> a -> b
$ [Char]
"Target of " [Char] -> [Char] -> [Char]
forall a. [a] -> [a] -> [a]
++ [Char]
sINLINE [Char] -> [Char] -> [Char]
forall a. [a] -> [a] -> [a]
++ [Char]
" pragma should be a function or constructor"
    where
      sINLINE :: [Char]
sINLINE    = if Bool
b then [Char]
"INLINE" else [Char]
"NOINLINE"
      notInScope :: TCMT IO [Pragma]
notInScope = [] [Pragma] -> TCMT IO () -> TCMT IO [Pragma]
forall a b. a -> TCMT IO b -> TCMT IO a
forall (f :: * -> *) a b. Functor f => a -> f b -> f a
<$ QName -> TCMT IO ()
notInScopeWarning QName
x
      ret :: QName -> TCMT IO [Pragma]
ret QName
y      = [Pragma] -> TCMT IO [Pragma]
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return [ Bool -> QName -> Pragma
A.InlinePragma Bool
b QName
y ]

  toAbstract (C.NotProjectionLikePragma Range
_ QName
x) = do
    (QName -> Pragma) -> [QName] -> [Pragma]
forall a b. (a -> b) -> [a] -> [b]
map QName -> Pragma
A.NotProjectionLikePragma ([QName] -> [Pragma])
-> (Maybe QName -> [QName]) -> Maybe QName -> [Pragma]
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Maybe QName -> [QName]
forall a. Maybe a -> [a]
maybeToList (Maybe QName -> [Pragma])
-> TCMT IO (Maybe QName) -> TCMT IO [Pragma]
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> do
      (QName -> IsAmbiguous -> Warning) -> QName -> TCMT IO (Maybe QName)
unambiguousDef ([Char] -> QName -> IsAmbiguous -> Warning
PragmaExpectsUnambiguousProjectionOrFunction [Char]
"NOT_PROJECTION_LIKE") QName
x

  toAbstract (C.OverlapPragma Range
_ [QName]
xs OverlapMode
i) = do
    (QName -> Pragma) -> [QName] -> [Pragma]
forall a b. (a -> b) -> [a] -> [b]
map ((QName -> OverlapMode -> Pragma) -> OverlapMode -> QName -> Pragma
forall a b c. (a -> b -> c) -> b -> a -> c
flip QName -> OverlapMode -> Pragma
A.OverlapPragma OverlapMode
i) ([QName] -> [Pragma])
-> ([Maybe QName] -> [QName]) -> [Maybe QName] -> [Pragma]
forall b c a. (b -> c) -> (a -> b) -> a -> c
. [Maybe QName] -> [QName]
forall a. [Maybe a] -> [a]
catMaybes ([Maybe QName] -> [Pragma])
-> TCMT IO [Maybe QName] -> TCMT IO [Pragma]
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> do
      (QName -> TCMT IO (Maybe QName))
-> [QName] -> TCMT IO [Maybe QName]
forall (t :: * -> *) (m :: * -> *) a b.
(Traversable t, Monad m) =>
(a -> m b) -> t a -> m (t b)
forall (m :: * -> *) a b. Monad m => (a -> m b) -> [a] -> m [b]
mapM ((QName -> IsAmbiguous -> Warning) -> QName -> TCMT IO (Maybe QName)
unambiguousConOrDef ((QName -> IsAmbiguous -> Warning)
 -> QName -> TCMT IO (Maybe QName))
-> (QName -> IsAmbiguous -> Warning)
-> QName
-> TCMT IO (Maybe QName)
forall a b. (a -> b) -> a -> b
$ [Char] -> QName -> IsAmbiguous -> Warning
PragmaExpectsUnambiguousConstructorOrFunction [Char]
pragma) [QName]
xs
    where
      pragma :: [Char]
pragma = case OverlapMode
i of
        OverlapMode
Overlappable -> [Char]
"OVERLAPPABLE"
        OverlapMode
Overlapping  -> [Char]
"OVERLAPPING"
        OverlapMode
Overlaps     -> [Char]
"OVERLAPS"
        OverlapMode
Incoherent   -> [Char]
"INCOHERENT"
        -- Never written by the user:
        OverlapMode
DefaultOverlap -> [Char]
forall a. HasCallStack => a
__IMPOSSIBLE__
        OverlapMode
FieldOverlap   -> [Char]
forall a. HasCallStack => a
__IMPOSSIBLE__

  toAbstract pragma :: Pragma
pragma@(C.BuiltinPragma Range
_ RString
rb QName
qx)
    | Just BuiltinId
b' <- Maybe BuiltinId
b, BuiltinId -> Bool
isUntypedBuiltin BuiltinId
b' = do
        q <- ResolveQName -> ScopeM (AbsOfCon ResolveQName)
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract (ResolveQName -> ScopeM (AbsOfCon ResolveQName))
-> ResolveQName -> ScopeM (AbsOfCon ResolveQName)
forall a b. (a -> b) -> a -> b
$ QName -> ResolveQName
ResolveQName QName
qx
        bindUntypedBuiltin b' q
        return [ A.BuiltinPragma rb q ]
        -- Andreas, 2015-02-14
        -- Some builtins cannot be given a valid Agda type,
        -- thus, they do not come with accompanying postulate or definition.
    | Just BuiltinId
b' <- Maybe BuiltinId
b, BuiltinId -> Bool
isBuiltinNoDef BuiltinId
b' = do
          case QName
qx of
            C.QName Name
x -> do
              -- The name shouldn't exist yet. If it does, we raise a warning
              -- and drop the existing definition.
              TCMT IO Bool -> TCMT IO () -> TCMT IO ()
forall (m :: * -> *). Monad m => m Bool -> m () -> m ()
unlessM ((ResolvedName
UnknownName ResolvedName -> ResolvedName -> Bool
forall a. Eq a => a -> a -> Bool
==) (ResolvedName -> Bool) -> ScopeM ResolvedName -> TCMT IO Bool
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> QName -> ScopeM ResolvedName
resolveName QName
qx) (TCMT IO () -> TCMT IO ()) -> TCMT IO () -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ do
                Warning -> TCMT IO ()
forall (m :: * -> *).
(HasCallStack, MonadWarning m) =>
Warning -> m ()
warning (Warning -> TCMT IO ()) -> Warning -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ BuiltinId -> Warning
BuiltinDeclaresIdentifier BuiltinId
b'
                (Scope -> Scope) -> TCMT IO ()
modifyCurrentScope ((Scope -> Scope) -> TCMT IO ()) -> (Scope -> Scope) -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ NameSpaceId -> Name -> Scope -> Scope
removeNameFromScope NameSpaceId
PublicNS Name
x
              -- We then happily bind the name
              y <- Name -> TCMT IO QName
freshAbstractQName' Name
x
              let kind = KindOfName -> Maybe KindOfName -> KindOfName
forall a. a -> Maybe a -> a
fromMaybe KindOfName
forall a. HasCallStack => a
__IMPOSSIBLE__ (Maybe KindOfName -> KindOfName) -> Maybe KindOfName -> KindOfName
forall a b. (a -> b) -> a -> b
$ BuiltinId -> Maybe KindOfName
builtinKindOfName BuiltinId
b'
              bindName PublicAccess kind x y
              return [ A.BuiltinNoDefPragma rb kind y ]
            QName
_ -> Pragma -> [Char] -> TCMT IO [Pragma]
forall p a. HasRange p => p -> [Char] -> ScopeM [a]
uselessPragma Pragma
pragma ([Char] -> TCMT IO [Pragma]) -> [Char] -> TCMT IO [Pragma]
forall a b. (a -> b) -> a -> b
$
              [Char]
"Pragma BUILTIN " [Char] -> [Char] -> [Char]
forall a. [a] -> [a] -> [a]
++ BuiltinId -> [Char]
forall a. IsBuiltin a => a -> [Char]
getBuiltinId BuiltinId
b' [Char] -> [Char] -> [Char]
forall a. [a] -> [a] -> [a]
++ [Char]
": expected unqualified identifier, " [Char] -> [Char] -> [Char]
forall a. [a] -> [a] -> [a]
++
              [Char]
"but found " [Char] -> [Char] -> [Char]
forall a. [a] -> [a] -> [a]
++ QName -> [Char]
forall a. Pretty a => a -> [Char]
prettyShow QName
qx
    | Bool
otherwise = do
          q0 <- ResolveQName -> ScopeM (AbsOfCon ResolveQName)
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract (ResolveQName -> ScopeM (AbsOfCon ResolveQName))
-> ResolveQName -> ScopeM (AbsOfCon ResolveQName)
forall a b. (a -> b) -> a -> b
$ QName -> ResolveQName
ResolveQName QName
qx

          -- Andreas, 2020-04-12, pr #4574.  For highlighting purposes:
          -- Rebind 'BuiltinPrim' as 'PrimName' and similar.
          q <- case (q0, b >>= builtinKindOfName, qx) of
            (DefinedName Access
acc AbstractName
y Suffix
suffix, Just KindOfName
kind, C.QName Name
x)
              | AbstractName -> KindOfName
anameKind AbstractName
y KindOfName -> KindOfName -> Bool
forall a. Eq a => a -> a -> Bool
/= KindOfName
kind
              , KindOfName
kind KindOfName -> [KindOfName] -> Bool
forall a. Eq a => a -> [a] -> Bool
forall (t :: * -> *) a. (Foldable t, Eq a) => a -> t a -> Bool
`elem` [ KindOfName
PrimName, KindOfName
AxiomName ] -> do
                  Access -> KindOfName -> Name -> QName -> TCMT IO ()
rebindName Access
acc KindOfName
kind Name
x (QName -> TCMT IO ()) -> QName -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ AbstractName -> QName
anameName AbstractName
y
                  ResolvedName -> ScopeM ResolvedName
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return (ResolvedName -> ScopeM ResolvedName)
-> ResolvedName -> ScopeM ResolvedName
forall a b. (a -> b) -> a -> b
$ Access -> AbstractName -> Suffix -> ResolvedName
DefinedName Access
acc AbstractName
y{ anameKind = kind } Suffix
suffix
            (ResolvedName, Maybe KindOfName, QName)
_ -> ResolvedName -> ScopeM ResolvedName
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return ResolvedName
q0

          return [ A.BuiltinPragma rb q ]
    where b :: Maybe BuiltinId
b = [Char] -> Maybe BuiltinId
builtinById (RString -> [Char]
forall a. Ranged a -> a
rangedThing RString
rb)

  toAbstract (C.EtaPragma Range
_ QName
x) = do
    (QName -> Pragma) -> [QName] -> [Pragma]
forall a b. (a -> b) -> [a] -> [b]
map QName -> Pragma
A.EtaPragma ([QName] -> [Pragma])
-> (Maybe QName -> [QName]) -> Maybe QName -> [Pragma]
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Maybe QName -> [QName]
forall a. Maybe a -> [a]
maybeToList (Maybe QName -> [Pragma])
-> TCMT IO (Maybe QName) -> TCMT IO [Pragma]
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> do
      (QName -> Warning) -> QName -> TCMT IO (Maybe QName)
scopeCheckDef ([Char] -> QName -> Warning
PragmaExpectsDefinedSymbol [Char]
"ETA") QName
x

  toAbstract pragma :: Pragma
pragma@(C.DisplayPragma Range
_ Pattern
lhs Expr
rhs) = do
    Maybe Pragma -> [Pragma]
forall a. Maybe a -> [a]
maybeToList (Maybe Pragma -> [Pragma])
-> TCMT IO (Maybe Pragma) -> TCMT IO [Pragma]
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> do
      TCMT IO (Maybe Pragma) -> TCMT IO (Maybe Pragma)
forall a. ScopeM a -> ScopeM a
withLocalVars (TCMT IO (Maybe Pragma) -> TCMT IO (Maybe Pragma))
-> TCMT IO (Maybe Pragma) -> TCMT IO (Maybe Pragma)
forall a b. (a -> b) -> a -> b
$ MaybeT (TCMT IO) Pragma -> TCMT IO (Maybe Pragma)
forall (m :: * -> *) a. MaybeT m a -> m (Maybe a)
runMaybeT do
        let err :: MaybeT (TCMT IO) a
err = [Char] -> MaybeT (TCMT IO) a
forall a. [Char] -> MaybeT (TCMT IO) a
failure [Char]
"DISPLAY pragma left-hand side must have form 'f e1 .. en'"
            getHead :: Pattern -> MaybeT (TCMT IO) QName
getHead (C.IdentP Bool
_ QName
x)              = QName -> MaybeT (TCMT IO) QName
forall a. a -> MaybeT (TCMT IO) a
forall (m :: * -> *) a. Monad m => a -> m a
return QName
x
            getHead (C.RawAppP Range
_ (List2 Pattern
p Pattern
_ [Pattern]
_)) = Pattern -> MaybeT (TCMT IO) QName
getHead Pattern
p
            getHead Pattern
_ = MaybeT (TCMT IO) QName
forall {a}. MaybeT (TCMT IO) a
err

        top <- Pattern -> MaybeT (TCMT IO) QName
getHead Pattern
lhs

        (isPatSyn, hd) <- do
          qx <- liftTCM $ resolveName' allKindsOfNames Nothing top
          case qx of
            VarName Name
x' BindingSource
_                -> (Bool, QName) -> MaybeT (TCMT IO) (Bool, QName)
forall a. a -> MaybeT (TCMT IO) a
forall (m :: * -> *) a. Monad m => a -> m a
return ((Bool, QName) -> MaybeT (TCMT IO) (Bool, QName))
-> (QName -> (Bool, QName))
-> QName
-> MaybeT (TCMT IO) (Bool, QName)
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (Bool
False,) (QName -> MaybeT (TCMT IO) (Bool, QName))
-> QName -> MaybeT (TCMT IO) (Bool, QName)
forall a b. (a -> b) -> a -> b
$ List1 Name -> QName
A.qnameFromList (List1 Name -> QName) -> List1 Name -> QName
forall a b. (a -> b) -> a -> b
$ Name -> List1 Name
forall el coll. Singleton el coll => el -> coll
singleton Name
x'
            DefinedName Access
_ AbstractName
d Suffix
NoSuffix    -> (Bool, QName) -> MaybeT (TCMT IO) (Bool, QName)
forall a. a -> MaybeT (TCMT IO) a
forall (m :: * -> *) a. Monad m => a -> m a
return ((Bool, QName) -> MaybeT (TCMT IO) (Bool, QName))
-> (QName -> (Bool, QName))
-> QName
-> MaybeT (TCMT IO) (Bool, QName)
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (Bool
False,) (QName -> MaybeT (TCMT IO) (Bool, QName))
-> QName -> MaybeT (TCMT IO) (Bool, QName)
forall a b. (a -> b) -> a -> b
$ AbstractName -> QName
anameName AbstractName
d
            DefinedName Access
_ AbstractName
d Suffix{}    -> [Char] -> MaybeT (TCMT IO) (Bool, QName)
forall a. [Char] -> MaybeT (TCMT IO) a
failure ([Char] -> MaybeT (TCMT IO) (Bool, QName))
-> [Char] -> MaybeT (TCMT IO) (Bool, QName)
forall a b. (a -> b) -> a -> b
$ [Char]
"Invalid pattern " [Char] -> [Char] -> [Char]
forall a. [a] -> [a] -> [a]
++ QName -> [Char]
forall a. Pretty a => a -> [Char]
prettyShow QName
top
            FieldName     (AbstractName
d :| [])     -> (Bool, QName) -> MaybeT (TCMT IO) (Bool, QName)
forall a. a -> MaybeT (TCMT IO) a
forall (m :: * -> *) a. Monad m => a -> m a
return ((Bool, QName) -> MaybeT (TCMT IO) (Bool, QName))
-> (QName -> (Bool, QName))
-> QName
-> MaybeT (TCMT IO) (Bool, QName)
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (Bool
False,) (QName -> MaybeT (TCMT IO) (Bool, QName))
-> QName -> MaybeT (TCMT IO) (Bool, QName)
forall a b. (a -> b) -> a -> b
$ AbstractName -> QName
anameName AbstractName
d
            FieldName List1 AbstractName
ds                -> [Char] -> MaybeT (TCMT IO) (Bool, QName)
forall a. [Char] -> MaybeT (TCMT IO) a
failure ([Char] -> MaybeT (TCMT IO) (Bool, QName))
-> [Char] -> MaybeT (TCMT IO) (Bool, QName)
forall a b. (a -> b) -> a -> b
$ [Char]
"Ambiguous projection " [Char] -> [Char] -> [Char]
forall a. [a] -> [a] -> [a]
++ QName -> [Char]
forall a. Pretty a => a -> [Char]
prettyShow QName
top [Char] -> [Char] -> [Char]
forall a. [a] -> [a] -> [a]
++ [Char]
": " [Char] -> [Char] -> [Char]
forall a. [a] -> [a] -> [a]
++ AmbiguousQName -> [Char]
forall a. Pretty a => a -> [Char]
prettyShow (List1 QName -> AmbiguousQName
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
ds)
            ConstructorName Set1 Induction
_ (AbstractName
d :| []) -> (Bool, QName) -> MaybeT (TCMT IO) (Bool, QName)
forall a. a -> MaybeT (TCMT IO) a
forall (m :: * -> *) a. Monad m => a -> m a
return ((Bool, QName) -> MaybeT (TCMT IO) (Bool, QName))
-> (QName -> (Bool, QName))
-> QName
-> MaybeT (TCMT IO) (Bool, QName)
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (Bool
False,) (QName -> MaybeT (TCMT IO) (Bool, QName))
-> QName -> MaybeT (TCMT IO) (Bool, QName)
forall a b. (a -> b) -> a -> b
$ AbstractName -> QName
anameName AbstractName
d
            ConstructorName Set1 Induction
_ List1 AbstractName
ds        -> [Char] -> MaybeT (TCMT IO) (Bool, QName)
forall a. [Char] -> MaybeT (TCMT IO) a
failure ([Char] -> MaybeT (TCMT IO) (Bool, QName))
-> [Char] -> MaybeT (TCMT IO) (Bool, QName)
forall a b. (a -> b) -> a -> b
$ [Char]
"Ambiguous constructor " [Char] -> [Char] -> [Char]
forall a. [a] -> [a] -> [a]
++ QName -> [Char]
forall a. Pretty a => a -> [Char]
prettyShow QName
top [Char] -> [Char] -> [Char]
forall a. [a] -> [a] -> [a]
++ [Char]
": " [Char] -> [Char] -> [Char]
forall a. [a] -> [a] -> [a]
++ AmbiguousQName -> [Char]
forall a. Pretty a => a -> [Char]
prettyShow (List1 QName -> AmbiguousQName
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
ds)
            ResolvedName
UnknownName                 -> do TCMT IO () -> MaybeT (TCMT IO) ()
forall a. TCM a -> MaybeT (TCMT IO) a
forall (tcm :: * -> *) a. MonadTCM tcm => TCM a -> tcm a
liftTCM (TCMT IO () -> MaybeT (TCMT IO) ())
-> TCMT IO () -> MaybeT (TCMT IO) ()
forall a b. (a -> b) -> a -> b
$ QName -> TCMT IO ()
notInScopeWarning QName
top; MaybeT (TCMT IO) (Bool, QName)
forall {a}. MaybeT (TCMT IO) a
forall (m :: * -> *) a. MonadPlus m => m a
mzero
            PatternSynResName (AbstractName
d :| []) -> (Bool, QName) -> MaybeT (TCMT IO) (Bool, QName)
forall a. a -> MaybeT (TCMT IO) a
forall (m :: * -> *) a. Monad m => a -> m a
return ((Bool, QName) -> MaybeT (TCMT IO) (Bool, QName))
-> (QName -> (Bool, QName))
-> QName
-> MaybeT (TCMT IO) (Bool, QName)
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (Bool
True,) (QName -> MaybeT (TCMT IO) (Bool, QName))
-> QName -> MaybeT (TCMT IO) (Bool, QName)
forall a b. (a -> b) -> a -> b
$ AbstractName -> QName
anameName AbstractName
d
            PatternSynResName List1 AbstractName
ds        -> [Char] -> MaybeT (TCMT IO) (Bool, QName)
forall a. [Char] -> MaybeT (TCMT IO) a
failure ([Char] -> MaybeT (TCMT IO) (Bool, QName))
-> [Char] -> MaybeT (TCMT IO) (Bool, QName)
forall a b. (a -> b) -> a -> b
$ [Char]
"Ambiguous pattern synonym" [Char] -> [Char] -> [Char]
forall a. [a] -> [a] -> [a]
++ QName -> [Char]
forall a. Pretty a => a -> [Char]
prettyShow QName
top [Char] -> [Char] -> [Char]
forall a. [a] -> [a] -> [a]
++ [Char]
": " [Char] -> [Char] -> [Char]
forall a. [a] -> [a] -> [a]
++ List1 QName -> [Char]
forall a. Pretty a => a -> [Char]
prettyShow ((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
ds)

        lhs <- liftTCM $ toAbstract $ LeftHandSide top lhs YesDisplayLHS
        ps  <- case lhs of
                 A.LHS LHSInfo
_ (A.LHSHead QName
_ NAPs Expr
ps) -> NAPs Expr -> MaybeT (TCMT IO) (NAPs Expr)
forall a. a -> MaybeT (TCMT IO) a
forall (m :: * -> *) a. Monad m => a -> m a
return NAPs Expr
ps
                 LHS
_ -> MaybeT (TCMT IO) (NAPs Expr)
forall {a}. MaybeT (TCMT IO) a
err

        -- Andreas, 2016-08-08, issue #2132
        -- Remove pattern synonyms on lhs
        (hd, ps) <- do
          p <- liftTCM $ expandPatternSynonyms $
            (if isPatSyn then A.PatternSynP else A.DefP) (PatRange $ getRange lhs) (unambiguous hd) ps
          case p of
            A.DefP PatInfo
_ AmbiguousQName
f NAPs Expr
ps | Just QName
hd <- AmbiguousQName -> Maybe QName
getUnambiguous AmbiguousQName
f -> (QName, NAPs Expr) -> MaybeT (TCMT IO) (QName, NAPs Expr)
forall a. a -> MaybeT (TCMT IO) a
forall (m :: * -> *) a. Monad m => a -> m a
return (QName
hd, NAPs Expr
ps)
            A.ConP ConPatInfo
_ AmbiguousQName
c NAPs Expr
ps | Just QName
hd <- AmbiguousQName -> Maybe QName
getUnambiguous AmbiguousQName
c -> (QName, NAPs Expr) -> MaybeT (TCMT IO) (QName, NAPs Expr)
forall a. a -> MaybeT (TCMT IO) a
forall (m :: * -> *) a. Monad m => a -> m a
return (QName
hd, NAPs Expr
ps)
            A.PatternSynP{} -> MaybeT (TCMT IO) (QName, NAPs Expr)
forall a. HasCallStack => a
__IMPOSSIBLE__
            Pattern
_ -> MaybeT (TCMT IO) (QName, NAPs Expr)
forall {a}. MaybeT (TCMT IO) a
err

        rhs <- liftTCM $ toAbstract rhs

        -- Andreas, 2024-10-06, issue #7533:
        -- Check that all pattern variables occur on the rhs.
        -- Otherwise, there might be a misunderstanding of what display forms do.
        let used = Expr -> Set Name
allUsedNames Expr
rhs
        List1.unlessNull (filter (not . (isNoName || (`Set.member` used))) $ patternVars ps) $
          warning . UnusedVariablesInDisplayForm

        return $ A.DisplayPragma hd ps rhs
    where
      failure :: forall a. String -> MaybeT ScopeM a
      failure :: forall a. [Char] -> MaybeT (TCMT IO) a
failure [Char]
msg = do Warning -> MaybeT (TCMT IO) ()
forall (m :: * -> *).
(HasCallStack, MonadWarning m) =>
Warning -> m ()
warning (Range -> Doc -> Warning
UselessPragma (Pragma -> Range
forall a. HasRange a => a -> Range
getRange Pragma
pragma) (Doc -> Warning) -> Doc -> Warning
forall a b. (a -> b) -> a -> b
$ [Char] -> Doc
P.fwords [Char]
msg); MaybeT (TCMT IO) a
forall {a}. MaybeT (TCMT IO) a
forall (m :: * -> *) a. MonadPlus m => m a
mzero

  -- A warning attached to an ambiguous name shall apply to all disambiguations.
  toAbstract pragma :: Pragma
pragma@(C.WarningOnUsage Range
_ QName
x Text
str) = do
    ys <- QName -> ScopeM ResolvedName
resolveName QName
x ScopeM ResolvedName
-> (ResolvedName -> TCMT IO [AbstractName])
-> TCMT IO [AbstractName]
forall a b. TCMT IO a -> (a -> TCMT IO b) -> TCMT IO b
forall (m :: * -> *) a b. Monad m => m a -> (a -> m b) -> m b
>>= \case
      ConstructorName Set1 Induction
_ List1 AbstractName
ds     -> [AbstractName] -> TCMT IO [AbstractName]
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return ([AbstractName] -> TCMT IO [AbstractName])
-> [AbstractName] -> TCMT IO [AbstractName]
forall a b. (a -> b) -> a -> b
$ List1 AbstractName -> [Item (List1 AbstractName)]
forall l. IsList l => l -> [Item l]
List1.toList List1 AbstractName
ds
      FieldName List1 AbstractName
ds             -> [AbstractName] -> TCMT IO [AbstractName]
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return ([AbstractName] -> TCMT IO [AbstractName])
-> [AbstractName] -> TCMT IO [AbstractName]
forall a b. (a -> b) -> a -> b
$ List1 AbstractName -> [Item (List1 AbstractName)]
forall l. IsList l => l -> [Item l]
List1.toList List1 AbstractName
ds
      PatternSynResName List1 AbstractName
ds     -> [AbstractName] -> TCMT IO [AbstractName]
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return ([AbstractName] -> TCMT IO [AbstractName])
-> [AbstractName] -> TCMT IO [AbstractName]
forall a b. (a -> b) -> a -> b
$ List1 AbstractName -> [Item (List1 AbstractName)]
forall l. IsList l => l -> [Item l]
List1.toList List1 AbstractName
ds
      DefinedName Access
_ AbstractName
d Suffix
NoSuffix -> [AbstractName] -> TCMT IO [AbstractName]
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return ([AbstractName] -> TCMT IO [AbstractName])
-> [AbstractName] -> TCMT IO [AbstractName]
forall a b. (a -> b) -> a -> b
$ AbstractName -> [AbstractName]
forall el coll. Singleton el coll => el -> coll
singleton AbstractName
d
      DefinedName Access
_ AbstractName
d Suffix{} -> [] [AbstractName] -> TCMT IO () -> TCMT IO [AbstractName]
forall a b. a -> TCMT IO b -> TCMT IO a
forall (f :: * -> *) a b. Functor f => a -> f b -> f a
<$ QName -> TCMT IO ()
notInScopeWarning QName
x
      ResolvedName
UnknownName              -> [] [AbstractName] -> TCMT IO () -> TCMT IO [AbstractName]
forall a b. a -> TCMT IO b -> TCMT IO a
forall (f :: * -> *) a b. Functor f => a -> f b -> f a
<$ QName -> TCMT IO ()
notInScopeWarning QName
x
      VarName Name
x BindingSource
_              -> [] [AbstractName] -> TCMT IO [Any] -> TCMT IO [AbstractName]
forall a b. a -> TCMT IO b -> TCMT IO a
forall (f :: * -> *) a b. Functor f => a -> f b -> f a
<$ do
        Pragma -> [Char] -> TCMT IO [Any]
forall p a. HasRange p => p -> [Char] -> ScopeM [a]
uselessPragma Pragma
pragma ([Char] -> TCMT IO [Any]) -> [Char] -> TCMT IO [Any]
forall a b. (a -> b) -> a -> b
$ [Char]
"Not a defined name: " [Char] -> [Char] -> [Char]
forall a. [a] -> [a] -> [a]
++ Name -> [Char]
forall a. Pretty a => a -> [Char]
prettyShow Name
x
    forM_ ys $ \ AbstractName
y -> (UserWarnings -> f UserWarnings) -> TCState -> f TCState
Lens' TCState UserWarnings
stLocalUserWarnings Lens' TCState UserWarnings
-> (UserWarnings -> UserWarnings) -> TCMT IO ()
forall (m :: * -> *) a.
MonadTCState m =>
Lens' TCState a -> (a -> a) -> m ()
`modifyTCLens` QName -> Text -> UserWarnings -> UserWarnings
forall k a. Ord k => k -> a -> Map k a -> Map k a
Map.insert (AbstractName -> QName
anameName AbstractName
y) Text
str
    return []

  toAbstract (C.WarningOnImport Range
_ Text
str) = do
    (Maybe Text -> f (Maybe Text)) -> TCState -> f TCState
Lens' TCState (Maybe Text)
stWarningOnImport Lens' TCState (Maybe Text) -> Maybe Text -> TCMT IO ()
forall (m :: * -> *) a.
MonadTCState m =>
Lens' TCState a -> a -> m ()
`setTCLens` Text -> Maybe Text
forall a. a -> Maybe a
Just Text
str
    [Pragma] -> TCMT IO [Pragma]
forall a. a -> TCMT IO a
forall (f :: * -> *) a. Applicative f => a -> f a
pure []

  -- Termination, Coverage, Positivity, Universe, and Catchall
  -- pragmes are handled by the nicifier
  toAbstract C.TerminationCheckPragma{}  = TCMT IO [Pragma]
ScopeM (AbsOfCon Pragma)
forall a. HasCallStack => a
__IMPOSSIBLE__
  toAbstract C.NoCoverageCheckPragma{}   = TCMT IO [Pragma]
ScopeM (AbsOfCon Pragma)
forall a. HasCallStack => a
__IMPOSSIBLE__
  toAbstract C.NoPositivityCheckPragma{} = TCMT IO [Pragma]
ScopeM (AbsOfCon Pragma)
forall a. HasCallStack => a
__IMPOSSIBLE__
  toAbstract C.NoUniverseCheckPragma{}   = TCMT IO [Pragma]
ScopeM (AbsOfCon Pragma)
forall a. HasCallStack => a
__IMPOSSIBLE__
  toAbstract C.CatchallPragma{}          = TCMT IO [Pragma]
ScopeM (AbsOfCon Pragma)
forall a. HasCallStack => a
__IMPOSSIBLE__

  -- Polarity pragmas are handled by the niceifier.
  toAbstract C.PolarityPragma{} = TCMT IO [Pragma]
ScopeM (AbsOfCon Pragma)
forall a. HasCallStack => a
__IMPOSSIBLE__

uselessPragma :: HasRange p => p -> String -> ScopeM [a]
uselessPragma :: forall p a. HasRange p => p -> [Char] -> ScopeM [a]
uselessPragma p
pragma = ([] [a] -> TCMT IO () -> TCMT IO [a]
forall a b. a -> TCMT IO b -> TCMT IO a
forall (f :: * -> *) a b. Functor f => a -> f b -> f a
<$) (TCMT IO () -> TCMT IO [a])
-> ([Char] -> TCMT IO ()) -> [Char] -> TCMT IO [a]
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Warning -> TCMT IO ()
forall (m :: * -> *).
(HasCallStack, MonadWarning m) =>
Warning -> m ()
warning (Warning -> TCMT IO ())
-> ([Char] -> Warning) -> [Char] -> TCMT IO ()
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Range -> Doc -> Warning
UselessPragma (p -> Range
forall a. HasRange a => a -> Range
getRange p
pragma) (Doc -> Warning) -> ([Char] -> Doc) -> [Char] -> Warning
forall b c a. (b -> c) -> (a -> b) -> a -> c
. [Char] -> Doc
P.fwords

unambiguousConOrDef :: (C.QName -> IsAmbiguous -> Warning) -> C.QName -> ScopeM (Maybe A.QName)
unambiguousConOrDef :: (QName -> IsAmbiguous -> Warning) -> QName -> TCMT IO (Maybe QName)
unambiguousConOrDef QName -> IsAmbiguous -> Warning
warn QName
x = do
    TCMT IO (Maybe Expr)
-> TCMT IO (Maybe QName)
-> (Expr -> TCMT IO (Maybe QName))
-> TCMT IO (Maybe QName)
forall (m :: * -> *) a b.
Monad m =>
m (Maybe a) -> m b -> (a -> m b) -> m b
caseMaybeM (MaybeOldQName -> ScopeM (AbsOfCon MaybeOldQName)
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract (MaybeOldQName -> ScopeM (AbsOfCon MaybeOldQName))
-> MaybeOldQName -> ScopeM (AbsOfCon MaybeOldQName)
forall a b. (a -> b) -> a -> b
$ OldQName -> MaybeOldQName
MaybeOldQName (OldQName -> MaybeOldQName) -> OldQName -> MaybeOldQName
forall a b. (a -> b) -> a -> b
$ QName -> Maybe (Set1 Name) -> OldQName
OldQName QName
x Maybe (Set1 Name)
forall a. Maybe a
Nothing) TCMT IO (Maybe QName)
notInScope ((Expr -> TCMT IO (Maybe QName)) -> TCMT IO (Maybe QName))
-> (Expr -> TCMT IO (Maybe QName)) -> TCMT IO (Maybe QName)
forall a b. (a -> b) -> a -> b
$ \case
      A.Def' QName
y Suffix
NoSuffix              -> QName -> TCMT IO (Maybe QName)
forall {a}. a -> TCMT IO (Maybe a)
ret QName
y
      A.Def' QName
y Suffix{}              -> IsAmbiguous -> TCMT IO (Maybe QName)
failure IsAmbiguous
NotAmbiguous
      A.Proj ProjOrigin
_ AmbiguousQName
p
        | Just QName
y <- AmbiguousQName -> Maybe QName
getUnambiguous AmbiguousQName
p -> QName -> TCMT IO (Maybe QName)
forall {a}. a -> TCMT IO (Maybe a)
ret QName
y
        | Bool
otherwise                  -> IsAmbiguous -> TCMT IO (Maybe QName)
failure (IsAmbiguous -> TCMT IO (Maybe QName))
-> IsAmbiguous -> TCMT IO (Maybe QName)
forall a b. (a -> b) -> a -> b
$ AmbiguousQName -> IsAmbiguous
YesAmbiguous AmbiguousQName
p
      A.Con AmbiguousQName
c
        | Just QName
y <- AmbiguousQName -> Maybe QName
getUnambiguous AmbiguousQName
c -> QName -> TCMT IO (Maybe QName)
forall {a}. a -> TCMT IO (Maybe a)
ret QName
y
        | Bool
otherwise                  -> IsAmbiguous -> TCMT IO (Maybe QName)
failure (IsAmbiguous -> TCMT IO (Maybe QName))
-> IsAmbiguous -> TCMT IO (Maybe QName)
forall a b. (a -> b) -> a -> b
$ AmbiguousQName -> IsAmbiguous
YesAmbiguous AmbiguousQName
c
      A.Var{}                        -> IsAmbiguous -> TCMT IO (Maybe QName)
failure IsAmbiguous
NotAmbiguous
      A.PatternSyn{}                 -> IsAmbiguous -> TCMT IO (Maybe QName)
failure IsAmbiguous
NotAmbiguous
      Expr
_ -> TCMT IO (Maybe QName)
forall a. HasCallStack => a
__IMPOSSIBLE__
  where
    notInScope :: TCMT IO (Maybe QName)
notInScope = Maybe QName
forall a. Maybe a
Nothing Maybe QName -> TCMT IO () -> TCMT IO (Maybe QName)
forall a b. a -> TCMT IO b -> TCMT IO a
forall (f :: * -> *) a b. Functor f => a -> f b -> f a
<$ QName -> TCMT IO ()
notInScopeWarning QName
x
    failure :: IsAmbiguous -> TCMT IO (Maybe QName)
failure = (Maybe QName
forall a. Maybe a
Nothing Maybe QName -> TCMT IO () -> TCMT IO (Maybe QName)
forall a b. a -> TCMT IO b -> TCMT IO a
forall (f :: * -> *) a b. Functor f => a -> f b -> f a
<$) (TCMT IO () -> TCMT IO (Maybe QName))
-> (IsAmbiguous -> TCMT IO ())
-> IsAmbiguous
-> TCMT IO (Maybe QName)
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Warning -> TCMT IO ()
forall (m :: * -> *).
(HasCallStack, MonadWarning m) =>
Warning -> m ()
warning (Warning -> TCMT IO ())
-> (IsAmbiguous -> Warning) -> IsAmbiguous -> TCMT IO ()
forall b c a. (b -> c) -> (a -> b) -> a -> c
. QName -> IsAmbiguous -> Warning
warn QName
x
    ret :: a -> TCMT IO (Maybe a)
ret = Maybe a -> TCMT IO (Maybe a)
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return (Maybe a -> TCMT IO (Maybe a))
-> (a -> Maybe a) -> a -> TCMT IO (Maybe a)
forall b c a. (b -> c) -> (a -> b) -> a -> c
. a -> Maybe a
forall a. a -> Maybe a
Just

unambiguousDef :: (C.QName -> IsAmbiguous -> Warning) -> C.QName -> ScopeM (Maybe A.QName)
unambiguousDef :: (QName -> IsAmbiguous -> Warning) -> QName -> TCMT IO (Maybe QName)
unambiguousDef QName -> IsAmbiguous -> Warning
warn QName
x = do
    TCMT IO (Maybe Expr)
-> TCMT IO (Maybe QName)
-> (Expr -> TCMT IO (Maybe QName))
-> TCMT IO (Maybe QName)
forall (m :: * -> *) a b.
Monad m =>
m (Maybe a) -> m b -> (a -> m b) -> m b
caseMaybeM (MaybeOldQName -> ScopeM (AbsOfCon MaybeOldQName)
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract (MaybeOldQName -> ScopeM (AbsOfCon MaybeOldQName))
-> MaybeOldQName -> ScopeM (AbsOfCon MaybeOldQName)
forall a b. (a -> b) -> a -> b
$ OldQName -> MaybeOldQName
MaybeOldQName (OldQName -> MaybeOldQName) -> OldQName -> MaybeOldQName
forall a b. (a -> b) -> a -> b
$ QName -> Maybe (Set1 Name) -> OldQName
OldQName QName
x Maybe (Set1 Name)
forall a. Maybe a
Nothing) TCMT IO (Maybe QName)
notInScope ((Expr -> TCMT IO (Maybe QName)) -> TCMT IO (Maybe QName))
-> (Expr -> TCMT IO (Maybe QName)) -> TCMT IO (Maybe QName)
forall a b. (a -> b) -> a -> b
$ \case
      A.Def' QName
y Suffix
NoSuffix              -> QName -> TCMT IO (Maybe QName)
forall {a}. a -> TCMT IO (Maybe a)
ret QName
y
      A.Def' QName
y Suffix{}              -> IsAmbiguous -> TCMT IO (Maybe QName)
failure IsAmbiguous
NotAmbiguous
      A.Proj ProjOrigin
_ AmbiguousQName
p
        | Just QName
y <- AmbiguousQName -> Maybe QName
getUnambiguous AmbiguousQName
p -> QName -> TCMT IO (Maybe QName)
forall {a}. a -> TCMT IO (Maybe a)
ret QName
y
        | Bool
otherwise                  -> IsAmbiguous -> TCMT IO (Maybe QName)
failure (IsAmbiguous -> TCMT IO (Maybe QName))
-> IsAmbiguous -> TCMT IO (Maybe QName)
forall a b. (a -> b) -> a -> b
$ AmbiguousQName -> IsAmbiguous
YesAmbiguous AmbiguousQName
p
      A.Con{}                        -> IsAmbiguous -> TCMT IO (Maybe QName)
failure IsAmbiguous
NotAmbiguous
      A.Var{}                        -> IsAmbiguous -> TCMT IO (Maybe QName)
failure IsAmbiguous
NotAmbiguous
      A.PatternSyn{}                 -> IsAmbiguous -> TCMT IO (Maybe QName)
failure IsAmbiguous
NotAmbiguous
      Expr
_ -> TCMT IO (Maybe QName)
forall a. HasCallStack => a
__IMPOSSIBLE__
  where
    notInScope :: TCMT IO (Maybe QName)
notInScope = Maybe QName
forall a. Maybe a
Nothing Maybe QName -> TCMT IO () -> TCMT IO (Maybe QName)
forall a b. a -> TCMT IO b -> TCMT IO a
forall (f :: * -> *) a b. Functor f => a -> f b -> f a
<$ QName -> TCMT IO ()
notInScopeWarning QName
x
    failure :: IsAmbiguous -> TCMT IO (Maybe QName)
failure = (Maybe QName
forall a. Maybe a
Nothing Maybe QName -> TCMT IO () -> TCMT IO (Maybe QName)
forall a b. a -> TCMT IO b -> TCMT IO a
forall (f :: * -> *) a b. Functor f => a -> f b -> f a
<$) (TCMT IO () -> TCMT IO (Maybe QName))
-> (IsAmbiguous -> TCMT IO ())
-> IsAmbiguous
-> TCMT IO (Maybe QName)
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Warning -> TCMT IO ()
forall (m :: * -> *).
(HasCallStack, MonadWarning m) =>
Warning -> m ()
warning (Warning -> TCMT IO ())
-> (IsAmbiguous -> Warning) -> IsAmbiguous -> TCMT IO ()
forall b c a. (b -> c) -> (a -> b) -> a -> c
. QName -> IsAmbiguous -> Warning
warn QName
x
    ret :: a -> TCMT IO (Maybe a)
ret = Maybe a -> TCMT IO (Maybe a)
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return (Maybe a -> TCMT IO (Maybe a))
-> (a -> Maybe a) -> a -> TCMT IO (Maybe a)
forall b c a. (b -> c) -> (a -> b) -> a -> c
. a -> Maybe a
forall a. a -> Maybe a
Just

scopeCheckDef :: (C.QName -> Warning) -> C.QName -> ScopeM (Maybe A.QName)
scopeCheckDef :: (QName -> Warning) -> QName -> TCMT IO (Maybe QName)
scopeCheckDef QName -> Warning
warn QName
x = do
    TCMT IO (Maybe Expr)
-> TCMT IO (Maybe QName)
-> (Expr -> TCMT IO (Maybe QName))
-> TCMT IO (Maybe QName)
forall (m :: * -> *) a b.
Monad m =>
m (Maybe a) -> m b -> (a -> m b) -> m b
caseMaybeM (MaybeOldQName -> ScopeM (AbsOfCon MaybeOldQName)
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract (MaybeOldQName -> ScopeM (AbsOfCon MaybeOldQName))
-> MaybeOldQName -> ScopeM (AbsOfCon MaybeOldQName)
forall a b. (a -> b) -> a -> b
$ OldQName -> MaybeOldQName
MaybeOldQName (OldQName -> MaybeOldQName) -> OldQName -> MaybeOldQName
forall a b. (a -> b) -> a -> b
$ QName -> Maybe (Set1 Name) -> OldQName
OldQName QName
x Maybe (Set1 Name)
forall a. Maybe a
Nothing) TCMT IO (Maybe QName)
notInScope ((Expr -> TCMT IO (Maybe QName)) -> TCMT IO (Maybe QName))
-> (Expr -> TCMT IO (Maybe QName)) -> TCMT IO (Maybe QName)
forall a b. (a -> b) -> a -> b
$ \case
      A.Def' QName
y Suffix
NoSuffix -> QName -> TCMT IO (Maybe QName)
forall {a}. a -> TCMT IO (Maybe a)
ret QName
y
      A.Def' QName
y Suffix{} -> TCMT IO (Maybe QName)
failure
      A.Proj{}          -> TCMT IO (Maybe QName)
failure
      A.Con{}           -> TCMT IO (Maybe QName)
failure
      A.Var{}           -> TCMT IO (Maybe QName)
failure
      A.PatternSyn{}    -> TCMT IO (Maybe QName)
failure
      Expr
_ -> TCMT IO (Maybe QName)
forall a. HasCallStack => a
__IMPOSSIBLE__
  where
    notInScope :: TCMT IO (Maybe QName)
notInScope = Maybe QName
forall a. Maybe a
Nothing Maybe QName -> TCMT IO () -> TCMT IO (Maybe QName)
forall a b. a -> TCMT IO b -> TCMT IO a
forall (f :: * -> *) a b. Functor f => a -> f b -> f a
<$ QName -> TCMT IO ()
notInScopeWarning QName
x
    failure :: TCMT IO (Maybe QName)
failure = Maybe QName
forall a. Maybe a
Nothing Maybe QName -> TCMT IO () -> TCMT IO (Maybe QName)
forall a b. a -> TCMT IO b -> TCMT IO a
forall (f :: * -> *) a b. Functor f => a -> f b -> f a
<$ do Warning -> TCMT IO ()
forall (m :: * -> *).
(HasCallStack, MonadWarning m) =>
Warning -> m ()
warning (Warning -> TCMT IO ()) -> Warning -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ QName -> Warning
warn QName
x
    ret :: a -> TCMT IO (Maybe a)
ret = Maybe a -> TCMT IO (Maybe a)
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return (Maybe a -> TCMT IO (Maybe a))
-> (a -> Maybe a) -> a -> TCMT IO (Maybe a)
forall b c a. (b -> c) -> (a -> b) -> a -> c
. a -> Maybe a
forall a. a -> Maybe a
Just

instance ToAbstract C.Clause where
  type AbsOfCon C.Clause = A.Clause

  toAbstract :: Clause -> ScopeM (AbsOfCon Clause)
toAbstract (C.Clause Name
top Bool
catchall lhs :: LHS
lhs@(C.LHS Pattern
p [RewriteEqn]
eqs [WithExpr]
with) RHS' Expr
rhs WhereClause' [Declaration]
wh [Clause]
wcs) = ScopeM (AbsOfCon Clause) -> ScopeM (AbsOfCon Clause)
forall a. ScopeM a -> ScopeM a
withLocalVars (ScopeM (AbsOfCon Clause) -> ScopeM (AbsOfCon Clause))
-> ScopeM (AbsOfCon Clause) -> ScopeM (AbsOfCon Clause)
forall a b. (a -> b) -> a -> b
$ do
    -- Jesper, 2018-12-10, #3095: pattern variables bound outside the
    -- module are locally treated as module parameters
    (ScopeInfo -> ScopeInfo) -> TCMT IO ()
forall (m :: * -> *).
MonadTCState m =>
(ScopeInfo -> ScopeInfo) -> m ()
modifyScope_ ((ScopeInfo -> ScopeInfo) -> TCMT IO ())
-> (ScopeInfo -> ScopeInfo) -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ ([(Name, LocalVar)] -> [(Name, LocalVar)])
-> ScopeInfo -> ScopeInfo
updateScopeLocals (([(Name, LocalVar)] -> [(Name, LocalVar)])
 -> ScopeInfo -> ScopeInfo)
-> ([(Name, LocalVar)] -> [(Name, LocalVar)])
-> ScopeInfo
-> ScopeInfo
forall a b. (a -> b) -> a -> b
$ ((Name, LocalVar) -> (Name, LocalVar))
-> [(Name, LocalVar)] -> [(Name, LocalVar)]
forall a b. (a -> b) -> [a] -> [b]
map (((Name, LocalVar) -> (Name, LocalVar))
 -> [(Name, LocalVar)] -> [(Name, LocalVar)])
-> ((Name, LocalVar) -> (Name, LocalVar))
-> [(Name, LocalVar)]
-> [(Name, LocalVar)]
forall a b. (a -> b) -> a -> b
$ (LocalVar -> LocalVar) -> (Name, LocalVar) -> (Name, LocalVar)
forall b c a. (b -> c) -> (a, b) -> (a, c)
forall (p :: * -> * -> *) b c a.
Bifunctor p =>
(b -> c) -> p a b -> p a c
second LocalVar -> LocalVar
patternToModuleBound
    -- Andreas, 2012-02-14: need to reset local vars before checking subclauses
    vars0 <- TCMT IO [(Name, LocalVar)]
forall (m :: * -> *). ReadTCState m => m [(Name, LocalVar)]
getLocalVars
    lhs' <- toAbstract $ LeftHandSide (C.QName top) p NoDisplayLHS
    printLocals 30 "after lhs:"
    vars1 <- getLocalVars
    eqs <- mapM (toAbstractCtx TopCtx) eqs
    vars2 <- getLocalVars
    let vars = Int -> [(Name, LocalVar)] -> [(Name, LocalVar)]
forall a. Int -> [a] -> [a]
dropEnd ([(Name, LocalVar)] -> Int
forall a. [a] -> Int
forall (t :: * -> *) a. Foldable t => t a -> Int
length [(Name, LocalVar)]
vars1) [(Name, LocalVar)]
vars2 [(Name, LocalVar)] -> [(Name, LocalVar)] -> [(Name, LocalVar)]
forall a. [a] -> [a] -> [a]
++ [(Name, LocalVar)]
vars0
    let wcs' = ([(Name, LocalVar)]
vars, [Clause]
wcs)

    -- Handle rewrite equations first.
    if not (null eqs)
      then do
        rhs <- toAbstractCtx TopCtx $ RightHandSide eqs with wcs' rhs wh
        rhs <- toAbstract rhs
        return $ A.Clause lhs' [] rhs A.noWhereDecls catchall
      else do
        -- the right hand side is checked with the module of the local definitions opened
        (rhs, ds) <- whereToAbstract (getRange wh) wh $
                       toAbstractCtx TopCtx $ RightHandSide [] with wcs' rhs NoWhere
        rhs <- toAbstract rhs
        return $ A.Clause lhs' [] rhs ds catchall


whereToAbstract
  :: Range                            -- ^ The range of the @where@ block.
  -> C.WhereClause                    -- ^ The @where@ block.
  -> ScopeM a                         -- ^ The scope-checking task to be run in the context of the @where@ module.
  -> ScopeM (a, A.WhereDeclarations)  -- ^ Additionally return the scope-checked contents of the @where@ module.
whereToAbstract :: forall a.
Range
-> WhereClause' [Declaration]
-> ScopeM a
-> ScopeM (a, WhereDeclarations)
whereToAbstract Range
r WhereClause' [Declaration]
wh ScopeM a
inner = do
  case WhereClause' [Declaration]
wh of
    WhereClause' [Declaration]
NoWhere       -> ScopeM (a, WhereDeclarations)
ret
    AnyWhere Range
_ [] -> ScopeM (a, WhereDeclarations)
warnEmptyWhere
    AnyWhere Range
_ [Declaration]
ds -> ScopeM (a, WhereDeclarations) -> ScopeM (a, WhereDeclarations)
forall a. ScopeM a -> ScopeM a
enter (ScopeM (a, WhereDeclarations) -> ScopeM (a, WhereDeclarations))
-> ScopeM (a, WhereDeclarations) -> ScopeM (a, WhereDeclarations)
forall a b. (a -> b) -> a -> b
$ do
      -- Andreas, 2016-07-17 issues #2081 and #2101
      -- where-declarations are automatically private.
      -- This allows their type signature to be checked InAbstractMode.
      Range
-> Erased
-> Maybe (Name, Access)
-> List1 Declaration
-> ScopeM a
-> ScopeM (a, WhereDeclarations)
forall a.
Range
-> Erased
-> Maybe (Name, Access)
-> List1 Declaration
-> ScopeM a
-> ScopeM (a, WhereDeclarations)
whereToAbstract1 Range
r Erased
defaultErased Maybe (Name, Access)
forall a. Maybe a
Nothing
        (Declaration -> List1 Declaration
forall el coll. Singleton el coll => el -> coll
singleton (Declaration -> List1 Declaration)
-> Declaration -> List1 Declaration
forall a b. (a -> b) -> a -> b
$ KwRange -> Origin -> [Declaration] -> Declaration
C.Private KwRange
forall a. Null a => a
empty Origin
Inserted [Declaration]
ds) ScopeM a
inner
    SomeWhere Range
_ Erased
e Name
m Access
a [Declaration]
ds0 -> ScopeM (a, WhereDeclarations) -> ScopeM (a, WhereDeclarations)
forall a. ScopeM a -> ScopeM a
enter (ScopeM (a, WhereDeclarations) -> ScopeM (a, WhereDeclarations))
-> ScopeM (a, WhereDeclarations) -> ScopeM (a, WhereDeclarations)
forall a b. (a -> b) -> a -> b
$
      [Declaration]
-> ScopeM (a, WhereDeclarations)
-> (List1 Declaration -> ScopeM (a, WhereDeclarations))
-> ScopeM (a, WhereDeclarations)
forall a b. [a] -> b -> (List1 a -> b) -> b
List1.ifNull [Declaration]
ds0 ScopeM (a, WhereDeclarations)
warnEmptyWhere {-else-} ((List1 Declaration -> ScopeM (a, WhereDeclarations))
 -> ScopeM (a, WhereDeclarations))
-> (List1 Declaration -> ScopeM (a, WhereDeclarations))
-> ScopeM (a, WhereDeclarations)
forall a b. (a -> b) -> a -> b
$ \List1 Declaration
ds -> do
      -- Named where-modules do not default to private.
      Range
-> Erased
-> Maybe (Name, Access)
-> List1 Declaration
-> ScopeM a
-> ScopeM (a, WhereDeclarations)
forall a.
Range
-> Erased
-> Maybe (Name, Access)
-> List1 Declaration
-> ScopeM a
-> ScopeM (a, WhereDeclarations)
whereToAbstract1 Range
r Erased
e ((Name, Access) -> Maybe (Name, Access)
forall a. a -> Maybe a
Just (Name
m, Access
a)) List1 Declaration
ds ScopeM a
inner
  where
  enter :: TCMT IO a -> TCMT IO a
enter = (TCEnv -> TCEnv) -> TCMT IO a -> TCMT IO a
forall a. (TCEnv -> TCEnv) -> TCMT IO a -> TCMT IO a
forall (m :: * -> *) a.
MonadTCEnv m =>
(TCEnv -> TCEnv) -> m a -> m a
localTC ((TCEnv -> TCEnv) -> TCMT IO a -> TCMT IO a)
-> (TCEnv -> TCEnv) -> TCMT IO a -> TCMT IO a
forall a b. (a -> b) -> a -> b
$ \TCEnv
env -> TCEnv
env { envCheckingWhere = True }
  ret :: ScopeM (a, WhereDeclarations)
ret = (,WhereDeclarations
A.noWhereDecls) (a -> (a, WhereDeclarations))
-> ScopeM a -> ScopeM (a, WhereDeclarations)
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> ScopeM a
inner
  warnEmptyWhere :: ScopeM (a, WhereDeclarations)
warnEmptyWhere = do
    Range -> TCMT IO () -> TCMT IO ()
forall (m :: * -> *) x a.
(MonadTrace m, HasRange x) =>
x -> m a -> m a
setCurrentRange Range
r (TCMT IO () -> TCMT IO ()) -> TCMT IO () -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ Warning -> TCMT IO ()
forall (m :: * -> *).
(HasCallStack, MonadWarning m) =>
Warning -> m ()
warning Warning
EmptyWhere
    ScopeM (a, WhereDeclarations)
ret

whereToAbstract1
  :: Range                            -- ^ The range of the @where@-block.
  -> Erased                           -- ^ Is the where module erased?
  -> Maybe (C.Name, Access)           -- ^ The name of the @where@ module (if any).
  -> List1 C.Declaration              -- ^ The contents of the @where@ module.
  -> ScopeM a                         -- ^ The scope-checking task to be run in the context of the @where@ module.
  -> ScopeM (a, A.WhereDeclarations)  -- ^ Additionally return the scope-checked contents of the @where@ module.
whereToAbstract1 :: forall a.
Range
-> Erased
-> Maybe (Name, Access)
-> List1 Declaration
-> ScopeM a
-> ScopeM (a, WhereDeclarations)
whereToAbstract1 Range
r Erased
e Maybe (Name, Access)
whname List1 Declaration
whds ScopeM a
inner = do
  -- ASR (16 November 2015) Issue 1137: We ban termination
  -- pragmas inside `where` clause.
  WhereOrRecord -> List1 Declaration -> TCMT IO ()
forall a. FoldDecl a => WhereOrRecord -> a -> TCMT IO ()
checkNoTerminationPragma WhereOrRecord
InWhereBlock List1 Declaration
whds

  -- Create a fresh concrete name if there isn't (a proper) one.
  (m, acc) <- do
    case Maybe (Name, Access)
whname of
      Just (Name
m, Access
acc) | Bool -> Bool
not (Name -> Bool
forall a. IsNoName a => a -> Bool
isNoName Name
m) -> (Name, Access) -> TCMT IO (Name, Access)
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return (Name
m, Access
acc)
      Maybe (Name, Access)
_ -> TCMT IO NameId
forall i (m :: * -> *). MonadFresh i m => m i
fresh TCMT IO NameId
-> (NameId -> (Name, Access)) -> TCMT IO (Name, Access)
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \ NameId
x -> (Range -> NameId -> Name
C.NoName (Maybe (Name, Access) -> Range
forall a. HasRange a => a -> Range
getRange Maybe (Name, Access)
whname) NameId
x, Access
privateAccessInserted)
           -- unnamed where's are private
  old <- getCurrentModule
  am  <- toAbstract (NewModuleName m)
  (scope, d) <- scopeCheckModule r e (C.QName m) am [] $
                toAbstract $ Declarations $ List1.toList whds
  setScope scope
  x <- inner
  setCurrentModule old
  bindModule acc m am
  -- Issue 848: if the module was anonymous (module _ where) open it public
  let anonymousSomeWhere = Bool -> ((Name, Access) -> Bool) -> Maybe (Name, Access) -> Bool
forall b a. b -> (a -> b) -> Maybe a -> b
maybe Bool
False (Name -> Bool
forall a. IsNoName a => a -> Bool
isNoName (Name -> Bool)
-> ((Name, Access) -> Name) -> (Name, Access) -> Bool
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (Name, Access) -> Name
forall a b. (a, b) -> a
fst) Maybe (Name, Access)
whname
  when anonymousSomeWhere $
   void $ -- We can ignore the returned default A.ImportDirective.
    openModule TopOpenModule (Just am) (C.QName m) $
      defaultImportDir { publicOpen = Just empty }
  return (x, A.WhereDecls (Just am) (isNothing whname) $ singleton d)

data TerminationOrPositivity = Termination | Positivity
  deriving (Int -> TerminationOrPositivity -> [Char] -> [Char]
[TerminationOrPositivity] -> [Char] -> [Char]
TerminationOrPositivity -> [Char]
(Int -> TerminationOrPositivity -> [Char] -> [Char])
-> (TerminationOrPositivity -> [Char])
-> ([TerminationOrPositivity] -> [Char] -> [Char])
-> Show TerminationOrPositivity
forall a.
(Int -> a -> [Char] -> [Char])
-> (a -> [Char]) -> ([a] -> [Char] -> [Char]) -> Show a
$cshowsPrec :: Int -> TerminationOrPositivity -> [Char] -> [Char]
showsPrec :: Int -> TerminationOrPositivity -> [Char] -> [Char]
$cshow :: TerminationOrPositivity -> [Char]
show :: TerminationOrPositivity -> [Char]
$cshowList :: [TerminationOrPositivity] -> [Char] -> [Char]
showList :: [TerminationOrPositivity] -> [Char] -> [Char]
Show)

data WhereOrRecord = InWhereBlock | InRecordDef

checkNoTerminationPragma :: FoldDecl a => WhereOrRecord -> a -> ScopeM ()
checkNoTerminationPragma :: forall a. FoldDecl a => WhereOrRecord -> a -> TCMT IO ()
checkNoTerminationPragma WhereOrRecord
b a
ds =
  -- foldDecl traverses into all sub-declarations.
  [(TerminationOrPositivity, Range)]
-> ((TerminationOrPositivity, Range) -> TCMT IO ()) -> TCMT IO ()
forall (t :: * -> *) (m :: * -> *) a b.
(Foldable t, Monad m) =>
t a -> (a -> m b) -> m ()
forM_ ((Declaration -> [(TerminationOrPositivity, Range)])
-> a -> [(TerminationOrPositivity, Range)]
forall m. Monoid m => (Declaration -> m) -> a -> m
forall a m. (FoldDecl a, Monoid m) => (Declaration -> m) -> a -> m
foldDecl (Declaration -> [Pragma]
forall m. CMaybe Pragma m => Declaration -> m
isPragma (Declaration -> [Pragma])
-> (Pragma -> [(TerminationOrPositivity, Range)])
-> Declaration
-> [(TerminationOrPositivity, Range)]
forall (m :: * -> *) a b c.
Monad m =>
(a -> m b) -> (b -> m c) -> a -> m c
>=> Pragma -> [(TerminationOrPositivity, Range)]
isTerminationPragma) a
ds) \ (TerminationOrPositivity
p, Range
r) ->
    Range -> TCMT IO () -> TCMT IO ()
forall (m :: * -> *) x a.
(MonadTrace m, HasRange x) =>
x -> m a -> m a
setCurrentRange Range
r (TCMT IO () -> TCMT IO ()) -> TCMT IO () -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ Warning -> TCMT IO ()
forall (m :: * -> *).
(HasCallStack, MonadWarning m) =>
Warning -> m ()
warning (Warning -> TCMT IO ()) -> Warning -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ Range -> Doc -> Warning
UselessPragma Range
r (Doc -> Warning) -> Doc -> Warning
forall a b. (a -> b) -> a -> b
$ [Doc] -> Doc
forall (t :: * -> *). Foldable t => t Doc -> Doc
P.vcat
      [ [Char] -> Doc
forall a. [Char] -> Doc a
P.text ([Char] -> Doc) -> [Char] -> Doc
forall a b. (a -> b) -> a -> b
$ TerminationOrPositivity -> [Char]
forall a. Show a => a -> [Char]
show TerminationOrPositivity
p [Char] -> [Char] -> [Char]
forall a. [a] -> [a] -> [a]
++ [Char]
" pragmas are ignored in " [Char] -> [Char] -> [Char]
forall a. [a] -> [a] -> [a]
++ WhereOrRecord -> [Char]
forall {a}. IsString a => WhereOrRecord -> a
what WhereOrRecord
b
      , [Char] -> Doc
forall a. [Char] -> Doc a
P.text ([Char] -> Doc) -> [Char] -> Doc
forall a b. (a -> b) -> a -> b
$ [Char]
"(see " [Char] -> [Char] -> [Char]
forall a. [a] -> [a] -> [a]
++ WhereOrRecord -> [Char]
issue WhereOrRecord
b [Char] -> [Char] -> [Char]
forall a. [a] -> [a] -> [a]
++ [Char]
")"
      ]
  where
    what :: WhereOrRecord -> a
what WhereOrRecord
InWhereBlock = a
"where clauses"
    what WhereOrRecord
InRecordDef  = a
"record definitions"
    github :: a -> [Char]
github a
n = [Char]
"https://github.com/agda/agda/issues/" [Char] -> [Char] -> [Char]
forall a. [a] -> [a] -> [a]
++ a -> [Char]
forall a. Show a => a -> [Char]
show a
n
    issue :: WhereOrRecord -> [Char]
issue WhereOrRecord
InWhereBlock = Integer -> [Char]
forall a. Show a => a -> [Char]
github Integer
3355
    issue WhereOrRecord
InRecordDef  = Integer -> [Char]
forall a. Show a => a -> [Char]
github Integer
3008

    isTerminationPragma :: C.Pragma -> [(TerminationOrPositivity, Range)]
    isTerminationPragma :: Pragma -> [(TerminationOrPositivity, Range)]
isTerminationPragma = \case
      C.TerminationCheckPragma Range
r TerminationCheck
_  -> [(TerminationOrPositivity
Termination, Range
r)]
      C.NoPositivityCheckPragma Range
r   -> [(TerminationOrPositivity
Positivity, Range
r)]
      C.OptionsPragma Range
_ [[Char]]
_           -> []
      C.BuiltinPragma Range
_ RString
_ QName
_         -> []
      C.RewritePragma Range
_ Range
_ [QName]
_         -> []
      C.ForeignPragma Range
_ Ranged Text
_ [Char]
_         -> []
      C.CompilePragma Range
_ Ranged Text
_ QName
_ [Char]
_       -> []
      C.StaticPragma Range
_ QName
_            -> []
      C.InlinePragma Range
_ Bool
_ QName
_          -> []
      C.ImpossiblePragma Range
_ [[Char]]
_        -> []
      C.EtaPragma Range
_ QName
_               -> []
      C.WarningOnUsage Range
_ QName
_ Text
_        -> []
      C.WarningOnImport Range
_ Text
_         -> []
      C.InjectivePragma Range
_ QName
_         -> []
      C.InjectiveForInferencePragma{} -> []
      C.DisplayPragma Range
_ Pattern
_ Expr
_         -> []
      C.CatchallPragma Range
_            -> []
      C.NoCoverageCheckPragma Range
_     -> []
      C.PolarityPragma Range
_ Name
_ [Occurrence]
_        -> []
      C.NoUniverseCheckPragma Range
_     -> []
      C.NotProjectionLikePragma Range
_ QName
_ -> []
      C.OverlapPragma Range
_ [QName]
_ OverlapMode
_         -> []

data RightHandSide = RightHandSide
  { RightHandSide -> [RewriteEqn' () BindName Pattern Expr]
_rhsRewriteEqn :: [RewriteEqn' () A.BindName A.Pattern A.Expr]
    -- ^ @rewrite e | with p <- e in eq@ (many)
  , RightHandSide -> [WithExpr]
_rhsWithExpr   :: [C.WithExpr]
    -- ^ @with e@ (many)
  , RightHandSide -> ([(Name, LocalVar)], [Clause])
_rhsSubclauses :: (LocalVars, [C.Clause])
    -- ^ the subclauses spawned by a with (monadic because we need to reset the local vars before checking these clauses)
  , RightHandSide -> RHS' Expr
_rhs           :: C.RHS
  , RightHandSide -> WhereClause' [Declaration]
_rhsWhere      :: WhereClause
      -- ^ @where@ module.
  }

data AbstractRHS
  = AbsurdRHS'
  | WithRHS' (List1 A.WithExpr) (List1 (ScopeM C.Clause))
    -- ^ The with clauses haven't been translated yet
  | RHS' A.Expr C.Expr
  | RewriteRHS' [RewriteEqn' () A.BindName A.Pattern A.Expr] AbstractRHS A.WhereDeclarations

qualifyName_ :: A.Name -> ScopeM A.QName
qualifyName_ :: Name -> TCMT IO QName
qualifyName_ Name
x = do
  m <- TCMT IO ModuleName
forall (m :: * -> *). ReadTCState m => m ModuleName
getCurrentModule
  return $ A.qualify m x

withFunctionName :: String -> ScopeM A.QName
withFunctionName :: [Char] -> TCMT IO QName
withFunctionName [Char]
s = do
  NameId i _ <- TCMT IO NameId
forall i (m :: * -> *). MonadFresh i m => m i
fresh
  qualifyName_ =<< freshName_ (s ++ show i)

instance ToAbstract (RewriteEqn' () A.BindName A.Pattern A.Expr) where
  type AbsOfCon (RewriteEqn' () A.BindName A.Pattern A.Expr) = A.RewriteEqn
  toAbstract :: RewriteEqn' () BindName Pattern Expr
-> ScopeM (AbsOfCon (RewriteEqn' () BindName Pattern Expr))
toAbstract = \case
    Rewrite List1 ((), Expr)
es -> (List1 (QName, Expr)
 -> AbsOfCon (RewriteEqn' () BindName Pattern Expr))
-> TCMT IO (List1 (QName, Expr))
-> ScopeM (AbsOfCon (RewriteEqn' () BindName Pattern Expr))
forall a b. (a -> b) -> TCMT IO a -> TCMT IO b
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
fmap List1 (QName, Expr) -> RewriteEqn' QName BindName Pattern Expr
List1 (QName, Expr)
-> AbsOfCon (RewriteEqn' () BindName Pattern Expr)
forall qn nm p e. List1 (qn, e) -> RewriteEqn' qn nm p e
Rewrite (TCMT IO (List1 (QName, Expr))
 -> ScopeM (AbsOfCon (RewriteEqn' () BindName Pattern Expr)))
-> TCMT IO (List1 (QName, Expr))
-> ScopeM (AbsOfCon (RewriteEqn' () BindName Pattern Expr))
forall a b. (a -> b) -> a -> b
$ List1 ((), Expr)
-> (((), Expr) -> TCMT IO (QName, Expr))
-> TCMT IO (List1 (QName, Expr))
forall (t :: * -> *) (m :: * -> *) a b.
(Traversable t, Monad m) =>
t a -> (a -> m b) -> m (t b)
forM List1 ((), Expr)
es ((((), Expr) -> TCMT IO (QName, Expr))
 -> TCMT IO (List1 (QName, Expr)))
-> (((), Expr) -> TCMT IO (QName, Expr))
-> TCMT IO (List1 (QName, Expr))
forall a b. (a -> b) -> a -> b
$ \ (()
_, Expr
e) -> do
      qn <- [Char] -> TCMT IO QName
withFunctionName [Char]
"-rewrite"
      pure (qn, e)
    Invert ()
_ List1 (Named BindName (Pattern, Expr))
pes -> do
      qn <- [Char] -> TCMT IO QName
withFunctionName [Char]
"-invert"
      pure $ Invert qn pes
    LeftLet List1 (Pattern, Expr)
pes -> AbsOfCon (RewriteEqn' () BindName Pattern Expr)
-> ScopeM (AbsOfCon (RewriteEqn' () BindName Pattern Expr))
forall a. a -> TCMT IO a
forall (f :: * -> *) a. Applicative f => a -> f a
pure (AbsOfCon (RewriteEqn' () BindName Pattern Expr)
 -> ScopeM (AbsOfCon (RewriteEqn' () BindName Pattern Expr)))
-> AbsOfCon (RewriteEqn' () BindName Pattern Expr)
-> ScopeM (AbsOfCon (RewriteEqn' () BindName Pattern Expr))
forall a b. (a -> b) -> a -> b
$ List1 (Pattern, Expr) -> RewriteEqn' QName BindName Pattern Expr
forall qn nm p e. List1 (p, e) -> RewriteEqn' qn nm p e
LeftLet List1 (Pattern, Expr)
pes

instance ToAbstract C.RewriteEqn where
  type AbsOfCon C.RewriteEqn = RewriteEqn' () A.BindName A.Pattern A.Expr
  toAbstract :: RewriteEqn -> ScopeM (AbsOfCon RewriteEqn)
toAbstract = \case
    Rewrite List1 ((), Expr)
es   -> List1 ((), Expr) -> RewriteEqn' () BindName Pattern Expr
forall qn nm p e. List1 (qn, e) -> RewriteEqn' qn nm p e
Rewrite (List1 ((), Expr) -> RewriteEqn' () BindName Pattern Expr)
-> TCMT IO (List1 ((), Expr))
-> TCMT IO (RewriteEqn' () BindName Pattern Expr)
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> (((), Expr) -> TCMT IO ((), Expr))
-> List1 ((), Expr) -> TCMT IO (List1 ((), Expr))
forall (t :: * -> *) (m :: * -> *) a b.
(Traversable t, Monad m) =>
(a -> m b) -> t a -> m (t b)
forall (m :: * -> *) a b.
Monad m =>
(a -> m b) -> NonEmpty a -> m (NonEmpty b)
mapM ((), Expr) -> TCMT IO ((), Expr)
((), Expr) -> ScopeM (AbsOfCon ((), Expr))
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract List1 ((), Expr)
es
    Invert ()
_ List1 (Named Name (Pattern, Expr))
npes -> ()
-> List1 (Named BindName (Pattern, Expr))
-> RewriteEqn' () BindName Pattern Expr
forall qn nm p e.
qn -> List1 (Named nm (p, e)) -> RewriteEqn' qn nm p e
Invert () (List1 (Named BindName (Pattern, Expr))
 -> RewriteEqn' () BindName Pattern Expr)
-> TCMT IO (List1 (Named BindName (Pattern, Expr)))
-> TCMT IO (RewriteEqn' () BindName Pattern Expr)
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> do
      -- Given a list of irrefutable with expressions of the form @p <- e in q@
      let (NonEmpty (Maybe Name, Pattern)
nps, List1 Expr
es) = NonEmpty ((Maybe Name, Pattern), Expr)
-> (NonEmpty (Maybe Name, Pattern), List1 Expr)
forall (f :: * -> *) a b. Functor f => f (a, b) -> (f a, f b)
List1.unzip
                    (NonEmpty ((Maybe Name, Pattern), Expr)
 -> (NonEmpty (Maybe Name, Pattern), List1 Expr))
-> NonEmpty ((Maybe Name, Pattern), Expr)
-> (NonEmpty (Maybe Name, Pattern), List1 Expr)
forall a b. (a -> b) -> a -> b
$ (Named Name (Pattern, Expr) -> ((Maybe Name, Pattern), Expr))
-> List1 (Named Name (Pattern, Expr))
-> NonEmpty ((Maybe Name, Pattern), Expr)
forall a b. (a -> b) -> NonEmpty a -> NonEmpty b
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
fmap (\ (Named Maybe Name
nm (Pattern
p, Expr
e)) -> ((Maybe Name
nm, Pattern
p), Expr
e)) List1 (Named Name (Pattern, Expr))
npes
      -- we first check the expressions @e@: the patterns may shadow some of the
      -- variables mentioned in them!
      es <- List1 Expr -> ScopeM (AbsOfCon (List1 Expr))
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract List1 Expr
es
      -- we then parse the pairs of patterns @p@ and names @q@ for the equality
      -- constraints of the form @p ≡ e@.
      nps <- forM nps $ \ (Maybe Name
n, Pattern
p) -> do
        -- first the pattern
        p <- Pattern -> TCMT IO Pattern
toAbsPat Pattern
p
        -- and then the name
        n <- toAbstract $ fmap (NewName WithBound . C.mkBoundName_) n
        pure (n, p)
      -- we finally reassemble the telescope
      pure $ List1.zipWith (\ (Maybe BindName
n,Pattern
p) Expr
e -> Maybe BindName -> (Pattern, Expr) -> Named BindName (Pattern, Expr)
forall name a. Maybe name -> a -> Named name a
Named Maybe BindName
n (Pattern
p, Expr
e)) nps es
    LeftLet List1 (Pattern, Expr)
pes -> (List1 (Pattern, Expr) -> AbsOfCon RewriteEqn)
-> TCMT IO (List1 (Pattern, Expr)) -> ScopeM (AbsOfCon RewriteEqn)
forall a b. (a -> b) -> TCMT IO a -> TCMT IO b
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
fmap List1 (Pattern, Expr) -> RewriteEqn' () BindName Pattern Expr
List1 (Pattern, Expr) -> AbsOfCon RewriteEqn
forall qn nm p e. List1 (p, e) -> RewriteEqn' qn nm p e
LeftLet (TCMT IO (List1 (Pattern, Expr)) -> ScopeM (AbsOfCon RewriteEqn))
-> TCMT IO (List1 (Pattern, Expr)) -> ScopeM (AbsOfCon RewriteEqn)
forall a b. (a -> b) -> a -> b
$ List1 (Pattern, Expr)
-> ((Pattern, Expr) -> TCMT IO (Pattern, Expr))
-> TCMT IO (List1 (Pattern, Expr))
forall (t :: * -> *) (m :: * -> *) a b.
(Traversable t, Monad m) =>
t a -> (a -> m b) -> m (t b)
forM List1 (Pattern, Expr)
pes (((Pattern, Expr) -> TCMT IO (Pattern, Expr))
 -> TCMT IO (List1 (Pattern, Expr)))
-> ((Pattern, Expr) -> TCMT IO (Pattern, Expr))
-> TCMT IO (List1 (Pattern, Expr))
forall a b. (a -> b) -> a -> b
$ \ (Pattern
p, Expr
e) -> do
      -- first check the expression: the pattern may shadow
      -- some of the variables mentioned in it!
      e <- Expr -> ScopeM (AbsOfCon Expr)
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract Expr
e
      p <- toAbsPat p
      pure (p, e)
    where
      toAbsPat :: Pattern -> TCMT IO Pattern
toAbsPat Pattern
p = do
        -- Expand puns if optHiddenArgumentPuns is True.
        p <- Pattern -> TCMT IO Pattern
expandPunsOpt Pattern
p
        p <- parsePattern p
        p <- toAbstract p
        checkPatternLinearity p (typeError . RepeatedVariablesInPattern)
        bindVarsToBind
        toAbstract p

instance ToAbstract AbstractRHS where
  type AbsOfCon AbstractRHS = A.RHS

  toAbstract :: AbstractRHS -> ScopeM (AbsOfCon AbstractRHS)
toAbstract AbstractRHS
AbsurdRHS'            = RHS -> TCMT IO RHS
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return RHS
A.AbsurdRHS
  toAbstract (RHS' Expr
e Expr
c)            = AbsOfCon AbstractRHS -> ScopeM (AbsOfCon AbstractRHS)
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return (AbsOfCon AbstractRHS -> ScopeM (AbsOfCon AbstractRHS))
-> AbsOfCon AbstractRHS -> ScopeM (AbsOfCon AbstractRHS)
forall a b. (a -> b) -> a -> b
$ Expr -> Maybe Expr -> RHS
A.RHS Expr
e (Maybe Expr -> RHS) -> Maybe Expr -> RHS
forall a b. (a -> b) -> a -> b
$ Expr -> Maybe Expr
forall a. a -> Maybe a
Just Expr
c
  toAbstract (RewriteRHS' [RewriteEqn' () BindName Pattern Expr]
eqs AbstractRHS
rhs WhereDeclarations
wh) = do
    eqs <- [RewriteEqn' () BindName Pattern Expr]
-> ScopeM (AbsOfCon [RewriteEqn' () BindName Pattern Expr])
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract [RewriteEqn' () BindName Pattern Expr]
eqs
    rhs <- toAbstract rhs
    return $ RewriteRHS eqs [] rhs wh
  toAbstract (WithRHS' List1 WithExpr
es List1 (TCMT IO Clause)
cs) = do
    aux <- [Char] -> TCMT IO QName
withFunctionName [Char]
"with-"
    A.WithRHS aux es <$> do toAbstract =<< sequence cs

instance ToAbstract RightHandSide where
  type AbsOfCon RightHandSide = AbstractRHS
  toAbstract :: RightHandSide -> TCMT IO (AbsOfCon RightHandSide)
toAbstract (RightHandSide eqs :: [RewriteEqn' () BindName Pattern Expr]
eqs@(RewriteEqn' () BindName Pattern Expr
_:[RewriteEqn' () BindName Pattern Expr]
_) [WithExpr]
es ([(Name, LocalVar)], [Clause])
cs RHS' Expr
rhs WhereClause' [Declaration]
wh)               = do
    (rhs, ds) <- Range
-> WhereClause' [Declaration]
-> TCMT IO AbstractRHS
-> ScopeM (AbstractRHS, WhereDeclarations)
forall a.
Range
-> WhereClause' [Declaration]
-> ScopeM a
-> ScopeM (a, WhereDeclarations)
whereToAbstract (WhereClause' [Declaration] -> Range
forall a. HasRange a => a -> Range
getRange WhereClause' [Declaration]
wh) WhereClause' [Declaration]
wh (TCMT IO AbstractRHS -> ScopeM (AbstractRHS, WhereDeclarations))
-> TCMT IO AbstractRHS -> ScopeM (AbstractRHS, WhereDeclarations)
forall a b. (a -> b) -> a -> b
$
                   RightHandSide -> TCMT IO (AbsOfCon RightHandSide)
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract ([RewriteEqn' () BindName Pattern Expr]
-> [WithExpr]
-> ([(Name, LocalVar)], [Clause])
-> RHS' Expr
-> WhereClause' [Declaration]
-> RightHandSide
RightHandSide [] [WithExpr]
es ([(Name, LocalVar)], [Clause])
cs RHS' Expr
rhs WhereClause' [Declaration]
forall decls. WhereClause' decls
NoWhere)
    return $ RewriteRHS' eqs rhs ds
  toAbstract (RightHandSide [] []    ([(Name, LocalVar)]
_  , Clause
_:[Clause]
_) RHS' Expr
_          WhereClause' [Declaration]
_)  = TCMT IO AbstractRHS
TCMT IO (AbsOfCon RightHandSide)
forall a. HasCallStack => a
__IMPOSSIBLE__
  toAbstract (RightHandSide [] (WithExpr
_:[WithExpr]
_) ([(Name, LocalVar)], [Clause])
_         (C.RHS Expr
_)   WhereClause' [Declaration]
_)  = TCMT IO AbstractRHS
TCMT IO (AbsOfCon RightHandSide)
forall a. HasCallStack => a
__IMPOSSIBLE__  -- see test/Fail/BothWithAndRHS: triggers MissingWithClauses first
  toAbstract (RightHandSide [] []    ([(Name, LocalVar)]
_  , []) RHS' Expr
rhs         WhereClause' [Declaration]
NoWhere) = RHS' Expr -> ScopeM (AbsOfCon (RHS' Expr))
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract RHS' Expr
rhs
  toAbstract (RightHandSide [] (WithExpr
z:[WithExpr]
zs)([(Name, LocalVar)]
lv , Clause
c:[Clause]
cs) RHS' Expr
C.AbsurdRHS WhereClause' [Declaration]
NoWhere) = do
    let (List1 (Maybe (NewName BoundName))
ns, List1 (Arg Expr)
es) = (WithExpr -> (Maybe (NewName BoundName), Arg Expr))
-> List1 WithExpr
-> (List1 (Maybe (NewName BoundName)), List1 (Arg Expr))
forall a b c. (a -> (b, c)) -> List1 a -> (List1 b, List1 c)
List1.unzipWith (\ (Named Maybe Name
nm Arg Expr
e) -> (BindingSource -> BoundName -> NewName BoundName
forall a. BindingSource -> a -> NewName a
NewName BindingSource
WithBound (BoundName -> NewName BoundName)
-> (Name -> BoundName) -> Name -> NewName BoundName
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Name -> BoundName
C.mkBoundName_ (Name -> NewName BoundName)
-> Maybe Name -> Maybe (NewName BoundName)
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> Maybe Name
nm, Arg Expr
e)) (List1 WithExpr
 -> (List1 (Maybe (NewName BoundName)), List1 (Arg Expr)))
-> List1 WithExpr
-> (List1 (Maybe (NewName BoundName)), List1 (Arg Expr))
forall a b. (a -> b) -> a -> b
$ WithExpr
z WithExpr -> [WithExpr] -> List1 WithExpr
forall a. a -> [a] -> NonEmpty a
:| [WithExpr]
zs
    es <- Precedence
-> List1 (Arg Expr) -> ScopeM (AbsOfCon (List1 (Arg Expr)))
forall c. ToAbstract c => Precedence -> c -> ScopeM (AbsOfCon c)
toAbstractCtx Precedence
TopCtx List1 (Arg Expr)
es
    lvars0 <- getLocalVars
    ns <- toAbstract ns
    lvars1 <- getLocalVars
    let lv' = Int -> [(Name, LocalVar)] -> [(Name, LocalVar)]
forall a. Int -> [a] -> [a]
dropEnd ([(Name, LocalVar)] -> Int
forall a. [a] -> Int
forall (t :: * -> *) a. Foldable t => t a -> Int
length [(Name, LocalVar)]
lvars0) [(Name, LocalVar)]
lvars1 [(Name, LocalVar)] -> [(Name, LocalVar)] -> [(Name, LocalVar)]
forall a. [a] -> [a] -> [a]
++ [(Name, LocalVar)]
lv
    let cs' = NonEmpty Clause
-> (Clause -> TCMT IO Clause) -> List1 (TCMT IO Clause)
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
for (Clause
c Clause -> [Clause] -> NonEmpty Clause
forall a. a -> [a] -> NonEmpty a
:| [Clause]
cs) ((Clause -> TCMT IO Clause) -> List1 (TCMT IO Clause))
-> (Clause -> TCMT IO Clause) -> List1 (TCMT IO Clause)
forall a b. (a -> b) -> a -> b
$ \ Clause
c -> [(Name, LocalVar)] -> TCMT IO ()
setLocalVars [(Name, LocalVar)]
lv' TCMT IO () -> Clause -> TCMT IO Clause
forall (f :: * -> *) a b. Functor f => f a -> b -> f b
$> Clause
c
    let nes = (Maybe BindName -> Arg Expr -> WithExpr)
-> NonEmpty (Maybe BindName)
-> NonEmpty (Arg Expr)
-> List1 WithExpr
forall a b c.
(a -> b -> c) -> NonEmpty a -> NonEmpty b -> NonEmpty c
List1.zipWith Maybe BindName -> Arg Expr -> WithExpr
forall name a. Maybe name -> a -> Named name a
Named NonEmpty (Maybe BindName)
ns NonEmpty (Arg Expr)
es
    return $ WithRHS' nes cs'
  -- TODO: some of these might be possible
  toAbstract (RightHandSide [] (WithExpr
_ : [WithExpr]
_) ([(Name, LocalVar)], [Clause])
_ RHS' Expr
C.AbsurdRHS  AnyWhere{}) = TCMT IO AbstractRHS
TCMT IO (AbsOfCon RightHandSide)
forall a. HasCallStack => a
__IMPOSSIBLE__
  toAbstract (RightHandSide [] (WithExpr
_ : [WithExpr]
_) ([(Name, LocalVar)], [Clause])
_ RHS' Expr
C.AbsurdRHS SomeWhere{}) = TCMT IO AbstractRHS
TCMT IO (AbsOfCon RightHandSide)
forall a. HasCallStack => a
__IMPOSSIBLE__
  toAbstract (RightHandSide [] (WithExpr
_ : [WithExpr]
_) ([(Name, LocalVar)], [Clause])
_ RHS' Expr
C.AbsurdRHS   NoWhere{}) = TCMT IO AbstractRHS
TCMT IO (AbsOfCon RightHandSide)
forall a. HasCallStack => a
__IMPOSSIBLE__
  toAbstract (RightHandSide [] []     ([(Name, LocalVar)]
_, []) RHS' Expr
C.AbsurdRHS  AnyWhere{}) = TCMT IO AbstractRHS
TCMT IO (AbsOfCon RightHandSide)
forall a. HasCallStack => a
__IMPOSSIBLE__
  toAbstract (RightHandSide [] []     ([(Name, LocalVar)]
_, []) RHS' Expr
C.AbsurdRHS SomeWhere{}) = TCMT IO AbstractRHS
TCMT IO (AbsOfCon RightHandSide)
forall a. HasCallStack => a
__IMPOSSIBLE__
  toAbstract (RightHandSide [] []     ([(Name, LocalVar)]
_, []) C.RHS{}      AnyWhere{}) = TCMT IO AbstractRHS
TCMT IO (AbsOfCon RightHandSide)
forall a. HasCallStack => a
__IMPOSSIBLE__
  toAbstract (RightHandSide [] []     ([(Name, LocalVar)]
_, []) C.RHS{}     SomeWhere{}) = TCMT IO AbstractRHS
TCMT IO (AbsOfCon RightHandSide)
forall a. HasCallStack => a
__IMPOSSIBLE__

instance ToAbstract C.RHS where
    type AbsOfCon C.RHS = AbstractRHS

    toAbstract :: RHS' Expr -> ScopeM (AbsOfCon (RHS' Expr))
toAbstract RHS' Expr
C.AbsurdRHS = AbstractRHS -> TCMT IO AbstractRHS
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return (AbstractRHS -> TCMT IO AbstractRHS)
-> AbstractRHS -> TCMT IO AbstractRHS
forall a b. (a -> b) -> a -> b
$ AbstractRHS
AbsurdRHS'
    toAbstract (C.RHS Expr
e)   = Expr -> Expr -> AbstractRHS
RHS' (Expr -> Expr -> AbstractRHS)
-> ScopeM Expr -> TCMT IO (Expr -> AbstractRHS)
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> Expr -> ScopeM (AbsOfCon Expr)
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract Expr
e TCMT IO (Expr -> AbstractRHS)
-> TCMT IO Expr -> TCMT IO AbstractRHS
forall a b. TCMT IO (a -> b) -> TCMT IO a -> TCMT IO b
forall (f :: * -> *) a b. Applicative f => f (a -> b) -> f a -> f b
<*> Expr -> TCMT IO Expr
forall a. a -> TCMT IO a
forall (f :: * -> *) a. Applicative f => a -> f a
pure Expr
e

-- | Wrapper to check lhs (possibly of a 'C.DisplayPragma').
--
data LeftHandSide = LeftHandSide
  C.QName
    -- ^ Name of the definition we are checking.
  C.Pattern
    -- ^ Full left hand side.
  DisplayLHS
    -- ^ Are we checking a 'C.DisplayPragma'?

instance ToAbstract LeftHandSide where
    type AbsOfCon LeftHandSide = A.LHS

    toAbstract :: LeftHandSide -> ScopeM (AbsOfCon LeftHandSide)
toAbstract (LeftHandSide QName
top Pattern
lhs DisplayLHS
displayLhs) =
      Call
-> ScopeM (AbsOfCon LeftHandSide) -> ScopeM (AbsOfCon LeftHandSide)
forall a. Call -> TCMT IO a -> TCMT IO a
forall (m :: * -> *) a. MonadTrace m => Call -> m a -> m a
traceCall (QName -> Pattern -> Call
ScopeCheckLHS QName
top Pattern
lhs) (ScopeM (AbsOfCon LeftHandSide) -> ScopeM (AbsOfCon LeftHandSide))
-> ScopeM (AbsOfCon LeftHandSide) -> ScopeM (AbsOfCon LeftHandSide)
forall a b. (a -> b) -> a -> b
$ do
        [Char] -> Int -> [Char] -> TCMT IO ()
forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> [Char] -> m ()
reportSLn [Char]
"scope.lhs" Int
25 ([Char] -> TCMT IO ()) -> [Char] -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ [Char]
"original lhs: " [Char] -> [Char] -> [Char]
forall a. [a] -> [a] -> [a]
++ Pattern -> [Char]
forall a. Pretty a => a -> [Char]
prettyShow Pattern
lhs
        [Char] -> Int -> [Char] -> TCMT IO ()
forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> [Char] -> m ()
reportSLn [Char]
"scope.lhs" Int
60 ([Char] -> TCMT IO ()) -> [Char] -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ [Char]
"patternQNames: " [Char] -> [Char] -> [Char]
forall a. [a] -> [a] -> [a]
++ [QName] -> [Char]
forall a. Pretty a => a -> [Char]
prettyShow (Pattern -> [QName]
forall p. CPatternLike p => p -> [QName]
patternQNames Pattern
lhs)
        [Char] -> Int -> [Char] -> TCMT IO ()
forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> [Char] -> m ()
reportSLn [Char]
"scope.lhs" Int
60 ([Char] -> TCMT IO ()) -> [Char] -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ [Char]
"original lhs (raw): " [Char] -> [Char] -> [Char]
forall a. [a] -> [a] -> [a]
++ Pattern -> [Char]
forall a. Show a => a -> [Char]
show Pattern
lhs

        -- Expand puns if optHiddenArgumentPuns is True. Note that pun
        -- expansion should happen before the left-hand side is
        -- parsed, because {(x)} is not treated as a pun, whereas {x}
        -- is.
        lhs  <- Pattern -> TCMT IO Pattern
expandPunsOpt Pattern
lhs
        reportSLn "scope.lhs" 25 $
          "lhs with expanded puns: " ++ prettyShow lhs
        reportSLn "scope.lhs" 60 $
          "lhs with expanded puns (raw): " ++ show lhs

        lhscore <- parseLHS displayLhs top lhs
        let ell = LHSCore -> ExpandedEllipsis
hasExpandedEllipsis LHSCore
lhscore
        reportSLn "scope.lhs" 25 $ "parsed lhs: " ++ prettyShow lhscore
        reportSLn "scope.lhs" 60 $ "parsed lhs (raw): " ++ show lhscore
        printLocals 30 "before lhs:"
        -- error if copattern parsed but --no-copatterns option
        unlessM (optCopatterns <$> pragmaOptions) $
          when (hasCopatterns lhscore) $
            typeError $ NeedOptionCopatterns
        -- scope check patterns except for dot patterns
        lhscore <- toAbstract $ CLHSCore displayLhs lhscore
        bindVarsToBind
        -- reportSLn "scope.lhs" 25 $ "parsed lhs patterns: " ++ prettyShow lhscore  -- TODO: Pretty A.LHSCore'
        reportSLn "scope.lhs" 60 $ "parsed lhs patterns: " ++ show lhscore
        printLocals 30 "checked pattern:"
        -- scope check dot patterns
        lhscore <- toAbstract lhscore
        -- reportSLn "scope.lhs" 25 $ "parsed lhs dot patterns: " ++ prettyShow lhscore  -- TODO: Pretty A.LHSCore'
        reportSLn "scope.lhs" 60 $ "parsed lhs dot patterns: " ++ show lhscore
        printLocals 30 "checked dots:"
        return $ A.LHS (LHSInfo (getRange lhs) ell) lhscore

-- | Expands hidden argument puns when option 'optHiddenArgumentPuns' is set.

expandPunsOpt :: C.Pattern -> ScopeM C.Pattern
expandPunsOpt :: Pattern -> TCMT IO Pattern
expandPunsOpt Pattern
p = do
  TCMT IO PragmaOptions
forall (m :: * -> *). HasOptions m => m PragmaOptions
pragmaOptions TCMT IO PragmaOptions -> (PragmaOptions -> Bool) -> TCMT IO Bool
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> PragmaOptions -> Bool
optHiddenArgumentPuns TCMT IO Bool -> (Bool -> Pattern) -> TCMT IO Pattern
forall (f :: * -> *) a b. Functor f => f a -> (a -> b) -> f b
<&> \case
    Bool
True  -> KillRangeT Pattern
expandPuns Pattern
p
    Bool
False -> Pattern
p

-- | Expands hidden argument puns.

expandPuns :: C.Pattern -> C.Pattern
expandPuns :: KillRangeT Pattern
expandPuns = KillRangeT Pattern -> KillRangeT Pattern
forall p. CPatternLike p => KillRangeT Pattern -> p -> p
mapCPattern \case
  C.HiddenP   Range
r Named NamedName Pattern
p -> Range -> Named NamedName Pattern -> Pattern
C.HiddenP   Range
r (Named NamedName Pattern -> Pattern)
-> Named NamedName Pattern -> Pattern
forall a b. (a -> b) -> a -> b
$ Named NamedName Pattern -> Named NamedName Pattern
expand Named NamedName Pattern
p
  C.InstanceP Range
r Named NamedName Pattern
p -> Range -> Named NamedName Pattern -> Pattern
C.InstanceP Range
r (Named NamedName Pattern -> Pattern)
-> Named NamedName Pattern -> Pattern
forall a b. (a -> b) -> a -> b
$ Named NamedName Pattern -> Named NamedName Pattern
expand Named NamedName Pattern
p
  Pattern
p -> Pattern
p
  where
  -- Only patterns of the form {x} or ⦃ x ⦄, where x is an unqualified
  -- name (not @_@), are interpreted as puns.
  expand :: Named_ C.Pattern -> Named_ C.Pattern
  expand :: Named NamedName Pattern -> Named NamedName Pattern
expand
   (Named { namedThing :: forall name a. Named name a -> a
namedThing = C.IdentP Bool
_ q :: QName
q@(C.QName x :: Name
x@C.Name{})
          , nameOf :: forall name a. Named name a -> Maybe name
nameOf     = Maybe NamedName
Nothing
          }) =
    Named { namedThing :: Pattern
namedThing = Bool -> QName -> Pattern
C.IdentP Bool
False QName
q
          , nameOf :: Maybe NamedName
nameOf     = NamedName -> Maybe NamedName
forall a. a -> Maybe a
Just (NamedName -> Maybe NamedName) -> NamedName -> Maybe NamedName
forall a b. (a -> b) -> a -> b
$
                         WithOrigin
                           { woOrigin :: Origin
woOrigin = Origin
ExpandedPun
                           , woThing :: RString
woThing  = [Char] -> RString
forall a. a -> Ranged a
unranged (Name -> [Char]
forall a. Pretty a => a -> [Char]
prettyShow Name
x)
                           }
          }
  expand Named NamedName Pattern
p = Named NamedName Pattern
p

hasExpandedEllipsis :: C.LHSCore -> ExpandedEllipsis
hasExpandedEllipsis :: LHSCore -> ExpandedEllipsis
hasExpandedEllipsis LHSCore
core = case LHSCore
core of
  C.LHSHead{}       -> ExpandedEllipsis
NoEllipsis
  C.LHSProj{}       -> LHSCore -> ExpandedEllipsis
hasExpandedEllipsis (LHSCore -> ExpandedEllipsis) -> LHSCore -> ExpandedEllipsis
forall a b. (a -> b) -> a -> b
$ NamedArg LHSCore -> LHSCore
forall a. NamedArg a -> a
namedArg (NamedArg LHSCore -> LHSCore) -> NamedArg LHSCore -> LHSCore
forall a b. (a -> b) -> a -> b
$ LHSCore -> NamedArg LHSCore
C.lhsFocus LHSCore
core -- can this ever be ExpandedEllipsis?
  C.LHSWith{}       -> LHSCore -> ExpandedEllipsis
hasExpandedEllipsis (LHSCore -> ExpandedEllipsis) -> LHSCore -> ExpandedEllipsis
forall a b. (a -> b) -> a -> b
$ LHSCore -> LHSCore
C.lhsHead LHSCore
core
  C.LHSEllipsis Range
r LHSCore
p -> case LHSCore
p of
    C.LHSWith LHSCore
p List1 Pattern
wps [NamedArg Pattern]
_ -> LHSCore -> ExpandedEllipsis
hasExpandedEllipsis LHSCore
p ExpandedEllipsis -> ExpandedEllipsis -> ExpandedEllipsis
forall a. Semigroup a => a -> a -> a
<> Range -> Int -> ExpandedEllipsis
ExpandedEllipsis Range
r (List1 Pattern -> Int
forall a. NonEmpty a -> Int
forall (t :: * -> *) a. Foldable t => t a -> Int
length List1 Pattern
wps)
    C.LHSHead{}       -> Range -> Int -> ExpandedEllipsis
ExpandedEllipsis Range
r Int
0
    C.LHSProj{}       -> Range -> Int -> ExpandedEllipsis
ExpandedEllipsis Range
r Int
0
    C.LHSEllipsis{}   -> ExpandedEllipsis
forall a. HasCallStack => a
__IMPOSSIBLE__

-- | Merges adjacent EqualP patterns into one:
-- type checking expects only one pattern for each domain in the telescope.
mergeEqualPs :: [NamedArg (Pattern' e)] -> ScopeM [NamedArg (Pattern' e)]
mergeEqualPs :: forall e. [NamedArg (Pattern' e)] -> ScopeM [NamedArg (Pattern' e)]
mergeEqualPs = (PatInfo, [(e, e)])
-> [Arg (Named NamedName (Pattern' e))]
-> TCMT IO [Arg (Named NamedName (Pattern' e))]
forall {e}.
(PatInfo, [(e, e)])
-> [Arg (Named NamedName (Pattern' e))]
-> TCMT IO [Arg (Named NamedName (Pattern' e))]
go (PatInfo
forall a. Null a => a
empty, [])
  where
    go :: (PatInfo, [(e, e)])
-> [Arg (Named NamedName (Pattern' e))]
-> TCMT IO [Arg (Named NamedName (Pattern' e))]
go (PatInfo, [(e, e)])
acc (p :: Arg (Named NamedName (Pattern' e))
p@(Arg ArgInfo
ai (Named Maybe NamedName
mn (A.EqualP PatInfo
r List1 (e, e)
es))) : [Arg (Named NamedName (Pattern' e))]
ps) = Arg (Named NamedName (Pattern' e))
-> TCMT IO [Arg (Named NamedName (Pattern' e))]
-> TCMT IO [Arg (Named NamedName (Pattern' e))]
forall (m :: * -> *) x a.
(MonadTrace m, HasRange x) =>
x -> m a -> m a
setCurrentRange Arg (Named NamedName (Pattern' e))
p (TCMT IO [Arg (Named NamedName (Pattern' e))]
 -> TCMT IO [Arg (Named NamedName (Pattern' e))])
-> TCMT IO [Arg (Named NamedName (Pattern' e))]
-> TCMT IO [Arg (Named NamedName (Pattern' e))]
forall a b. (a -> b) -> a -> b
$ do
      -- Face constraint patterns must be defaultNamedArg; check this:
      Bool -> TCMT IO () -> TCMT IO ()
forall b (m :: * -> *). (IsBool b, Monad m) => b -> m () -> m ()
unless (ArgInfo -> Modality
forall a. LensModality a => a -> Modality
getModality ArgInfo
ai Modality -> Modality -> Bool
forall a. Eq a => a -> a -> Bool
== Modality
defaultModality) TCMT IO ()
forall a. HasCallStack => a
__IMPOSSIBLE__
      Bool -> TCMT IO () -> TCMT IO ()
forall b (m :: * -> *). (IsBool b, Monad m) => b -> m () -> m ()
when (ArgInfo -> Bool
forall a. LensHiding a => a -> Bool
notVisible ArgInfo
ai) (TCMT IO () -> TCMT IO ()) -> TCMT IO () -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$
        Warning -> TCMT IO ()
forall (m :: * -> *).
(HasCallStack, MonadWarning m) =>
Warning -> m ()
warning (Warning -> TCMT IO ()) -> Warning -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ ArgInfo -> Warning
FaceConstraintCannotBeHidden ArgInfo
ai
      Maybe NamedName -> (NamedName -> TCMT IO ()) -> TCMT IO ()
forall (m :: * -> *) a. Monad m => Maybe a -> (a -> m ()) -> m ()
whenJust Maybe NamedName
mn ((NamedName -> TCMT IO ()) -> TCMT IO ())
-> (NamedName -> TCMT IO ()) -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ \ NamedName
x -> NamedName -> TCMT IO () -> TCMT IO ()
forall (m :: * -> *) x a.
(MonadTrace m, HasRange x) =>
x -> m a -> m a
setCurrentRange NamedName
x (TCMT IO () -> TCMT IO ()) -> TCMT IO () -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$
        Warning -> TCMT IO ()
forall (m :: * -> *).
(HasCallStack, MonadWarning m) =>
Warning -> m ()
warning (Warning -> TCMT IO ()) -> Warning -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ NamedName -> Warning
FaceConstraintCannotBeNamed NamedName
x
      (PatInfo, [(e, e)])
-> [Arg (Named NamedName (Pattern' e))]
-> TCMT IO [Arg (Named NamedName (Pattern' e))]
go ((PatInfo, [(e, e)])
acc (PatInfo, [(e, e)]) -> (PatInfo, [(e, e)]) -> (PatInfo, [(e, e)])
forall a. Monoid a => a -> a -> a
`mappend` (PatInfo
r, List1 (e, e) -> [Item (List1 (e, e))]
forall l. IsList l => l -> [Item l]
List1.toList List1 (e, e)
es)) [Arg (Named NamedName (Pattern' e))]
ps
    go (PatInfo
r, ((e, e)
e:[(e, e)]
es))   [Arg (Named NamedName (Pattern' e))]
ps = (Pattern' e -> Arg (Named NamedName (Pattern' e))
forall a. a -> NamedArg a
defaultNamedArg (PatInfo -> List1 (e, e) -> Pattern' e
forall e. PatInfo -> List1 (e, e) -> Pattern' e
A.EqualP PatInfo
r (List1 (e, e) -> Pattern' e) -> List1 (e, e) -> Pattern' e
forall a b. (a -> b) -> a -> b
$ (e, e)
e (e, e) -> [(e, e)] -> List1 (e, e)
forall a. a -> [a] -> NonEmpty a
:| [(e, e)]
es) Arg (Named NamedName (Pattern' e))
-> [Arg (Named NamedName (Pattern' e))]
-> [Arg (Named NamedName (Pattern' e))]
forall a. a -> [a] -> [a]
:) ([Arg (Named NamedName (Pattern' e))]
 -> [Arg (Named NamedName (Pattern' e))])
-> TCMT IO [Arg (Named NamedName (Pattern' e))]
-> TCMT IO [Arg (Named NamedName (Pattern' e))]
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> [Arg (Named NamedName (Pattern' e))]
-> TCMT IO [Arg (Named NamedName (Pattern' e))]
forall e. [NamedArg (Pattern' e)] -> ScopeM [NamedArg (Pattern' e)]
mergeEqualPs [Arg (Named NamedName (Pattern' e))]
ps
    go (PatInfo
_, [])       [] = [Arg (Named NamedName (Pattern' e))]
-> TCMT IO [Arg (Named NamedName (Pattern' e))]
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return []
    go (PatInfo
_, []) (Arg (Named NamedName (Pattern' e))
p : [Arg (Named NamedName (Pattern' e))]
ps) = (Arg (Named NamedName (Pattern' e))
p Arg (Named NamedName (Pattern' e))
-> [Arg (Named NamedName (Pattern' e))]
-> [Arg (Named NamedName (Pattern' e))]
forall a. a -> [a] -> [a]
:) ([Arg (Named NamedName (Pattern' e))]
 -> [Arg (Named NamedName (Pattern' e))])
-> TCMT IO [Arg (Named NamedName (Pattern' e))]
-> TCMT IO [Arg (Named NamedName (Pattern' e))]
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> [Arg (Named NamedName (Pattern' e))]
-> TCMT IO [Arg (Named NamedName (Pattern' e))]
forall e. [NamedArg (Pattern' e)] -> ScopeM [NamedArg (Pattern' e)]
mergeEqualPs [Arg (Named NamedName (Pattern' e))]
ps


-- | Scope-check a 'C.LHSCore' (of possibly a 'C.DisplayForm').

data CLHSCore = CLHSCore
  DisplayLHS
    -- ^ Are we checking the left hand side of a 'C.DisplayForm'?
  C.LHSCore
    -- ^ The lhs to scope-check.

-- | Scope-check a 'C.LHSCore' not of a 'C.DisplayForm'.

instance ToAbstract C.LHSCore where
  type AbsOfCon C.LHSCore = A.LHSCore' C.Expr

  toAbstract :: LHSCore -> ScopeM (AbsOfCon LHSCore)
toAbstract = CLHSCore -> TCMT IO (LHSCore' Expr)
CLHSCore -> ScopeM (AbsOfCon CLHSCore)
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract (CLHSCore -> TCMT IO (LHSCore' Expr))
-> (LHSCore -> CLHSCore) -> LHSCore -> TCMT IO (LHSCore' Expr)
forall b c a. (b -> c) -> (a -> b) -> a -> c
. DisplayLHS -> LHSCore -> CLHSCore
CLHSCore DisplayLHS
NoDisplayLHS

-- does not check pattern linearity
instance ToAbstract CLHSCore where
  type AbsOfCon CLHSCore = A.LHSCore' C.Expr

  toAbstract :: CLHSCore -> ScopeM (AbsOfCon CLHSCore)
toAbstract (CLHSCore DisplayLHS
displayLhs LHSCore
core0) = case LHSCore
core0 of

    C.LHSHead QName
x [NamedArg Pattern]
ps -> do
        x <- TCMT IO QName -> TCMT IO QName
forall a. ScopeM a -> ScopeM a
withLocalVars do
          [(Name, LocalVar)] -> TCMT IO ()
setLocalVars []
          OldName QName -> ScopeM (AbsOfCon (OldName QName))
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract (QName -> OldName QName
forall a. a -> OldName a
OldName QName
x)
        ps <- toAbstract $ (fmap . fmap . fmap) (CPattern displayLhs) ps
        A.LHSHead x <$> mergeEqualPs ps

    C.LHSProj QName
d [NamedArg Pattern]
ps1 NamedArg LHSCore
core [NamedArg Pattern]
ps2 -> do
        Bool -> TCMT IO () -> TCMT IO ()
forall b (m :: * -> *). (IsBool b, Monad m) => b -> m () -> m ()
unless ([NamedArg Pattern] -> Bool
forall a. Null a => a -> Bool
null [NamedArg Pattern]
ps1) (TCMT IO () -> TCMT IO ()) -> TCMT IO () -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ TypeError -> TCMT IO ()
forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
TypeError -> m a
typeError (TypeError -> TCMT IO ()) -> TypeError -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ Pattern -> TypeError
IllformedProjectionPatternConcrete ((Pattern -> NamedArg Pattern -> Pattern)
-> Pattern -> [NamedArg Pattern] -> Pattern
forall b a. (b -> a -> b) -> b -> [a] -> b
forall (t :: * -> *) b a.
Foldable t =>
(b -> a -> b) -> b -> t a -> b
foldl Pattern -> NamedArg Pattern -> Pattern
C.AppP (Bool -> QName -> Pattern
C.IdentP Bool
True QName
d) [NamedArg Pattern]
ps1)
        ds <- QName -> ScopeM ResolvedName
resolveName QName
d ScopeM ResolvedName
-> (ResolvedName -> TCMT IO (List1 QName)) -> TCMT IO (List1 QName)
forall a b. TCMT IO a -> (a -> TCMT IO b) -> TCMT IO b
forall (m :: * -> *) a b. Monad m => m a -> (a -> m b) -> m b
>>= \case
          FieldName List1 AbstractName
ds -> List1 QName -> TCMT IO (List1 QName)
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return (List1 QName -> TCMT IO (List1 QName))
-> List1 QName -> TCMT IO (List1 QName)
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
ds
          ResolvedName
UnknownName  -> QName -> TCMT IO (List1 QName)
forall a. QName -> TCM a
notInScopeError QName
d
          ResolvedName
_            -> TypeError -> TCMT IO (List1 QName)
forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
TypeError -> m a
typeError (TypeError -> TCMT IO (List1 QName))
-> TypeError -> TCMT IO (List1 QName)
forall a b. (a -> b) -> a -> b
$ QName -> TypeError
CopatternHeadNotProjection QName
d
        core <- toAbstract $ (fmap . fmap) (CLHSCore displayLhs) core
        ps2  <- toAbstract $ (fmap . fmap . fmap) (CPattern displayLhs) ps2
        A.LHSProj (AmbQ ds) core <$> mergeEqualPs ps2

    C.LHSWith LHSCore
core List1 Pattern
wps [NamedArg Pattern]
ps -> do
      -- DISPLAY pragmas cannot have @with@, so no need to pass on @displayLhs@.
      core <- LHSCore -> ScopeM (AbsOfCon LHSCore)
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract LHSCore
core
      wps  <- fmap defaultArg <$> toAbstract wps
      ps   <- toAbstract ps
      return $ A.lhsCoreApp (A.lhsCoreWith core wps) ps

    -- In case of a part of the LHS which was expanded from an ellipsis,
    -- we flush the @scopeVarsToBind@ in order to allow variables bound
    -- in the ellipsis to be shadowed.
    C.LHSEllipsis Range
_ LHSCore
core -> do
      core <- LHSCore -> ScopeM (AbsOfCon LHSCore)
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract LHSCore
core  -- Cannot come from a DISPLAY pragma.
      bindVarsToBind
      return core

instance ToAbstract c => ToAbstract (WithHiding c) where
  type AbsOfCon (WithHiding c) = WithHiding (AbsOfCon c)
  toAbstract :: WithHiding c -> ScopeM (AbsOfCon (WithHiding c))
toAbstract (WithHiding Hiding
h c
a) = Hiding -> AbsOfCon c -> WithHiding (AbsOfCon c)
forall a. Hiding -> a -> WithHiding a
WithHiding Hiding
h (AbsOfCon c -> WithHiding (AbsOfCon c))
-> TCMT IO (AbsOfCon c) -> TCMT IO (WithHiding (AbsOfCon c))
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> Hiding -> c -> TCMT IO (AbsOfCon c)
forall h c.
(LensHiding h, ToAbstract c) =>
h -> c -> ScopeM (AbsOfCon c)
toAbstractHiding Hiding
h c
a

instance ToAbstract c => ToAbstract (Arg c) where
    type AbsOfCon (Arg c) = Arg (AbsOfCon c)
    toAbstract :: Arg c -> ScopeM (AbsOfCon (Arg c))
toAbstract (Arg ArgInfo
info c
e) =
        ArgInfo -> AbsOfCon c -> Arg (AbsOfCon c)
forall e. ArgInfo -> e -> Arg e
Arg ArgInfo
info (AbsOfCon c -> Arg (AbsOfCon c))
-> TCMT IO (AbsOfCon c) -> TCMT IO (Arg (AbsOfCon c))
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> ArgInfo -> c -> TCMT IO (AbsOfCon c)
forall h c.
(LensHiding h, ToAbstract c) =>
h -> c -> ScopeM (AbsOfCon c)
toAbstractHiding ArgInfo
info c
e

instance ToAbstract c => ToAbstract (Named name c) where
    type AbsOfCon (Named name c) = Named name (AbsOfCon c)
    toAbstract :: Named name c -> ScopeM (AbsOfCon (Named name c))
toAbstract = (c -> TCMT IO (AbsOfCon c))
-> Named name c -> TCMT IO (Named name (AbsOfCon c))
forall (t :: * -> *) (f :: * -> *) a b.
(Traversable t, Applicative f) =>
(a -> f b) -> t a -> f (t b)
forall (f :: * -> *) a b.
Applicative f =>
(a -> f b) -> Named name a -> f (Named name b)
traverse c -> TCMT IO (AbsOfCon c)
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract

instance ToAbstract c => ToAbstract (Ranged c) where
    type AbsOfCon (Ranged c) = Ranged (AbsOfCon c)
    toAbstract :: Ranged c -> ScopeM (AbsOfCon (Ranged c))
toAbstract = (c -> TCMT IO (AbsOfCon c))
-> Ranged c -> TCMT IO (Ranged (AbsOfCon c))
forall (t :: * -> *) (f :: * -> *) a b.
(Traversable t, Applicative f) =>
(a -> f b) -> t a -> f (t b)
forall (f :: * -> *) a b.
Applicative f =>
(a -> f b) -> Ranged a -> f (Ranged b)
traverse c -> TCMT IO (AbsOfCon c)
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract

{- DOES NOT WORK ANYMORE with pattern synonyms
instance ToAbstract c a => ToAbstract (A.LHSCore' c) (A.LHSCore' a) where
    toAbstract = mapM toAbstract
-}

instance ToAbstract (A.LHSCore' C.Expr) where
    type AbsOfCon (A.LHSCore' C.Expr) = A.LHSCore' A.Expr
    toAbstract :: LHSCore' Expr -> ScopeM (AbsOfCon (LHSCore' Expr))
toAbstract (A.LHSHead QName
f [NamedArg (Pattern' Expr)]
ps)         = QName -> NAPs Expr -> LHSCore' Expr
forall e. QName -> [NamedArg (Pattern' e)] -> LHSCore' e
A.LHSHead QName
f (NAPs Expr -> LHSCore' Expr)
-> TCMT IO (NAPs Expr) -> TCMT IO (LHSCore' Expr)
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> (NamedArg (Pattern' Expr) -> TCMT IO (Arg (Named_ Pattern)))
-> [NamedArg (Pattern' Expr)] -> TCMT IO (NAPs Expr)
forall (t :: * -> *) (m :: * -> *) a b.
(Traversable t, Monad m) =>
(a -> m b) -> t a -> m (t b)
forall (m :: * -> *) a b. Monad m => (a -> m b) -> [a] -> m [b]
mapM NamedArg (Pattern' Expr) -> TCMT IO (Arg (Named_ Pattern))
NamedArg (Pattern' Expr)
-> ScopeM (AbsOfCon (NamedArg (Pattern' Expr)))
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract [NamedArg (Pattern' Expr)]
ps
    toAbstract (A.LHSProj AmbiguousQName
d NamedArg (LHSCore' Expr)
lhscore [NamedArg (Pattern' Expr)]
ps) = AmbiguousQName
-> NamedArg (LHSCore' Expr) -> NAPs Expr -> LHSCore' Expr
forall e.
AmbiguousQName
-> NamedArg (LHSCore' e) -> [NamedArg (Pattern' e)] -> LHSCore' e
A.LHSProj AmbiguousQName
d (NamedArg (LHSCore' Expr) -> NAPs Expr -> LHSCore' Expr)
-> TCMT IO (NamedArg (LHSCore' Expr))
-> TCMT IO (NAPs Expr -> LHSCore' Expr)
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> (Named_ (LHSCore' Expr) -> TCMT IO (Named_ (LHSCore' Expr)))
-> NamedArg (LHSCore' Expr) -> TCMT IO (NamedArg (LHSCore' Expr))
forall (t :: * -> *) (m :: * -> *) a b.
(Traversable t, Monad m) =>
(a -> m b) -> t a -> m (t b)
forall (m :: * -> *) a b.
Monad m =>
(a -> m b) -> Arg a -> m (Arg b)
mapM Named_ (LHSCore' Expr) -> TCMT IO (Named_ (LHSCore' Expr))
Named_ (LHSCore' Expr)
-> ScopeM (AbsOfCon (Named_ (LHSCore' Expr)))
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract NamedArg (LHSCore' Expr)
lhscore TCMT IO (NAPs Expr -> LHSCore' Expr)
-> TCMT IO (NAPs Expr) -> TCMT IO (LHSCore' Expr)
forall a b. TCMT IO (a -> b) -> TCMT IO a -> TCMT IO b
forall (f :: * -> *) a b. Applicative f => f (a -> b) -> f a -> f b
<*> (NamedArg (Pattern' Expr) -> TCMT IO (Arg (Named_ Pattern)))
-> [NamedArg (Pattern' Expr)] -> TCMT IO (NAPs Expr)
forall (t :: * -> *) (m :: * -> *) a b.
(Traversable t, Monad m) =>
(a -> m b) -> t a -> m (t b)
forall (m :: * -> *) a b. Monad m => (a -> m b) -> [a] -> m [b]
mapM NamedArg (Pattern' Expr) -> TCMT IO (Arg (Named_ Pattern))
NamedArg (Pattern' Expr)
-> ScopeM (AbsOfCon (NamedArg (Pattern' Expr)))
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract [NamedArg (Pattern' Expr)]
ps
    toAbstract (A.LHSWith LHSCore' Expr
core NonEmpty (Arg (Pattern' Expr))
wps [NamedArg (Pattern' Expr)]
ps)  = (LHSCore' Expr
 -> List1 (Arg Pattern) -> NAPs Expr -> LHSCore' Expr)
-> TCMT IO (LHSCore' Expr)
-> TCMT IO (List1 (Arg Pattern))
-> TCMT IO (NAPs Expr)
-> TCMT IO (LHSCore' Expr)
forall (f :: * -> *) a b c d.
Applicative f =>
(a -> b -> c -> d) -> f a -> f b -> f c -> f d
liftA3 LHSCore' Expr -> List1 (Arg Pattern) -> NAPs Expr -> LHSCore' Expr
forall e.
LHSCore' e
-> List1 (Arg (Pattern' e))
-> [NamedArg (Pattern' e)]
-> LHSCore' e
A.LHSWith (LHSCore' Expr -> ScopeM (AbsOfCon (LHSCore' Expr))
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract LHSCore' Expr
core) (NonEmpty (Arg (Pattern' Expr))
-> ScopeM (AbsOfCon (NonEmpty (Arg (Pattern' Expr))))
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract NonEmpty (Arg (Pattern' Expr))
wps) ([NamedArg (Pattern' Expr)]
-> ScopeM (AbsOfCon [NamedArg (Pattern' Expr)])
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract [NamedArg (Pattern' Expr)]
ps)

-- Patterns are done in two phases. First everything but the dot patterns, and
-- then the dot patterns. This is because dot patterns can refer to variables
-- bound anywhere in the pattern.

instance ToAbstract (A.Pattern' C.Expr) where
  type AbsOfCon (A.Pattern' C.Expr) = A.Pattern' A.Expr
  toAbstract :: Pattern' Expr -> ScopeM (AbsOfCon (Pattern' Expr))
toAbstract = (Expr -> ScopeM Expr) -> Pattern' Expr -> TCMT IO Pattern
forall (t :: * -> *) (f :: * -> *) a b.
(Traversable t, Applicative f) =>
(a -> f b) -> t a -> f (t b)
forall (f :: * -> *) a b.
Applicative f =>
(a -> f b) -> Pattern' a -> f (Pattern' b)
traverse ((Expr -> ScopeM Expr) -> Pattern' Expr -> TCMT IO Pattern)
-> (Expr -> ScopeM Expr) -> Pattern' Expr -> TCMT IO Pattern
forall a b. (a -> b) -> a -> b
$ ScopeM Expr -> ScopeM Expr
forall a. ScopeM a -> ScopeM a
insideDotPattern (ScopeM Expr -> ScopeM Expr)
-> (Expr -> ScopeM Expr) -> Expr -> ScopeM Expr
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Precedence -> Expr -> ScopeM (AbsOfCon Expr)
forall c. ToAbstract c => Precedence -> c -> ScopeM (AbsOfCon c)
toAbstractCtx Precedence
DotPatternCtx  -- Issue #3033

resolvePatternIdentifier ::
     Bool
       -- ^ Is the identifier allowed to refer to a constructor (or a pattern synonym)?
       --
       --   Value 'False' is only used when 'optHiddenArgumentPuns' is 'True'.
       --   In this case, error 'InvalidPun' is thrown on identifiers that are not variables.
  -> DisplayLHS
       -- ^ Are definitions to be treated as constructors?
       --   'True' when we are checking a 'C.DisplayForm'.
  -> Hiding
       -- ^ Is the pattern variable hidden?
  -> C.QName
       -- ^ Identifier.
  -> Maybe (Set1 A.Name)
       -- ^ Possibly precomputed resolutions of the identifier (from the operator parser).
  -> ScopeM (A.Pattern' C.Expr)
resolvePatternIdentifier :: Bool
-> DisplayLHS
-> Hiding
-> QName
-> Maybe (Set1 Name)
-> TCMT IO (Pattern' Expr)
resolvePatternIdentifier Bool
canBeConstructor DisplayLHS
displayLhs Hiding
h QName
x Maybe (Set1 Name)
ns = do
  [Char] -> Int -> [Char] -> TCMT IO ()
forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> [Char] -> m ()
reportSLn [Char]
"scope.pat" Int
60 ([Char] -> TCMT IO ()) -> [Char] -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ [Char]
"resolvePatternIdentifier " [Char] -> [Char] -> [Char]
forall a. [a] -> [a] -> [a]
++ QName -> [Char]
forall a. Pretty a => a -> [Char]
prettyShow QName
x [Char] -> [Char] -> [Char]
forall a. [a] -> [a] -> [a]
++ [Char]
" at source position " [Char] -> [Char] -> [Char]
forall a. [a] -> [a] -> [a]
++ Range -> [Char]
forall a. Pretty a => a -> [Char]
prettyShow Range
r
  PatName -> ScopeM (AbsOfCon PatName)
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract (QName -> Maybe (Set1 Name) -> Hiding -> DisplayLHS -> PatName
PatName QName
x Maybe (Set1 Name)
ns Hiding
h DisplayLHS
displayLhs) TCMT IO APatName
-> (APatName -> TCMT IO (Pattern' Expr)) -> TCMT IO (Pattern' Expr)
forall a b. TCMT IO a -> (a -> TCMT IO b) -> TCMT IO b
forall (m :: * -> *) a b. Monad m => m a -> (a -> m b) -> m b
>>= \case

    VarPatName Name
y -> do
      [Char] -> Int -> [Char] -> TCMT IO ()
forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> [Char] -> m ()
reportSLn [Char]
"scope.pat" Int
60 ([Char] -> TCMT IO ()) -> [Char] -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ [Char]
"  resolved to VarPatName " [Char] -> [Char] -> [Char]
forall a. [a] -> [a] -> [a]
++ Name -> [Char]
forall a. Pretty a => a -> [Char]
prettyShow Name
y [Char] -> [Char] -> [Char]
forall a. [a] -> [a] -> [a]
++ [Char]
" with range " [Char] -> [Char] -> [Char]
forall a. [a] -> [a] -> [a]
++ Range -> [Char]
forall a. Pretty a => a -> [Char]
prettyShow (Name -> Range
forall a. HasRange a => a -> Range
getRange Name
y)
      Pattern' Expr -> TCMT IO (Pattern' Expr)
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return (Pattern' Expr -> TCMT IO (Pattern' Expr))
-> Pattern' Expr -> TCMT IO (Pattern' Expr)
forall a b. (a -> b) -> a -> b
$ BindName -> Pattern' Expr
forall e. BindName -> Pattern' e
VarP (BindName -> Pattern' Expr) -> BindName -> Pattern' Expr
forall a b. (a -> b) -> a -> b
$ Name -> BindName
A.mkBindName Name
y

    ConPatName List1 AbstractName
ds -> do
      Bool -> TCMT IO () -> TCMT IO ()
forall b (m :: * -> *). (IsBool b, Monad m) => b -> m () -> m ()
unless Bool
canBeConstructor (TCMT IO () -> TCMT IO ()) -> TCMT IO () -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ ConstructorOrPatternSynonym -> TCMT IO ()
err ConstructorOrPatternSynonym
IsConstructor
      Pattern' Expr -> TCMT IO (Pattern' Expr)
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return (Pattern' Expr -> TCMT IO (Pattern' Expr))
-> Pattern' Expr -> TCMT IO (Pattern' Expr)
forall a b. (a -> b) -> a -> b
$ ConPatInfo
-> AmbiguousQName -> [NamedArg (Pattern' Expr)] -> Pattern' Expr
forall e. ConPatInfo -> AmbiguousQName -> NAPs e -> Pattern' e
ConP (ConOrigin -> PatInfo -> ConPatLazy -> ConPatInfo
ConPatInfo ConOrigin
ConOCon PatInfo
info ConPatLazy
ConPatEager) (List1 QName -> AmbiguousQName
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
ds) []

    PatternSynPatName List1 AbstractName
ds -> do
      Bool -> TCMT IO () -> TCMT IO ()
forall b (m :: * -> *). (IsBool b, Monad m) => b -> m () -> m ()
unless Bool
canBeConstructor (TCMT IO () -> TCMT IO ()) -> TCMT IO () -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ ConstructorOrPatternSynonym -> TCMT IO ()
err ConstructorOrPatternSynonym
IsPatternSynonym
      Pattern' Expr -> TCMT IO (Pattern' Expr)
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return (Pattern' Expr -> TCMT IO (Pattern' Expr))
-> Pattern' Expr -> TCMT IO (Pattern' Expr)
forall a b. (a -> b) -> a -> b
$ PatInfo
-> AmbiguousQName -> [NamedArg (Pattern' Expr)] -> Pattern' Expr
forall e. PatInfo -> AmbiguousQName -> NAPs e -> Pattern' e
PatternSynP PatInfo
info (List1 QName -> AmbiguousQName
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
ds) []

    DefPatName AbstractName
d -> do
      DisplayLHS -> TCMT IO () -> TCMT IO ()
forall b (m :: * -> *). (IsBool b, Monad m) => b -> m () -> m ()
unless DisplayLHS
displayLhs TCMT IO ()
forall a. HasCallStack => a
__IMPOSSIBLE__
      Pattern' Expr -> TCMT IO (Pattern' Expr)
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return (Pattern' Expr -> TCMT IO (Pattern' Expr))
-> Pattern' Expr -> TCMT IO (Pattern' Expr)
forall a b. (a -> b) -> a -> b
$ PatInfo
-> AmbiguousQName -> [NamedArg (Pattern' Expr)] -> Pattern' Expr
forall e. PatInfo -> AmbiguousQName -> NAPs e -> Pattern' e
DefP PatInfo
info (List1 QName -> AmbiguousQName
AmbQ (List1 QName -> AmbiguousQName) -> List1 QName -> AmbiguousQName
forall a b. (a -> b) -> a -> b
$ QName -> List1 QName
forall el coll. Singleton el coll => el -> coll
singleton (QName -> List1 QName) -> QName -> List1 QName
forall a b. (a -> b) -> a -> b
$ AbstractName -> QName
anameName AbstractName
d) []

  where
  r :: Range
r = QName -> Range
forall a. HasRange a => a -> Range
getRange QName
x
  info :: PatInfo
info = Range -> PatInfo
PatRange Range
r
  err :: ConstructorOrPatternSynonym -> TCMT IO ()
err ConstructorOrPatternSynonym
s = Range -> TCMT IO () -> TCMT IO ()
forall (m :: * -> *) x a.
(MonadTrace m, HasRange x) =>
x -> m a -> m a
setCurrentRange Range
r (TCMT IO () -> TCMT IO ()) -> TCMT IO () -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ TypeError -> TCMT IO ()
forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
TypeError -> m a
typeError (TypeError -> TCMT IO ()) -> TypeError -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ ConstructorOrPatternSynonym -> QName -> TypeError
InvalidPun ConstructorOrPatternSynonym
s QName
x

-- | Apply an abstract syntax pattern head to pattern arguments.
--
--   Fails with 'InvalidPattern' if head is not a constructor pattern
--   (or similar) that can accept arguments.
--
applyAPattern
  :: C.Pattern            -- ^ The application pattern in concrete syntax.
  -> A.Pattern' C.Expr    -- ^ Head of application.
  -> NAPs1 C.Expr         -- ^ Arguments of application.
  -> ScopeM (A.Pattern' C.Expr)
applyAPattern :: Pattern -> Pattern' Expr -> NAPs1 Expr -> TCMT IO (Pattern' Expr)
applyAPattern Pattern
p0 Pattern' Expr
p NAPs1 Expr
ps1 = do
  let ps :: [Item (NAPs1 Expr)]
ps = NAPs1 Expr -> [Item (NAPs1 Expr)]
forall l. IsList l => l -> [Item l]
List1.toList NAPs1 Expr
ps1
  Range -> Pattern' Expr -> Pattern' Expr
forall a. SetRange a => Range -> a -> a
setRange (Pattern -> Range
forall a. HasRange a => a -> Range
getRange Pattern
p0) (Pattern' Expr -> Pattern' Expr)
-> TCMT IO (Pattern' Expr) -> TCMT IO (Pattern' Expr)
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> do
    case Pattern' Expr
p of
      A.ConP ConPatInfo
i AmbiguousQName
x [NamedArg (Pattern' Expr)]
as        -> Pattern' Expr -> TCMT IO (Pattern' Expr)
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return (Pattern' Expr -> TCMT IO (Pattern' Expr))
-> Pattern' Expr -> TCMT IO (Pattern' Expr)
forall a b. (a -> b) -> a -> b
$ ConPatInfo
-> AmbiguousQName -> [NamedArg (Pattern' Expr)] -> Pattern' Expr
forall e. ConPatInfo -> AmbiguousQName -> NAPs e -> Pattern' e
A.ConP        ConPatInfo
i AmbiguousQName
x ([NamedArg (Pattern' Expr)]
as [NamedArg (Pattern' Expr)]
-> [NamedArg (Pattern' Expr)] -> [NamedArg (Pattern' Expr)]
forall a. [a] -> [a] -> [a]
++ [Item (NAPs1 Expr)]
[NamedArg (Pattern' Expr)]
ps)
      A.DefP PatInfo
i AmbiguousQName
x [NamedArg (Pattern' Expr)]
as        -> Pattern' Expr -> TCMT IO (Pattern' Expr)
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return (Pattern' Expr -> TCMT IO (Pattern' Expr))
-> Pattern' Expr -> TCMT IO (Pattern' Expr)
forall a b. (a -> b) -> a -> b
$ PatInfo
-> AmbiguousQName -> [NamedArg (Pattern' Expr)] -> Pattern' Expr
forall e. PatInfo -> AmbiguousQName -> NAPs e -> Pattern' e
A.DefP        PatInfo
i AmbiguousQName
x ([NamedArg (Pattern' Expr)]
as [NamedArg (Pattern' Expr)]
-> [NamedArg (Pattern' Expr)] -> [NamedArg (Pattern' Expr)]
forall a. [a] -> [a] -> [a]
++ [Item (NAPs1 Expr)]
[NamedArg (Pattern' Expr)]
ps)
      A.PatternSynP PatInfo
i AmbiguousQName
x [NamedArg (Pattern' Expr)]
as -> Pattern' Expr -> TCMT IO (Pattern' Expr)
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return (Pattern' Expr -> TCMT IO (Pattern' Expr))
-> Pattern' Expr -> TCMT IO (Pattern' Expr)
forall a b. (a -> b) -> a -> b
$ PatInfo
-> AmbiguousQName -> [NamedArg (Pattern' Expr)] -> Pattern' Expr
forall e. PatInfo -> AmbiguousQName -> NAPs e -> Pattern' e
A.PatternSynP PatInfo
i AmbiguousQName
x ([NamedArg (Pattern' Expr)]
as [NamedArg (Pattern' Expr)]
-> [NamedArg (Pattern' Expr)] -> [NamedArg (Pattern' Expr)]
forall a. [a] -> [a] -> [a]
++ [Item (NAPs1 Expr)]
[NamedArg (Pattern' Expr)]
ps)
      -- Dotted constructors are turned into "lazy" constructor patterns.
      A.DotP PatInfo
i (Ident QName
x)   -> QName -> ScopeM ResolvedName
resolveName QName
x ScopeM ResolvedName
-> (ResolvedName -> TCMT IO (Pattern' Expr))
-> TCMT IO (Pattern' Expr)
forall a b. TCMT IO a -> (a -> TCMT IO b) -> TCMT IO b
forall (m :: * -> *) a b. Monad m => m a -> (a -> m b) -> m b
>>= \case
        ConstructorName Set1 Induction
_ List1 AbstractName
ds -> do
          let cpi :: ConPatInfo
cpi = ConOrigin -> PatInfo -> ConPatLazy -> ConPatInfo
ConPatInfo ConOrigin
ConOCon PatInfo
i ConPatLazy
ConPatLazy
              c :: AmbiguousQName
c   = List1 QName -> AmbiguousQName
AmbQ ((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
ds)
          Pattern' Expr -> TCMT IO (Pattern' Expr)
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return (Pattern' Expr -> TCMT IO (Pattern' Expr))
-> Pattern' Expr -> TCMT IO (Pattern' Expr)
forall a b. (a -> b) -> a -> b
$ ConPatInfo
-> AmbiguousQName -> [NamedArg (Pattern' Expr)] -> Pattern' Expr
forall e. ConPatInfo -> AmbiguousQName -> NAPs e -> Pattern' e
A.ConP ConPatInfo
cpi AmbiguousQName
c [Item (NAPs1 Expr)]
[NamedArg (Pattern' Expr)]
ps
        ResolvedName
_ -> TCMT IO (Pattern' Expr)
failure
      A.DotP{}    -> TCMT IO (Pattern' Expr)
failure
      A.VarP{}    -> TCMT IO (Pattern' Expr)
failure
      A.ProjP{}   -> TCMT IO (Pattern' Expr)
failure
      A.WildP{}   -> TCMT IO (Pattern' Expr)
failure
      A.AsP{}     -> TCMT IO (Pattern' Expr)
failure
      A.AbsurdP{} -> TCMT IO (Pattern' Expr)
failure
      A.LitP{}    -> TCMT IO (Pattern' Expr)
failure
      A.RecP{}    -> TCMT IO (Pattern' Expr)
failure
      A.EqualP{}  -> TCMT IO (Pattern' Expr)
failure
      A.WithP{}   -> TCMT IO (Pattern' Expr)
failure
  where
    failure :: TCMT IO (Pattern' Expr)
failure = TypeError -> TCMT IO (Pattern' Expr)
forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
TypeError -> m a
typeError (TypeError -> TCMT IO (Pattern' Expr))
-> TypeError -> TCMT IO (Pattern' Expr)
forall a b. (a -> b) -> a -> b
$ Pattern -> TypeError
InvalidPattern Pattern
p0

-- | Throw-away wrapper type for pattern translation.
data WithHidingInfo a = WithHidingInfo Hiding a

propagateHidingInfo :: NamedArg a -> NamedArg (WithHidingInfo a)
propagateHidingInfo :: forall a. NamedArg a -> NamedArg (WithHidingInfo a)
propagateHidingInfo NamedArg a
a = (Named NamedName a -> Named_ (WithHidingInfo a))
-> NamedArg a -> Arg (Named_ (WithHidingInfo a))
forall a b. (a -> b) -> Arg a -> Arg b
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
fmap ((a -> WithHidingInfo a)
-> Named NamedName a -> Named_ (WithHidingInfo a)
forall a b. (a -> b) -> Named NamedName a -> Named NamedName b
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
fmap ((a -> WithHidingInfo a)
 -> Named NamedName a -> Named_ (WithHidingInfo a))
-> (a -> WithHidingInfo a)
-> Named NamedName a
-> Named_ (WithHidingInfo a)
forall a b. (a -> b) -> a -> b
$ Hiding -> a -> WithHidingInfo a
forall a. Hiding -> a -> WithHidingInfo a
WithHidingInfo (Hiding -> a -> WithHidingInfo a)
-> Hiding -> a -> WithHidingInfo a
forall a b. (a -> b) -> a -> b
$ NamedArg a -> Hiding
forall a. LensHiding a => a -> Hiding
getHiding NamedArg a
a) NamedArg a
a

-- | Hiding info is only used for pattern variables.
instance ToAbstract (WithHidingInfo C.Pattern) where
    type AbsOfCon (WithHidingInfo C.Pattern) = A.Pattern' C.Expr

    toAbstract :: WithHidingInfo Pattern
-> ScopeM (AbsOfCon (WithHidingInfo Pattern))
toAbstract (WithHidingInfo Hiding
h (C.IdentP Bool
canBeConstructor QName
x)) =
      Bool
-> DisplayLHS
-> Hiding
-> QName
-> Maybe (Set1 Name)
-> TCMT IO (Pattern' Expr)
resolvePatternIdentifier Bool
canBeConstructor DisplayLHS
NoDisplayLHS Hiding
h QName
x Maybe (Set1 Name)
forall a. Maybe a
Nothing

    toAbstract (WithHidingInfo Hiding
_ Pattern
p) = Pattern -> ScopeM (AbsOfCon Pattern)
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract Pattern
p

-- | Scope check a 'C.Pattern' (of possibly a 'C.DisplayForm').
--
data CPattern = CPattern
  DisplayLHS
    -- ^ Are we checking a 'C.DisplayForm'?
  C.Pattern
    -- ^ The pattern to scope-check.

-- | Scope check a 'C.Pattern' not belonging to a 'C.DisplayForm'.
--
instance ToAbstract C.Pattern where
  type AbsOfCon C.Pattern = A.Pattern' C.Expr

  toAbstract :: Pattern -> ScopeM (AbsOfCon Pattern)
toAbstract = CPattern -> TCMT IO (Pattern' Expr)
CPattern -> ScopeM (AbsOfCon CPattern)
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract (CPattern -> TCMT IO (Pattern' Expr))
-> (Pattern -> CPattern) -> Pattern -> TCMT IO (Pattern' Expr)
forall b c a. (b -> c) -> (a -> b) -> a -> c
. DisplayLHS -> Pattern -> CPattern
CPattern DisplayLHS
NoDisplayLHS

instance ToAbstract CPattern where
  type AbsOfCon CPattern = A.Pattern' C.Expr

  toAbstract :: CPattern -> ScopeM (AbsOfCon CPattern)
toAbstract (CPattern DisplayLHS
displayLhs Pattern
p0) = case Pattern
p0 of

    C.IdentP Bool
canBeConstructor QName
x ->
      Bool
-> DisplayLHS
-> Hiding
-> QName
-> Maybe (Set1 Name)
-> TCMT IO (Pattern' Expr)
resolvePatternIdentifier Bool
canBeConstructor DisplayLHS
displayLhs Hiding
forall a. Null a => a
empty QName
x Maybe (Set1 Name)
forall a. Maybe a
Nothing

    QuoteP Range
_r ->
      TypeError -> ScopeM (AbsOfCon CPattern)
forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
TypeError -> m a
typeError (TypeError -> ScopeM (AbsOfCon CPattern))
-> TypeError -> ScopeM (AbsOfCon CPattern)
forall a b. (a -> b) -> a -> b
$ CannotQuote -> TypeError
CannotQuote CannotQuote
CannotQuoteNothing

    AppP (QuoteP Range
_) NamedArg Pattern
p
      | IdentP Bool
_ QName
x <- NamedArg Pattern -> Pattern
forall a. NamedArg a -> a
namedArg NamedArg Pattern
p -> do
          if NamedArg Pattern -> Bool
forall a. LensHiding a => a -> Bool
visible NamedArg Pattern
p then do
            e <- OldQName -> ScopeM (AbsOfCon OldQName)
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract (QName -> Maybe (Set1 Name) -> OldQName
OldQName QName
x Maybe (Set1 Name)
forall a. Maybe a
Nothing)
            A.LitP (PatRange $ getRange x) . LitQName <$> quotedName e
          else TypeError -> ScopeM (AbsOfCon CPattern)
forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
TypeError -> m a
typeError (TypeError -> ScopeM (AbsOfCon CPattern))
-> TypeError -> ScopeM (AbsOfCon CPattern)
forall a b. (a -> b) -> a -> b
$ CannotQuote -> TypeError
CannotQuote CannotQuote
CannotQuoteHidden
      | Bool
otherwise -> TypeError -> ScopeM (AbsOfCon CPattern)
forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
TypeError -> m a
typeError (TypeError -> ScopeM (AbsOfCon CPattern))
-> TypeError -> ScopeM (AbsOfCon CPattern)
forall a b. (a -> b) -> a -> b
$ CannotQuote -> TypeError
CannotQuote (CannotQuote -> TypeError) -> CannotQuote -> TypeError
forall a b. (a -> b) -> a -> b
$ NamedArg Pattern -> CannotQuote
CannotQuotePattern NamedArg Pattern
p

    AppP Pattern
p NamedArg Pattern
q -> do
        [Char] -> Int -> [Char] -> TCMT IO ()
forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> [Char] -> m ()
reportSLn [Char]
"scope.pat" Int
50 ([Char] -> TCMT IO ()) -> [Char] -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ [Char]
"distributeDots before = " [Char] -> [Char] -> [Char]
forall a. [a] -> [a] -> [a]
++ Pattern -> [Char]
forall a. Show a => a -> [Char]
show Pattern
p
        p <- Pattern -> TCMT IO Pattern
distributeDots Pattern
p
        reportSLn "scope.pat" 50 $ "distributeDots after  = " ++ show p
        p' <- toAbstract (wrap p)
        -- Remember hiding info in argument to propagate to 'PatternBound'.
        q' <- ifThenElse displayLhs
          {-then-} (toAbstract $ (fmap . fmap) wrap q)
          {-else-} (toAbstract $ propagateHidingInfo q)
        applyAPattern p0 p' $ singleton q'

        where
            distributeDots :: C.Pattern -> ScopeM C.Pattern
            distributeDots :: Pattern -> TCMT IO Pattern
distributeDots p :: Pattern
p@(C.DotP KwRange
kwr Range
r Expr
e) = KwRange -> Range -> Expr -> TCMT IO Pattern
distributeDotsExpr KwRange
kwr Range
r Expr
e
            distributeDots Pattern
p = Pattern -> TCMT IO Pattern
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return Pattern
p

            distributeDotsExpr :: KwRange -> Range -> C.Expr -> ScopeM C.Pattern
            distributeDotsExpr :: KwRange -> Range -> Expr -> TCMT IO Pattern
distributeDotsExpr KwRange
kwr Range
r Expr
e = Expr -> TCMT IO Expr
parseRawApp Expr
e TCMT IO Expr -> (Expr -> TCMT IO Pattern) -> TCMT IO Pattern
forall a b. TCMT IO a -> (a -> TCMT IO b) -> TCMT IO b
forall (m :: * -> *) a b. Monad m => m a -> (a -> m b) -> m b
>>= \case
              C.App Range
r Expr
e NamedArg Expr
a     ->
                Pattern -> NamedArg Pattern -> Pattern
AppP (Pattern -> NamedArg Pattern -> Pattern)
-> TCMT IO Pattern -> TCMT IO (NamedArg Pattern -> Pattern)
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> KwRange -> Range -> Expr -> TCMT IO Pattern
distributeDotsExpr KwRange
forall a. Null a => a
empty Range
r Expr
e
                     TCMT IO (NamedArg Pattern -> Pattern)
-> TCMT IO (NamedArg Pattern) -> TCMT IO Pattern
forall a b. TCMT IO (a -> b) -> TCMT IO a -> TCMT IO b
forall (f :: * -> *) a b. Applicative f => f (a -> b) -> f a -> f b
<*> ((Named NamedName Expr -> TCMT IO (Named NamedName Pattern))
-> NamedArg Expr -> TCMT IO (NamedArg Pattern)
forall (t :: * -> *) (f :: * -> *) a b.
(Traversable t, Applicative f) =>
(a -> f b) -> t a -> f (t b)
forall (f :: * -> *) a b.
Applicative f =>
(a -> f b) -> Arg a -> f (Arg b)
traverse ((Named NamedName Expr -> TCMT IO (Named NamedName Pattern))
 -> NamedArg Expr -> TCMT IO (NamedArg Pattern))
-> ((Expr -> TCMT IO Pattern)
    -> Named NamedName Expr -> TCMT IO (Named NamedName Pattern))
-> (Expr -> TCMT IO Pattern)
-> NamedArg Expr
-> TCMT IO (NamedArg Pattern)
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (Expr -> TCMT IO Pattern)
-> Named NamedName Expr -> TCMT IO (Named NamedName Pattern)
forall (t :: * -> *) (f :: * -> *) a b.
(Traversable t, Applicative f) =>
(a -> f b) -> t a -> f (t b)
forall (f :: * -> *) a b.
Applicative f =>
(a -> f b) -> Named NamedName a -> f (Named NamedName b)
traverse) (KwRange -> Range -> Expr -> TCMT IO Pattern
distributeDotsExpr KwRange
forall a. Null a => a
empty Range
r) NamedArg Expr
a
              OpApp Range
r QName
q Set1 Name
ns OpAppArgs
as ->
                case ((Arg (Named NamedName (MaybePlaceholder (OpApp Expr)))
 -> Maybe (NamedArg Expr))
-> OpAppArgs -> Maybe (NonEmpty (NamedArg Expr))
forall (t :: * -> *) (f :: * -> *) a b.
(Traversable t, Applicative f) =>
(a -> f b) -> t a -> f (t b)
forall (f :: * -> *) a b.
Applicative f =>
(a -> f b) -> NonEmpty a -> f (NonEmpty b)
traverse ((Arg (Named NamedName (MaybePlaceholder (OpApp Expr)))
  -> Maybe (NamedArg Expr))
 -> OpAppArgs -> Maybe (NonEmpty (NamedArg Expr)))
-> ((MaybePlaceholder (OpApp Expr) -> Maybe Expr)
    -> Arg (Named NamedName (MaybePlaceholder (OpApp Expr)))
    -> Maybe (NamedArg Expr))
-> (MaybePlaceholder (OpApp Expr) -> Maybe Expr)
-> OpAppArgs
-> Maybe (NonEmpty (NamedArg Expr))
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (Named NamedName (MaybePlaceholder (OpApp Expr))
 -> Maybe (Named NamedName Expr))
-> Arg (Named NamedName (MaybePlaceholder (OpApp Expr)))
-> Maybe (NamedArg Expr)
forall (t :: * -> *) (f :: * -> *) a b.
(Traversable t, Applicative f) =>
(a -> f b) -> t a -> f (t b)
forall (f :: * -> *) a b.
Applicative f =>
(a -> f b) -> Arg a -> f (Arg b)
traverse ((Named NamedName (MaybePlaceholder (OpApp Expr))
  -> Maybe (Named NamedName Expr))
 -> Arg (Named NamedName (MaybePlaceholder (OpApp Expr)))
 -> Maybe (NamedArg Expr))
-> ((MaybePlaceholder (OpApp Expr) -> Maybe Expr)
    -> Named NamedName (MaybePlaceholder (OpApp Expr))
    -> Maybe (Named NamedName Expr))
-> (MaybePlaceholder (OpApp Expr) -> Maybe Expr)
-> Arg (Named NamedName (MaybePlaceholder (OpApp Expr)))
-> Maybe (NamedArg Expr)
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (MaybePlaceholder (OpApp Expr) -> Maybe Expr)
-> Named NamedName (MaybePlaceholder (OpApp Expr))
-> Maybe (Named NamedName Expr)
forall (t :: * -> *) (f :: * -> *) a b.
(Traversable t, Applicative f) =>
(a -> f b) -> t a -> f (t b)
forall (f :: * -> *) a b.
Applicative f =>
(a -> f b) -> Named NamedName a -> f (Named NamedName b)
traverse) MaybePlaceholder (OpApp Expr) -> Maybe Expr
forall a. MaybePlaceholder (OpApp a) -> Maybe a
fromNoPlaceholder OpAppArgs
as of
                  Just NonEmpty (NamedArg Expr)
as -> Range -> QName -> Set1 Name -> List1 (NamedArg Pattern) -> Pattern
OpAppP Range
r QName
q Set1 Name
ns (List1 (NamedArg Pattern) -> Pattern)
-> TCMT IO (List1 (NamedArg Pattern)) -> TCMT IO Pattern
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$>
                    ((NamedArg Expr -> TCMT IO (NamedArg Pattern))
-> NonEmpty (NamedArg Expr) -> TCMT IO (List1 (NamedArg Pattern))
forall (t :: * -> *) (f :: * -> *) a b.
(Traversable t, Applicative f) =>
(a -> f b) -> t a -> f (t b)
forall (f :: * -> *) a b.
Applicative f =>
(a -> f b) -> NonEmpty a -> f (NonEmpty b)
traverse ((NamedArg Expr -> TCMT IO (NamedArg Pattern))
 -> NonEmpty (NamedArg Expr) -> TCMT IO (List1 (NamedArg Pattern)))
-> ((Expr -> TCMT IO Pattern)
    -> NamedArg Expr -> TCMT IO (NamedArg Pattern))
-> (Expr -> TCMT IO Pattern)
-> NonEmpty (NamedArg Expr)
-> TCMT IO (List1 (NamedArg Pattern))
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (Named NamedName Expr -> TCMT IO (Named NamedName Pattern))
-> NamedArg Expr -> TCMT IO (NamedArg Pattern)
forall (t :: * -> *) (f :: * -> *) a b.
(Traversable t, Applicative f) =>
(a -> f b) -> t a -> f (t b)
forall (f :: * -> *) a b.
Applicative f =>
(a -> f b) -> Arg a -> f (Arg b)
traverse ((Named NamedName Expr -> TCMT IO (Named NamedName Pattern))
 -> NamedArg Expr -> TCMT IO (NamedArg Pattern))
-> ((Expr -> TCMT IO Pattern)
    -> Named NamedName Expr -> TCMT IO (Named NamedName Pattern))
-> (Expr -> TCMT IO Pattern)
-> NamedArg Expr
-> TCMT IO (NamedArg Pattern)
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (Expr -> TCMT IO Pattern)
-> Named NamedName Expr -> TCMT IO (Named NamedName Pattern)
forall (t :: * -> *) (f :: * -> *) a b.
(Traversable t, Applicative f) =>
(a -> f b) -> t a -> f (t b)
forall (f :: * -> *) a b.
Applicative f =>
(a -> f b) -> Named NamedName a -> f (Named NamedName b)
traverse) (KwRange -> Range -> Expr -> TCMT IO Pattern
distributeDotsExpr KwRange
forall a. Null a => a
empty Range
r) NonEmpty (NamedArg Expr)
as
                  Maybe (NonEmpty (NamedArg Expr))
Nothing -> Pattern -> TCMT IO Pattern
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return (Pattern -> TCMT IO Pattern) -> Pattern -> TCMT IO Pattern
forall a b. (a -> b) -> a -> b
$ KwRange -> Range -> Expr -> Pattern
C.DotP KwRange
forall a. Null a => a
empty Range
r Expr
e
              Paren Range
r Expr
e -> Range -> KillRangeT Pattern
ParenP Range
r KillRangeT Pattern -> TCMT IO Pattern -> TCMT IO Pattern
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> KwRange -> Range -> Expr -> TCMT IO Pattern
distributeDotsExpr KwRange
forall a. Null a => a
empty Range
r Expr
e
              Expr
_ -> Pattern -> TCMT IO Pattern
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return (Pattern -> TCMT IO Pattern) -> Pattern -> TCMT IO Pattern
forall a b. (a -> b) -> a -> b
$ KwRange -> Range -> Expr -> Pattern
C.DotP KwRange
kwr Range
r Expr
e

            fromNoPlaceholder :: MaybePlaceholder (OpApp a) -> Maybe a
            fromNoPlaceholder :: forall a. MaybePlaceholder (OpApp a) -> Maybe a
fromNoPlaceholder (NoPlaceholder Maybe PositionInName
_ (Ordinary a
e)) = a -> Maybe a
forall a. a -> Maybe a
Just a
e
            fromNoPlaceholder MaybePlaceholder (OpApp a)
_ = Maybe a
forall a. Maybe a
Nothing

            parseRawApp :: C.Expr -> ScopeM C.Expr
            parseRawApp :: Expr -> TCMT IO Expr
parseRawApp (RawApp Range
r List2 Expr
es) = List2 Expr -> TCMT IO Expr
parseApplication List2 Expr
es
            parseRawApp Expr
e             = Expr -> TCMT IO Expr
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return Expr
e

    OpAppP Range
r QName
op Set1 Name
ns List1 (NamedArg Pattern)
ps -> do
        [Char] -> Int -> [Char] -> TCMT IO ()
forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> [Char] -> m ()
reportSLn [Char]
"scope.pat" Int
60 ([Char] -> TCMT IO ()) -> [Char] -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ [Char]
"ConcreteToAbstract.toAbstract OpAppP{}: " [Char] -> [Char] -> [Char]
forall a. [a] -> [a] -> [a]
++ Pattern -> [Char]
forall a. Show a => a -> [Char]
show Pattern
p0
        p  <- Bool
-> DisplayLHS
-> Hiding
-> QName
-> Maybe (Set1 Name)
-> TCMT IO (Pattern' Expr)
resolvePatternIdentifier Bool
True DisplayLHS
displayLhs Hiding
forall a. Null a => a
empty QName
op (Set1 Name -> Maybe (Set1 Name)
forall a. a -> Maybe a
Just Set1 Name
ns)
        -- Remember hiding info in arguments to propagate to 'PatternBound'.
        ps <- ifThenElse displayLhs
          {-then-} (toAbstract $ (fmap . fmap . fmap) wrap ps)
          {-else-} (toAbstract $ fmap propagateHidingInfo ps)
        applyAPattern p0 p ps

    EllipsisP Range
_ Maybe Pattern
mp -> TCMT IO (Pattern' Expr)
-> (Pattern -> TCMT IO (Pattern' Expr))
-> Maybe Pattern
-> TCMT IO (Pattern' Expr)
forall b a. b -> (a -> b) -> Maybe a -> b
maybe TCMT IO (Pattern' Expr)
forall a. HasCallStack => a
__IMPOSSIBLE__ Pattern -> TCMT IO (Pattern' Expr)
Pattern -> ScopeM (AbsOfCon Pattern)
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract Maybe Pattern
mp  -- Not in DISPLAY pragma

    -- Removed when parsing
    HiddenP Range
_ Named NamedName Pattern
_    -> TCMT IO (Pattern' Expr)
ScopeM (AbsOfCon CPattern)
forall a. HasCallStack => a
__IMPOSSIBLE__
    InstanceP Range
_ Named NamedName Pattern
_  -> TCMT IO (Pattern' Expr)
ScopeM (AbsOfCon CPattern)
forall a. HasCallStack => a
__IMPOSSIBLE__
    RawAppP Range
_ List2 Pattern
_    -> TCMT IO (Pattern' Expr)
ScopeM (AbsOfCon CPattern)
forall a. HasCallStack => a
__IMPOSSIBLE__

    C.WildP Range
r      -> AbsOfCon CPattern -> ScopeM (AbsOfCon CPattern)
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return (AbsOfCon CPattern -> ScopeM (AbsOfCon CPattern))
-> AbsOfCon CPattern -> ScopeM (AbsOfCon CPattern)
forall a b. (a -> b) -> a -> b
$ PatInfo -> Pattern' Expr
forall e. PatInfo -> Pattern' e
A.WildP (PatInfo -> Pattern' Expr) -> PatInfo -> Pattern' Expr
forall a b. (a -> b) -> a -> b
$ Range -> PatInfo
PatRange Range
r
    -- Andreas, 2015-05-28 futile attempt to fix issue 819: repeated variable on lhs "_"
    -- toAbstract p@(C.WildP r)    = A.VarP <$> freshName r "_"
    C.ParenP Range
_ Pattern
p   -> CPattern -> ScopeM (AbsOfCon CPattern)
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract (CPattern -> ScopeM (AbsOfCon CPattern))
-> CPattern -> ScopeM (AbsOfCon CPattern)
forall a b. (a -> b) -> a -> b
$ Pattern -> CPattern
wrap Pattern
p  -- Andreas, 2024-09-27 not impossible
    C.LitP Range
r Literal
l     -> Range -> ScopeM (AbsOfCon CPattern) -> ScopeM (AbsOfCon CPattern)
forall (m :: * -> *) x a.
(MonadTrace m, HasRange x) =>
x -> m a -> m a
setCurrentRange Range
r (ScopeM (AbsOfCon CPattern) -> ScopeM (AbsOfCon CPattern))
-> ScopeM (AbsOfCon CPattern) -> ScopeM (AbsOfCon CPattern)
forall a b. (a -> b) -> a -> b
$ PatInfo -> Literal -> Pattern' Expr
forall e. PatInfo -> Literal -> Pattern' e
A.LitP (Range -> PatInfo
PatRange Range
r) Literal
l Pattern' Expr -> TCMT IO () -> TCMT IO (Pattern' Expr)
forall a b. a -> TCMT IO b -> TCMT IO a
forall (f :: * -> *) a b. Functor f => a -> f b -> f a
<$ Literal -> TCMT IO ()
checkLiteral Literal
l

    C.AsP Range
r Name
x Pattern
p -> do
        -- Andreas, 2018-06-30, issue #3147: as-variables can be non-linear a priori!
        -- x <- toAbstract (NewName PatternBound x)
        -- Andreas, 2020-05-01, issue #4631: as-variables should not shadow constructors.
        -- x <- bindPatternVariable x
      PatName -> ScopeM (AbsOfCon PatName)
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract (QName -> Maybe (Set1 Name) -> Hiding -> DisplayLHS -> PatName
PatName (Name -> QName
C.QName Name
x) Maybe (Set1 Name)
forall a. Maybe a
Nothing Hiding
forall a. Null a => a
empty DisplayLHS
NoDisplayLHS) TCMT IO APatName
-> (APatName -> TCMT IO (Pattern' Expr)) -> TCMT IO (Pattern' Expr)
forall a b. TCMT IO a -> (a -> TCMT IO b) -> TCMT IO b
forall (m :: * -> *) a b. Monad m => m a -> (a -> m b) -> m b
>>= \case
        VarPatName Name
x        -> PatInfo -> BindName -> Pattern' Expr -> Pattern' Expr
forall e. PatInfo -> BindName -> Pattern' e -> Pattern' e
A.AsP (Range -> PatInfo
PatRange Range
r) (Name -> BindName
A.mkBindName Name
x) (Pattern' Expr -> Pattern' Expr)
-> TCMT IO (Pattern' Expr) -> TCMT IO (Pattern' Expr)
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> CPattern -> ScopeM (AbsOfCon CPattern)
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract (Pattern -> CPattern
wrap Pattern
p)
        ConPatName{}        -> ConstructorOrPatternSynonym -> TCMT IO (Pattern' Expr)
ignoreAsPat ConstructorOrPatternSynonym
IsConstructor
        PatternSynPatName{} -> ConstructorOrPatternSynonym -> TCMT IO (Pattern' Expr)
ignoreAsPat ConstructorOrPatternSynonym
IsPatternSynonym
        DefPatName{}        -> TCMT IO (Pattern' Expr)
forall a. HasCallStack => a
__IMPOSSIBLE__  -- because of @False@ in @PatName@
      where
      -- An @-bound name which shadows a constructor is illegal and becomes dead code.
      ignoreAsPat :: ConstructorOrPatternSynonym -> TCMT IO (Pattern' Expr)
ignoreAsPat ConstructorOrPatternSynonym
b = do
        Name -> TCMT IO () -> TCMT IO ()
forall (m :: * -> *) x a.
(MonadTrace m, HasRange x) =>
x -> m a -> m a
setCurrentRange Name
x (TCMT IO () -> TCMT IO ()) -> TCMT IO () -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ Warning -> TCMT IO ()
forall (m :: * -> *).
(HasCallStack, MonadWarning m) =>
Warning -> m ()
warning (Warning -> TCMT IO ()) -> Warning -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ ConstructorOrPatternSynonym -> Warning
AsPatternShadowsConstructorOrPatternSynonym ConstructorOrPatternSynonym
b
        CPattern -> ScopeM (AbsOfCon CPattern)
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract (CPattern -> ScopeM (AbsOfCon CPattern))
-> CPattern -> ScopeM (AbsOfCon CPattern)
forall a b. (a -> b) -> a -> b
$ Pattern -> CPattern
wrap Pattern
p

    C.EqualP Range
r List1 (Expr, Expr)
es -> AbsOfCon CPattern -> ScopeM (AbsOfCon CPattern)
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return (AbsOfCon CPattern -> ScopeM (AbsOfCon CPattern))
-> AbsOfCon CPattern -> ScopeM (AbsOfCon CPattern)
forall a b. (a -> b) -> a -> b
$ PatInfo -> List1 (Expr, Expr) -> Pattern' Expr
forall e. PatInfo -> List1 (e, e) -> Pattern' e
A.EqualP (Range -> PatInfo
PatRange Range
r) List1 (Expr, Expr)
es

    -- We have to do dot patterns at the end since they can
    -- refer to the variables bound by the other patterns.
    C.DotP KwRange
_kwr Range
r Expr
e -> do
      let fallback :: TCMT IO (Pattern' Expr)
fallback = Pattern' Expr -> TCMT IO (Pattern' Expr)
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return (Pattern' Expr -> TCMT IO (Pattern' Expr))
-> Pattern' Expr -> TCMT IO (Pattern' Expr)
forall a b. (a -> b) -> a -> b
$ PatInfo -> Expr -> Pattern' Expr
forall e. PatInfo -> e -> Pattern' e
A.DotP (Range -> PatInfo
PatRange Range
r) Expr
e
      case Expr
e of
        C.Ident QName
x -> QName -> ScopeM ResolvedName
resolveName QName
x ScopeM ResolvedName
-> (ResolvedName -> TCMT IO (Pattern' Expr))
-> TCMT IO (Pattern' Expr)
forall a b. TCMT IO a -> (a -> TCMT IO b) -> TCMT IO b
forall (m :: * -> *) a b. Monad m => m a -> (a -> m b) -> m b
>>= \case
          -- Andreas, 2018-06-19, #3130
          -- We interpret .x as postfix projection if x is a field name in scope
          FieldName List1 AbstractName
xs -> Pattern' Expr -> TCMT IO (Pattern' Expr)
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return (Pattern' Expr -> TCMT IO (Pattern' Expr))
-> Pattern' Expr -> TCMT IO (Pattern' Expr)
forall a b. (a -> b) -> a -> b
$ PatInfo -> ProjOrigin -> AmbiguousQName -> Pattern' Expr
forall e. PatInfo -> ProjOrigin -> AmbiguousQName -> Pattern' e
A.ProjP (Range -> PatInfo
PatRange Range
r) ProjOrigin
ProjPostfix (AmbiguousQName -> Pattern' Expr)
-> AmbiguousQName -> Pattern' Expr
forall a b. (a -> b) -> a -> b
$ List1 QName -> AmbiguousQName
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
_ -> TCMT IO (Pattern' Expr)
fallback
        Expr
_ -> TCMT IO (Pattern' Expr)
ScopeM (AbsOfCon CPattern)
fallback

    C.AbsurdP Range
r -> AbsOfCon CPattern -> ScopeM (AbsOfCon CPattern)
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return (AbsOfCon CPattern -> ScopeM (AbsOfCon CPattern))
-> AbsOfCon CPattern -> ScopeM (AbsOfCon CPattern)
forall a b. (a -> b) -> a -> b
$ PatInfo -> Pattern' Expr
forall e. PatInfo -> Pattern' e
A.AbsurdP (PatInfo -> Pattern' Expr) -> PatInfo -> Pattern' Expr
forall a b. (a -> b) -> a -> b
$ Range -> PatInfo
PatRange Range
r
    C.RecP Range
r [FieldAssignment' Pattern]
fs -> ConPatInfo -> [FieldAssignment' (Pattern' Expr)] -> Pattern' Expr
forall e.
ConPatInfo -> [FieldAssignment' (Pattern' e)] -> Pattern' e
A.RecP (ConOrigin -> PatInfo -> ConPatLazy -> ConPatInfo
ConPatInfo ConOrigin
ConORec (Range -> PatInfo
PatRange Range
r) ConPatLazy
ConPatEager) ([FieldAssignment' (Pattern' Expr)] -> Pattern' Expr)
-> TCMT IO [FieldAssignment' (Pattern' Expr)]
-> TCMT IO (Pattern' Expr)
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> (FieldAssignment' Pattern
 -> TCMT IO (FieldAssignment' (Pattern' Expr)))
-> [FieldAssignment' Pattern]
-> TCMT IO [FieldAssignment' (Pattern' Expr)]
forall (t :: * -> *) (m :: * -> *) a b.
(Traversable t, Monad m) =>
(a -> m b) -> t a -> m (t b)
forall (m :: * -> *) a b. Monad m => (a -> m b) -> [a] -> m [b]
mapM ((Pattern -> TCMT IO (Pattern' Expr))
-> FieldAssignment' Pattern
-> TCMT IO (FieldAssignment' (Pattern' Expr))
forall (t :: * -> *) (f :: * -> *) a b.
(Traversable t, Applicative f) =>
(a -> f b) -> t a -> f (t b)
forall (f :: * -> *) a b.
Applicative f =>
(a -> f b) -> FieldAssignment' a -> f (FieldAssignment' b)
traverse ((Pattern -> TCMT IO (Pattern' Expr))
 -> FieldAssignment' Pattern
 -> TCMT IO (FieldAssignment' (Pattern' Expr)))
-> (Pattern -> TCMT IO (Pattern' Expr))
-> FieldAssignment' Pattern
-> TCMT IO (FieldAssignment' (Pattern' Expr))
forall a b. (a -> b) -> a -> b
$ CPattern -> TCMT IO (Pattern' Expr)
CPattern -> ScopeM (AbsOfCon CPattern)
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract (CPattern -> TCMT IO (Pattern' Expr))
-> (Pattern -> CPattern) -> Pattern -> TCMT IO (Pattern' Expr)
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Pattern -> CPattern
wrap) [FieldAssignment' Pattern]
fs
    C.WithP Range
r Pattern
p -> PatInfo -> Pattern' Expr -> Pattern' Expr
forall e. PatInfo -> Pattern' e -> Pattern' e
A.WithP (Range -> PatInfo
PatRange Range
r) (Pattern' Expr -> Pattern' Expr)
-> TCMT IO (Pattern' Expr) -> TCMT IO (Pattern' Expr)
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> Pattern -> ScopeM (AbsOfCon Pattern)
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract Pattern
p  -- not in DISPLAY pragma

    where
      -- Pass on @displayLhs@ context
      wrap :: Pattern -> CPattern
wrap = DisplayLHS -> Pattern -> CPattern
CPattern DisplayLHS
displayLhs

-- | An argument @OpApp C.Expr@ to an operator can have binders,
--   in case the operator is some @syntax@-notation.
--   For these binders, we have to create lambda-abstractions.
toAbstractOpArg :: Precedence -> OpApp C.Expr -> ScopeM A.Expr
toAbstractOpArg :: Precedence -> OpApp Expr -> ScopeM Expr
toAbstractOpArg Precedence
ctx (Ordinary Expr
e)                 = Precedence -> Expr -> ScopeM (AbsOfCon Expr)
forall c. ToAbstract c => Precedence -> c -> ScopeM (AbsOfCon c)
toAbstractCtx Precedence
ctx Expr
e
toAbstractOpArg Precedence
ctx (SyntaxBindingLambda Range
r NonEmpty (LamBinding' TypedBinding)
bs Expr
e) = Range
-> NonEmpty (LamBinding' TypedBinding)
-> Expr
-> Precedence
-> ScopeM Expr
toAbstractLam Range
r NonEmpty (LamBinding' TypedBinding)
bs Expr
e Precedence
ctx

-- | Turn an operator application into abstract syntax. Make sure to
-- record the right precedences for the various arguments.
toAbstractOpApp :: C.QName -> Set1 A.Name -> OpAppArgs -> ScopeM A.Expr
toAbstractOpApp :: QName -> Set1 Name -> OpAppArgs -> ScopeM Expr
toAbstractOpApp QName
op Set1 Name
ns OpAppArgs
es = do
    -- Replace placeholders with bound variables.
    (binders, es) <- OpAppArgs0 Expr
-> ScopeM ([LamBinding], [NamedArg (Either Expr (OpApp Expr))])
forall e.
OpAppArgs0 e
-> ScopeM ([LamBinding], [NamedArg (Either Expr (OpApp e))])
replacePlaceholders (OpAppArgs0 Expr
 -> ScopeM ([LamBinding], [NamedArg (Either Expr (OpApp Expr))]))
-> OpAppArgs0 Expr
-> ScopeM ([LamBinding], [NamedArg (Either Expr (OpApp Expr))])
forall a b. (a -> b) -> a -> b
$ OpAppArgs -> [Item OpAppArgs]
forall l. IsList l => l -> [Item l]
List1.toList OpAppArgs
es
    -- Get the notation for the operator.
    nota <- getNotation op ns
    let parts = NewNotation -> Notation
notation NewNotation
nota
    -- We can throw away the @VarPart@s, since binders
    -- have been preprocessed into @OpApp C.Expr@.
    let nonBindingParts = (NotationPart -> Bool) -> Notation -> Notation
forall a. (a -> Bool) -> [a] -> [a]
filter (Bool -> Bool
not (Bool -> Bool) -> (NotationPart -> Bool) -> NotationPart -> Bool
forall b c a. (b -> c) -> (a -> b) -> a -> c
. NotationPart -> Bool
isBinder) Notation
parts
    -- We should be left with as many holes as we have been given args @es@.
    -- If not, crash.
    unless (length (filter isAHole nonBindingParts) == length es) __IMPOSSIBLE__
    -- Translate operator and its arguments (each in the right context).
    op <- toAbstract (OldQName op (Just ns))
    es <- left (notaFixity nota) nonBindingParts es
    -- Prepend the generated section binders (if any).
    let body = (Expr -> (ParenPreference, NamedArg Expr) -> Expr)
-> Expr -> [(ParenPreference, NamedArg Expr)] -> Expr
forall b a. (b -> a -> b) -> b -> [a] -> b
forall (t :: * -> *) b a.
Foldable t =>
(b -> a -> b) -> b -> t a -> b
List.foldl' Expr -> (ParenPreference, NamedArg Expr) -> Expr
app Expr
op [(ParenPreference, NamedArg Expr)]
es
    return $ foldr (A.Lam (ExprRange (getRange body))) body binders
  where
    -- Build an application in the abstract syntax, with correct Range.
    app :: Expr -> (ParenPreference, NamedArg Expr) -> Expr
app Expr
e (ParenPreference
pref, NamedArg Expr
arg) = AppInfo -> Expr -> NamedArg Expr -> Expr
A.App AppInfo
info Expr
e NamedArg Expr
arg
      where info :: AppInfo
info = (Range -> AppInfo
defaultAppInfo Range
r) { appOrigin = getOrigin arg
                                      , appParens = pref }
            r :: Range
r = Expr -> NamedArg Expr -> Range
forall u t. (HasRange u, HasRange t) => u -> t -> Range
fuseRange Expr
e NamedArg Expr
arg

    inferParenPref :: NamedArg (Either A.Expr (OpApp C.Expr)) -> ParenPreference
    inferParenPref :: NamedArg (Either Expr (OpApp Expr)) -> ParenPreference
inferParenPref NamedArg (Either Expr (OpApp Expr))
e =
      case NamedArg (Either Expr (OpApp Expr)) -> Either Expr (OpApp Expr)
forall a. NamedArg a -> a
namedArg NamedArg (Either Expr (OpApp Expr))
e of
        Right (Ordinary Expr
e) -> Expr -> ParenPreference
inferParenPreference Expr
e
        Left{}             -> ParenPreference
PreferParenless  -- variable inserted by section expansion
        Right{}            -> ParenPreference
PreferParenless  -- syntax lambda

    -- Translate an argument. Returns the paren preference for the argument, so
    -- we can build the correct info for the A.App node.
    toAbsOpArg :: Precedence ->
                  NamedArg (Either A.Expr (OpApp C.Expr)) ->
                  ScopeM (ParenPreference, NamedArg A.Expr)
    toAbsOpArg :: Precedence
-> NamedArg (Either Expr (OpApp Expr))
-> ScopeM (ParenPreference, NamedArg Expr)
toAbsOpArg Precedence
cxt NamedArg (Either Expr (OpApp Expr))
e = (ParenPreference
pref,) (NamedArg Expr -> (ParenPreference, NamedArg Expr))
-> TCMT IO (NamedArg Expr)
-> ScopeM (ParenPreference, NamedArg Expr)
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> ((Named NamedName (Either Expr (OpApp Expr))
 -> TCMT IO (Named NamedName Expr))
-> NamedArg (Either Expr (OpApp Expr)) -> TCMT IO (NamedArg Expr)
forall (t :: * -> *) (f :: * -> *) a b.
(Traversable t, Applicative f) =>
(a -> f b) -> t a -> f (t b)
forall (f :: * -> *) a b.
Applicative f =>
(a -> f b) -> Arg a -> f (Arg b)
traverse ((Named NamedName (Either Expr (OpApp Expr))
  -> TCMT IO (Named NamedName Expr))
 -> NamedArg (Either Expr (OpApp Expr)) -> TCMT IO (NamedArg Expr))
-> ((Either Expr (OpApp Expr) -> ScopeM Expr)
    -> Named NamedName (Either Expr (OpApp Expr))
    -> TCMT IO (Named NamedName Expr))
-> (Either Expr (OpApp Expr) -> ScopeM Expr)
-> NamedArg (Either Expr (OpApp Expr))
-> TCMT IO (NamedArg Expr)
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (Either Expr (OpApp Expr) -> ScopeM Expr)
-> Named NamedName (Either Expr (OpApp Expr))
-> TCMT IO (Named NamedName Expr)
forall (t :: * -> *) (f :: * -> *) a b.
(Traversable t, Applicative f) =>
(a -> f b) -> t a -> f (t b)
forall (f :: * -> *) a b.
Applicative f =>
(a -> f b) -> Named NamedName a -> f (Named NamedName b)
traverse) ((Expr -> ScopeM Expr)
-> (OpApp Expr -> ScopeM Expr)
-> Either Expr (OpApp Expr)
-> ScopeM Expr
forall a c b. (a -> c) -> (b -> c) -> Either a b -> c
either Expr -> ScopeM Expr
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return (Precedence -> OpApp Expr -> ScopeM Expr
toAbstractOpArg Precedence
cxt)) NamedArg (Either Expr (OpApp Expr))
e
      where pref :: ParenPreference
pref = NamedArg (Either Expr (OpApp Expr)) -> ParenPreference
inferParenPref NamedArg (Either Expr (OpApp Expr))
e

    -- The hole left to the first @IdPart@ is filled with an expression in @LeftOperandCtx@.
    left :: Fixity
         -> [NotationPart]
         -> [NamedArg (Either A.Expr (OpApp C.Expr))]
         -> ScopeM [(ParenPreference, NamedArg A.Expr)]
    left :: Fixity
-> Notation
-> [NamedArg (Either Expr (OpApp Expr))]
-> ScopeM [(ParenPreference, NamedArg Expr)]
left Fixity
f (IdPart RString
_ : Notation
xs) [NamedArg (Either Expr (OpApp Expr))]
es = Fixity
-> Notation
-> [NamedArg (Either Expr (OpApp Expr))]
-> ScopeM [(ParenPreference, NamedArg Expr)]
inside Fixity
f Notation
xs [NamedArg (Either Expr (OpApp Expr))]
es
    left Fixity
f (NotationPart
_ : Notation
xs) (NamedArg (Either Expr (OpApp Expr))
e : [NamedArg (Either Expr (OpApp Expr))]
es) = do
        e  <- Precedence
-> NamedArg (Either Expr (OpApp Expr))
-> ScopeM (ParenPreference, NamedArg Expr)
toAbsOpArg (Fixity -> Precedence
LeftOperandCtx Fixity
f) NamedArg (Either Expr (OpApp Expr))
e
        es <- inside f xs es
        return (e : es)
    left Fixity
f (NotationPart
_  : Notation
_)  [] = ScopeM [(ParenPreference, NamedArg Expr)]
forall a. HasCallStack => a
__IMPOSSIBLE__
    left Fixity
f []        [NamedArg (Either Expr (OpApp Expr))]
_  = ScopeM [(ParenPreference, NamedArg Expr)]
forall a. HasCallStack => a
__IMPOSSIBLE__

    -- The holes in between the @IdPart@s are filled with an expression in @InsideOperandCtx@.
    inside :: Fixity
           -> [NotationPart]
           -> [NamedArg (Either A.Expr (OpApp C.Expr))]
           -> ScopeM [(ParenPreference, NamedArg A.Expr)]
    inside :: Fixity
-> Notation
-> [NamedArg (Either Expr (OpApp Expr))]
-> ScopeM [(ParenPreference, NamedArg Expr)]
inside Fixity
f [NotationPart
x]             [NamedArg (Either Expr (OpApp Expr))]
es = Fixity
-> NotationPart
-> [NamedArg (Either Expr (OpApp Expr))]
-> ScopeM [(ParenPreference, NamedArg Expr)]
right Fixity
f NotationPart
x [NamedArg (Either Expr (OpApp Expr))]
es
    inside Fixity
f (IdPart RString
_ : Notation
xs) [NamedArg (Either Expr (OpApp Expr))]
es = Fixity
-> Notation
-> [NamedArg (Either Expr (OpApp Expr))]
-> ScopeM [(ParenPreference, NamedArg Expr)]
inside Fixity
f Notation
xs [NamedArg (Either Expr (OpApp Expr))]
es
    inside Fixity
f (NotationPart
_  : Notation
xs) (NamedArg (Either Expr (OpApp Expr))
e : [NamedArg (Either Expr (OpApp Expr))]
es) = do
        e  <- Precedence
-> NamedArg (Either Expr (OpApp Expr))
-> ScopeM (ParenPreference, NamedArg Expr)
toAbsOpArg Precedence
InsideOperandCtx NamedArg (Either Expr (OpApp Expr))
e
        es <- inside f xs es
        return (e : es)
    inside Fixity
_ []      [] = [(ParenPreference, NamedArg Expr)]
-> ScopeM [(ParenPreference, NamedArg Expr)]
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return []
    inside Fixity
_ (NotationPart
_ : Notation
_) [] = ScopeM [(ParenPreference, NamedArg Expr)]
forall a. HasCallStack => a
__IMPOSSIBLE__
    inside Fixity
_ [] (NamedArg (Either Expr (OpApp Expr))
_ : [NamedArg (Either Expr (OpApp Expr))]
_) = ScopeM [(ParenPreference, NamedArg Expr)]
forall a. HasCallStack => a
__IMPOSSIBLE__

    -- The hole right of the last @IdPart@ is filled with an expression in @RightOperandCtx@.
    right :: Fixity
          -> NotationPart
          -> [NamedArg (Either A.Expr (OpApp C.Expr))]
          -> ScopeM [(ParenPreference, NamedArg A.Expr)]
    right :: Fixity
-> NotationPart
-> [NamedArg (Either Expr (OpApp Expr))]
-> ScopeM [(ParenPreference, NamedArg Expr)]
right Fixity
_ (IdPart RString
_)  [] = [(ParenPreference, NamedArg Expr)]
-> ScopeM [(ParenPreference, NamedArg Expr)]
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return []
    right Fixity
f NotationPart
_          [NamedArg (Either Expr (OpApp Expr))
e] = do
        let pref :: ParenPreference
pref = NamedArg (Either Expr (OpApp Expr)) -> ParenPreference
inferParenPref NamedArg (Either Expr (OpApp Expr))
e
        e <- Precedence
-> NamedArg (Either Expr (OpApp Expr))
-> ScopeM (ParenPreference, NamedArg Expr)
toAbsOpArg (Fixity -> ParenPreference -> Precedence
RightOperandCtx Fixity
f ParenPreference
pref) NamedArg (Either Expr (OpApp Expr))
e
        return [e]
    right Fixity
_ NotationPart
_     [NamedArg (Either Expr (OpApp Expr))]
_  = ScopeM [(ParenPreference, NamedArg Expr)]
forall a. HasCallStack => a
__IMPOSSIBLE__

    replacePlaceholders ::
      OpAppArgs0 e ->
      ScopeM ([A.LamBinding], [NamedArg (Either A.Expr (OpApp e))])
    replacePlaceholders :: forall e.
OpAppArgs0 e
-> ScopeM ([LamBinding], [NamedArg (Either Expr (OpApp e))])
replacePlaceholders []       = ([LamBinding], [NamedArg (Either Expr (OpApp e))])
-> TCMT IO ([LamBinding], [NamedArg (Either Expr (OpApp e))])
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return ([], [])
    replacePlaceholders (NamedArg (MaybePlaceholder (OpApp e))
a : [NamedArg (MaybePlaceholder (OpApp e))]
as) = case NamedArg (MaybePlaceholder (OpApp e)) -> MaybePlaceholder (OpApp e)
forall a. NamedArg a -> a
namedArg NamedArg (MaybePlaceholder (OpApp e))
a of
      NoPlaceholder Maybe PositionInName
_ OpApp e
x -> ([NamedArg (Either Expr (OpApp e))]
 -> [NamedArg (Either Expr (OpApp e))])
-> ([LamBinding], [NamedArg (Either Expr (OpApp e))])
-> ([LamBinding], [NamedArg (Either Expr (OpApp e))])
forall b c a. (b -> c) -> (a, b) -> (a, c)
mapSnd (Either Expr (OpApp e)
-> NamedArg (MaybePlaceholder (OpApp e))
-> NamedArg (Either Expr (OpApp e))
forall a b. a -> NamedArg b -> NamedArg a
set (OpApp e -> Either Expr (OpApp e)
forall a b. b -> Either a b
Right OpApp e
x) NamedArg (MaybePlaceholder (OpApp e))
a NamedArg (Either Expr (OpApp e))
-> [NamedArg (Either Expr (OpApp e))]
-> [NamedArg (Either Expr (OpApp e))]
forall a. a -> [a] -> [a]
:) (([LamBinding], [NamedArg (Either Expr (OpApp e))])
 -> ([LamBinding], [NamedArg (Either Expr (OpApp e))]))
-> TCMT IO ([LamBinding], [NamedArg (Either Expr (OpApp e))])
-> TCMT IO ([LamBinding], [NamedArg (Either Expr (OpApp e))])
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$>
                             [NamedArg (MaybePlaceholder (OpApp e))]
-> TCMT IO ([LamBinding], [NamedArg (Either Expr (OpApp e))])
forall e.
OpAppArgs0 e
-> ScopeM ([LamBinding], [NamedArg (Either Expr (OpApp e))])
replacePlaceholders [NamedArg (MaybePlaceholder (OpApp e))]
as
      Placeholder PositionInName
_     -> do
        x <- Range -> [Char] -> ScopeM Name
forall (m :: * -> *).
MonadFresh NameId m =>
Range -> [Char] -> m Name
freshName Range
forall a. Range' a
noRange [Char]
"section"
        let i = Origin -> ArgInfo -> ArgInfo
forall a. LensOrigin a => Origin -> a -> a
setOrigin Origin
Inserted (ArgInfo -> ArgInfo) -> ArgInfo -> ArgInfo
forall a b. (a -> b) -> a -> b
$ NamedArg (MaybePlaceholder (OpApp e)) -> ArgInfo
forall e. Arg e -> ArgInfo
argInfo NamedArg (MaybePlaceholder (OpApp e))
a
        (ls, ns) <- replacePlaceholders as
        return ( A.mkDomainFree (unnamedArg i $ A.insertedBinder_ x) : ls
               , set (Left (Var x)) a : ns
               )
      where
      set :: a -> NamedArg b -> NamedArg a
      set :: forall a b. a -> NamedArg b -> NamedArg a
set a
x NamedArg b
arg = (Named NamedName b -> Named_ a) -> NamedArg b -> Arg (Named_ a)
forall a b. (a -> b) -> Arg a -> Arg b
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
fmap ((b -> a) -> Named NamedName b -> Named_ a
forall a b. (a -> b) -> Named NamedName a -> Named NamedName b
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
fmap (a -> b -> a
forall a b. a -> b -> a
const a
x)) NamedArg b
arg

-- | Raises an error if the list of attributes contains an unsupported
-- attribute.

checkAttributes :: Attributes -> ScopeM ()
checkAttributes :: Attributes -> TCMT IO ()
checkAttributes []                     = () -> TCMT IO ()
forall a. a -> TCMT IO a
forall (m :: * -> *) a. Monad m => a -> m a
return ()
checkAttributes ((Attribute
attr, Range
r, [Char]
s) : Attributes
attrs) =
  case Attribute
attr of
    RelevanceAttribute{}    -> TCMT IO ()
cont
    CA.TacticAttribute{}    -> TCMT IO ()
cont
    LockAttribute Lock
IsNotLock -> TCMT IO ()
cont
    LockAttribute IsLock{}  -> do
      TCMT IO Bool -> TCMT IO () -> TCMT IO ()
forall (m :: * -> *). Monad m => m Bool -> m () -> m ()
unlessM (PragmaOptions -> Bool
optGuarded (PragmaOptions -> Bool) -> TCMT IO PragmaOptions -> TCMT IO Bool
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> TCMT IO PragmaOptions
forall (m :: * -> *). HasOptions m => m PragmaOptions
pragmaOptions) (TCMT IO () -> TCMT IO ()) -> TCMT IO () -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$
        Range -> TCMT IO () -> TCMT IO ()
forall (m :: * -> *) x a.
(MonadTrace m, HasRange x) =>
x -> m a -> m a
setCurrentRange Range
r (TCMT IO () -> TCMT IO ()) -> TCMT IO () -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ TypeError -> TCMT IO ()
forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
TypeError -> m a
typeError (TypeError -> TCMT IO ()) -> TypeError -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ [Char] -> [Char] -> [Char] -> TypeError
AttributeKindNotEnabled [Char]
"Lock" [Char]
"--guarded" [Char]
s
      TCMT IO ()
cont
    QuantityAttribute Quantityω{} -> TCMT IO ()
cont
    QuantityAttribute Quantity1{} -> TCMT IO ()
forall a. HasCallStack => a
__IMPOSSIBLE__
    QuantityAttribute Quantity0{} -> do
      TCMT IO Bool -> TCMT IO () -> TCMT IO ()
forall (m :: * -> *). Monad m => m Bool -> m () -> m ()
unlessM (PragmaOptions -> Bool
optErasure (PragmaOptions -> Bool) -> TCMT IO PragmaOptions -> TCMT IO Bool
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> TCMT IO PragmaOptions
forall (m :: * -> *). HasOptions m => m PragmaOptions
pragmaOptions) (TCMT IO () -> TCMT IO ()) -> TCMT IO () -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$
        Range -> TCMT IO () -> TCMT IO ()
forall (m :: * -> *) x a.
(MonadTrace m, HasRange x) =>
x -> m a -> m a
setCurrentRange Range
r (TCMT IO () -> TCMT IO ()) -> TCMT IO () -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ TypeError -> TCMT IO ()
forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
TypeError -> m a
typeError (TypeError -> TCMT IO ()) -> TypeError -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ [Char] -> [Char] -> [Char] -> TypeError
AttributeKindNotEnabled [Char]
"Erasure" [Char]
"--erasure" [Char]
s
      TCMT IO ()
cont
    CohesionAttribute{} -> do
      TCMT IO Bool -> TCMT IO () -> TCMT IO ()
forall (m :: * -> *). Monad m => m Bool -> m () -> m ()
unlessM (PragmaOptions -> Bool
optCohesion (PragmaOptions -> Bool) -> TCMT IO PragmaOptions -> TCMT IO Bool
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> TCMT IO PragmaOptions
forall (m :: * -> *). HasOptions m => m PragmaOptions
pragmaOptions) (TCMT IO () -> TCMT IO ()) -> TCMT IO () -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$
        Range -> TCMT IO () -> TCMT IO ()
forall (m :: * -> *) x a.
(MonadTrace m, HasRange x) =>
x -> m a -> m a
setCurrentRange Range
r (TCMT IO () -> TCMT IO ()) -> TCMT IO () -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ TypeError -> TCMT IO ()
forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
TypeError -> m a
typeError (TypeError -> TCMT IO ()) -> TypeError -> TCMT IO ()
forall a b. (a -> b) -> a -> b
$ [Char] -> [Char] -> [Char] -> TypeError
AttributeKindNotEnabled [Char]
"Cohesion" [Char]
"--cohesion" [Char]
s
      TCMT IO ()
cont
  where
  cont :: TCMT IO ()
cont = Attributes -> TCMT IO ()
checkAttributes Attributes
attrs

{--------------------------------------------------------------------------
    Things we parse but are not part of the Agda file syntax
 --------------------------------------------------------------------------}

-- | Content of interaction hole.

instance ToAbstract C.HoleContent where
  type AbsOfCon C.HoleContent = A.HoleContent
  toAbstract :: HoleContent -> ScopeM (AbsOfCon HoleContent)
toAbstract = \case
    HoleContentExpr Expr
e     -> Expr -> HoleContent' () BindName Pattern Expr
forall qn nm p e. e -> HoleContent' qn nm p e
HoleContentExpr (Expr -> HoleContent' () BindName Pattern Expr)
-> ScopeM Expr -> TCMT IO (HoleContent' () BindName Pattern Expr)
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> Expr -> ScopeM (AbsOfCon Expr)
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract Expr
e
    HoleContentRewrite [RewriteEqn]
es -> [RewriteEqn' () BindName Pattern Expr]
-> HoleContent' () BindName Pattern Expr
forall qn nm p e. [RewriteEqn' qn nm p e] -> HoleContent' qn nm p e
HoleContentRewrite ([RewriteEqn' () BindName Pattern Expr]
 -> HoleContent' () BindName Pattern Expr)
-> TCMT IO [RewriteEqn' () BindName Pattern Expr]
-> TCMT IO (HoleContent' () BindName Pattern Expr)
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> [RewriteEqn] -> ScopeM (AbsOfCon [RewriteEqn])
forall c. ToAbstract c => c -> ScopeM (AbsOfCon c)
toAbstract [RewriteEqn]
es