{-# LANGUAGE NondecreasingIndentation #-}

module Agda.TypeChecking.Rules.Def where

import Prelude hiding ( null )

import Control.Monad        ( forM, forM_ )
import Control.Monad.Except ( MonadError(..) )

import Data.Bifunctor
import Data.Function (on)
import Data.IntSet (IntSet)
import qualified Data.IntSet as IntSet
import qualified Data.List as List
import Data.Maybe
import Data.Semigroup (Semigroup((<>)))

import Agda.Interaction.Options

import Agda.Syntax.Common
import qualified Agda.Syntax.Concrete as C
import qualified Agda.Syntax.Concrete.Pretty as C
import Agda.Syntax.Position
import Agda.Syntax.Abstract.Pattern as A
import qualified Agda.Syntax.Abstract as A
import qualified Agda.Syntax.Abstract.Views as A
import Agda.Syntax.Internal as I
import Agda.Syntax.Internal.Pattern as I
import Agda.Syntax.Internal.MetaVars (allMetasList)
import qualified Agda.Syntax.Info as Info
import Agda.Syntax.Info hiding (defAbstract)

import Agda.TypeChecking.Monad
import qualified Agda.TypeChecking.Monad.Benchmark as Bench
import Agda.TypeChecking.Warnings ( warning, genericWarning )

import Agda.TypeChecking.Constraints
import Agda.TypeChecking.Conversion
import Agda.TypeChecking.Coverage.SplitTree
import Agda.TypeChecking.Inlining
import Agda.TypeChecking.Reduce
import Agda.TypeChecking.Patterns.Abstract (expandPatternSynonyms)
import Agda.TypeChecking.Pretty
import Agda.TypeChecking.Substitute
import Agda.TypeChecking.CheckInternal
import Agda.TypeChecking.With
import Agda.TypeChecking.Telescope
import Agda.TypeChecking.Telescope.Path
import Agda.TypeChecking.Injectivity
import Agda.TypeChecking.SizedTypes.Solve
import Agda.TypeChecking.Rewriting.Confluence
import Agda.TypeChecking.CompiledClause (CompiledClauses'(..), hasProjectionPatterns)
import Agda.TypeChecking.CompiledClause.Compile
import Agda.TypeChecking.Primitive hiding (Nat)
import Agda.TypeChecking.RecordPatterns ( recordRHSToCopatterns )
import Agda.TypeChecking.Sort

import Agda.TypeChecking.Rules.Term
import Agda.TypeChecking.Rules.LHS                 ( checkLeftHandSide, LHSResult(..), bindAsPatterns )
import {-# SOURCE #-} Agda.TypeChecking.Rules.Decl ( checkDecls )

import Agda.Utils.Functor
import Agda.Utils.Lens
import Agda.Utils.List
import Agda.Utils.List1 ( List1, pattern (:|), (<|) )
import qualified Agda.Utils.List1 as List1
import Agda.Utils.Maybe
import Agda.Utils.Monad
import Agda.Utils.Null
import Agda.Utils.Permutation
import Agda.Syntax.Common.Pretty ( prettyShow )
import qualified Agda.Syntax.Common.Pretty as P
import Agda.Utils.Singleton
import Agda.Utils.Size
import qualified Agda.Utils.SmallSet as SmallSet
import Agda.Utils.Update

import Agda.Utils.Impossible

---------------------------------------------------------------------------
-- * Definitions by pattern matching
---------------------------------------------------------------------------

checkFunDef :: A.DefInfo -> QName -> [A.Clause] -> TCM ()
checkFunDef :: DefInfo -> QName -> [Clause' LHS] -> TCM ()
checkFunDef DefInfo
i QName
name [Clause' LHS]
cs = do
        -- Reset blocking tag (in case a previous attempt was blocked)
        forall (m :: * -> *).
MonadTCState m =>
(Signature -> Signature) -> m ()
modifySignature forall a b. (a -> b) -> a -> b
$ QName -> (Definition -> Definition) -> Signature -> Signature
updateDefinition QName
name forall a b. (a -> b) -> a -> b
$ (Blocked_ -> Blocked_) -> Definition -> Definition
updateDefBlocked forall a b. (a -> b) -> a -> b
$ forall a b. a -> b -> a
const forall a b. (a -> b) -> a -> b
$
          forall t a. NotBlocked' t -> a -> Blocked' t a
NotBlocked (forall t. QName -> NotBlocked' t
MissingClauses QName
name) ()
        -- Get the type and relevance of the function
        Definition
def <- forall (m :: * -> *).
(Functor m, HasConstInfo m, HasOptions m, ReadTCState m,
 MonadTCEnv m, MonadDebug m) =>
Definition -> m Definition
instantiateDef forall (m :: * -> *) a b. Monad m => (a -> m b) -> m a -> m b
=<< forall (m :: * -> *). HasConstInfo m => QName -> m Definition
getConstInfo QName
name
        let t :: Type
t    = Definition -> Type
defType Definition
def
        let info :: ArgInfo
info = forall a. LensArgInfo a => a -> ArgInfo
getArgInfo Definition
def

        -- If the function is erased, then hard compile-time mode is
        -- entered.
        forall q a. LensQuantity q => q -> TCM a -> TCM a
setHardCompileTimeModeIfErased' ArgInfo
info forall a b. (a -> b) -> a -> b
$ do

        case [Clause' LHS] -> Type -> Maybe (Expr, Maybe Expr, MetaId)
isAlias [Clause' LHS]
cs Type
t of  -- #418: Don't use checkAlias for abstract definitions, since the type
                              -- of an abstract function must not be informed by its definition.
          Just (Expr
e, Maybe Expr
mc, MetaId
x)
            | forall t. DefInfo' t -> IsAbstract
Info.defAbstract DefInfo
i forall a. Eq a => a -> a -> Bool
== IsAbstract
ConcreteDef, forall t. DefInfo' t -> IsOpaque
Info.defOpaque DefInfo
i forall a. Eq a => a -> a -> Bool
== IsOpaque
TransparentDef ->
              forall (m :: * -> *) a. MonadTrace m => Call -> m a -> m a
traceCall (Range -> QName -> [Clause' LHS] -> Bool -> Call
CheckFunDefCall (forall a. HasRange a => a -> Range
getRange DefInfo
i) QName
name [Clause' LHS]
cs Bool
True) forall a b. (a -> b) -> a -> b
$ do
                -- Andreas, 2012-11-22: if the alias is in an abstract block
                -- it has been frozen.  We unfreeze it to enable type inference.
                -- See issue 729.
                -- Ulf, 2021-02-09: also unfreeze metas in the sort of this type
                forall (m :: * -> *). Monad m => m Bool -> m () -> m ()
whenM (forall (m :: * -> *).
(HasCallStack, MonadDebug m, ReadTCState m) =>
MetaId -> m Bool
isFrozen MetaId
x) forall a b. (a -> b) -> a -> b
$ do
                  [MetaId]
xs <- forall a. AllMetas a => a -> [MetaId]
allMetasList forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall a. Judgement a -> Type
jMetaType forall b c a. (b -> c) -> (a -> b) -> a -> c
. MetaVariable -> Judgement MetaId
mvJudgement forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> forall (m :: * -> *).
(HasCallStack, MonadDebug m, ReadTCState m) =>
MetaId -> m MetaVariable
lookupLocalMeta MetaId
x
                  forall (t :: * -> *) (m :: * -> *) a b.
(Foldable t, Monad m) =>
(a -> m b) -> t a -> m ()
mapM_ forall a (m :: * -> *).
(UnFreezeMeta a, MonadMetaSolver m) =>
a -> m ()
unfreezeMeta (MetaId
x forall a. a -> [a] -> [a]
: [MetaId]
xs)
                Type -> ArgInfo -> DefInfo -> QName -> Expr -> Maybe Expr -> TCM ()
checkAlias Type
t ArgInfo
info DefInfo
i QName
name Expr
e Maybe Expr
mc
            | Bool
otherwise -> do -- Warn about abstract alias (will never work!)
              -- Ulf, 2021-11-18, #5620: Don't warn if the meta is solved. A more intuitive solution
              -- would be to not treat definitions with solved meta types as aliases, but in mutual
              -- blocks you might actually have solved the type of an alias by the time you get to
              -- the definition. See test/Succeed/SizeInfinity.agda for an example where this
              -- happens.
              let
                what :: [Char]
what
                  | forall t. DefInfo' t -> IsOpaque
Info.defOpaque DefInfo
i forall a. Eq a => a -> a -> Bool
== IsOpaque
TransparentDef = [Char]
"abstract"
                  | Bool
otherwise                          = [Char]
"opaque"
              forall (m :: * -> *). Monad m => m Bool -> m () -> m ()
whenM (MetaInstantiation -> Bool
isOpenMeta forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> forall (m :: * -> *).
ReadTCState m =>
MetaId -> m MetaInstantiation
lookupMetaInstantiation MetaId
x) forall a b. (a -> b) -> a -> b
$
                forall (m :: * -> *) x a.
(MonadTrace m, HasRange x) =>
x -> m a -> m a
setCurrentRange DefInfo
i forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *). MonadWarning m => Doc -> m ()
genericWarning forall (m :: * -> *) a b. Monad m => (a -> m b) -> m a -> m b
=<<
                  TCMT IO Doc
"Missing type signature for" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall (m :: * -> *). Applicative m => [Char] -> m Doc
text [Char]
what forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> TCMT IO Doc
"definition" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> (forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM QName
name forall a. Semigroup a => a -> a -> a
<> TCMT IO Doc
".") forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
$$
                  forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
fsep (forall (m :: * -> *). Applicative m => [Char] -> [m Doc]
pwords ([Char]
"Types of " forall a. [a] -> [a] -> [a]
++ [Char]
what forall a. [a] -> [a] -> [a]
++ [Char]
" definitions are never inferred since this would leak") forall a. [a] -> [a] -> [a]
++
                        forall (m :: * -> *). Applicative m => [Char] -> [m Doc]
pwords ([Char]
"information that should be " forall a. [a] -> [a] -> [a]
++ [Char]
what forall a. [a] -> [a] -> [a]
++ [Char]
"."))
              Type
-> ArgInfo
-> Maybe ExtLamInfo
-> Maybe QName
-> DefInfo
-> QName
-> [Clause' LHS]
-> TCM ()
checkFunDef' Type
t ArgInfo
info forall a. Maybe a
Nothing forall a. Maybe a
Nothing DefInfo
i QName
name [Clause' LHS]
cs
          Maybe (Expr, Maybe Expr, MetaId)
_ -> Type
-> ArgInfo
-> Maybe ExtLamInfo
-> Maybe QName
-> DefInfo
-> QName
-> [Clause' LHS]
-> TCM ()
checkFunDef' Type
t ArgInfo
info forall a. Maybe a
Nothing forall a. Maybe a
Nothing DefInfo
i QName
name [Clause' LHS]
cs

        -- If it's a macro check that it ends in Term → TC ⊤
        let ismacro :: Bool
ismacro = Defn -> Bool
isMacro forall b c a. (b -> c) -> (a -> b) -> a -> c
. Definition -> Defn
theDef forall a b. (a -> b) -> a -> b
$ Definition
def
        forall (f :: * -> *). Applicative f => Bool -> f () -> f ()
when (Bool
ismacro Bool -> Bool -> Bool
|| forall t. DefInfo' t -> IsMacro
Info.defMacro DefInfo
i forall a. Eq a => a -> a -> Bool
== IsMacro
MacroDef) forall a b. (a -> b) -> a -> b
$ Type -> TCM ()
checkMacroType Type
t
    forall a. TCM a -> ((TCErr, Blocker) -> TCM a) -> TCM a
`catchIlltypedPatternBlockedOnMeta` \ (TCErr
err, Blocker
blocker) -> do
        forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> TCMT IO Doc -> m ()
reportSDoc [Char]
"tc.def" Int
20 forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
vcat forall a b. (a -> b) -> a -> b
$
          [ TCMT IO Doc
"checking function definition got stuck on: " forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall (m :: * -> *) a. (Applicative m, Pretty a) => a -> m Doc
pretty Blocker
blocker ]
        forall (m :: * -> *).
MonadTCState m =>
(Signature -> Signature) -> m ()
modifySignature forall a b. (a -> b) -> a -> b
$ QName -> (Definition -> Definition) -> Signature -> Signature
updateDefinition QName
name forall a b. (a -> b) -> a -> b
$ (Blocked_ -> Blocked_) -> Definition -> Definition
updateDefBlocked forall a b. (a -> b) -> a -> b
$ forall a b. a -> b -> a
const forall a b. (a -> b) -> a -> b
$ forall t a. Blocker -> a -> Blocked' t a
Blocked Blocker
blocker ()
        forall (m :: * -> *).
MonadConstraint m =>
Blocker -> Constraint -> m ()
addConstraint Blocker
blocker forall a b. (a -> b) -> a -> b
$ DefInfo -> QName -> [Clause' LHS] -> TCErr -> Constraint
CheckFunDef DefInfo
i QName
name [Clause' LHS]
cs TCErr
err

checkMacroType :: Type -> TCM ()
checkMacroType :: Type -> TCM ()
checkMacroType Type
t = do
  TelV Tele (Dom Type)
tel Type
tr <- forall (m :: * -> *).
(MonadReduce m, MonadAddContext m) =>
Type -> m (TelV Type)
telView Type
t

  let telList :: [Dom ([Char], Type)]
telList = forall t. Tele (Dom t) -> [Dom ([Char], t)]
telToList Tele (Dom Type)
tel
      resType :: Type
resType = forall t. Abstract t => Tele (Dom Type) -> t -> t
abstract ([Dom ([Char], Type)] -> Tele (Dom Type)
telFromList (forall a. Int -> [a] -> [a]
drop (forall (t :: * -> *) a. Foldable t => t a -> Int
length [Dom ([Char], Type)]
telList forall a. Num a => a -> a -> a
- Int
1) [Dom ([Char], Type)]
telList)) Type
tr
  Type
expectedType <- forall (m :: * -> *). Functor m => m Term -> m Type
el forall (m :: * -> *).
(HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) =>
m Term
primAgdaTerm forall (m :: * -> *). Applicative m => m Type -> m Type -> m Type
--> forall (m :: * -> *). Functor m => m Term -> m Type
el (forall (m :: * -> *).
(HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) =>
m Term
primAgdaTCM forall (m :: * -> *). Applicative m => m Term -> m Term -> m Term
<#> forall (m :: * -> *).
(HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) =>
m Term
primLevelZero forall (m :: * -> *). Applicative m => m Term -> m Term -> m Term
<@> forall (m :: * -> *).
(HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) =>
m Term
primUnit)
  forall (m :: * -> *). MonadConversion m => Type -> Type -> m ()
equalType Type
resType Type
expectedType
    forall e (m :: * -> *) a.
MonadError e m =>
m a -> (e -> m a) -> m a
`catchError` \ TCErr
_ -> forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
TypeError -> m a
typeError forall a b. (a -> b) -> a -> b
$ Type -> TypeError
MacroResultTypeMismatch Type
expectedType

-- | A single clause without arguments and without type signature is an alias.
isAlias :: [A.Clause] -> Type -> Maybe (A.Expr, Maybe C.Expr, MetaId)
isAlias :: [Clause' LHS] -> Type -> Maybe (Expr, Maybe Expr, MetaId)
isAlias [Clause' LHS]
cs Type
t =
        case [Clause' LHS] -> Maybe (Expr, Maybe Expr)
trivialClause [Clause' LHS]
cs of
          -- if we have just one clause without pattern matching and
          -- without a type signature, then infer, to allow
          -- "aliases" for things starting with hidden abstractions
          Just (Expr
e, Maybe Expr
mc) | Just MetaId
x <- Term -> Maybe MetaId
isMeta (forall t a. Type'' t a -> a
unEl Type
t) -> forall a. a -> Maybe a
Just (Expr
e, Maybe Expr
mc, MetaId
x)
          Maybe (Expr, Maybe Expr)
_ -> forall a. Maybe a
Nothing
  where
    isMeta :: Term -> Maybe MetaId
isMeta (MetaV MetaId
x Elims
_) = forall a. a -> Maybe a
Just MetaId
x
    isMeta Term
_           = forall a. Maybe a
Nothing
    trivialClause :: [Clause' LHS] -> Maybe (Expr, Maybe Expr)
trivialClause [A.Clause (A.LHS LHSInfo
i (A.LHSHead QName
f [])) [ProblemEq]
_ (A.RHS Expr
e Maybe Expr
mc) WhereDeclarations
wh Bool
_]
      | forall a. Null a => a -> Bool
null WhereDeclarations
wh     = forall a. a -> Maybe a
Just (Expr
e, Maybe Expr
mc)
    trivialClause [Clause' LHS]
_ = forall a. Maybe a
Nothing

-- | Check a trivial definition of the form @f = e@
checkAlias :: Type -> ArgInfo -> A.DefInfo -> QName -> A.Expr -> Maybe C.Expr -> TCM ()
checkAlias :: Type -> ArgInfo -> DefInfo -> QName -> Expr -> Maybe Expr -> TCM ()
checkAlias Type
t ArgInfo
ai DefInfo
i QName
name Expr
e Maybe Expr
mc =
  let clause :: SpineClause
clause = A.Clause { clauseLHS :: SpineLHS
clauseLHS          = LHSInfo -> QName -> [NamedArg (Pattern' Expr)] -> SpineLHS
A.SpineLHS (Range -> ExpandedEllipsis -> LHSInfo
LHSInfo (forall a. HasRange a => a -> Range
getRange DefInfo
i) ExpandedEllipsis
NoEllipsis) QName
name []
                        , clauseStrippedPats :: [ProblemEq]
clauseStrippedPats = []
                        , clauseRHS :: RHS
clauseRHS          = Expr -> Maybe Expr -> RHS
A.RHS Expr
e Maybe Expr
mc
                        , clauseWhereDecls :: WhereDeclarations
clauseWhereDecls   = WhereDeclarations
A.noWhereDecls
                        , clauseCatchall :: Bool
clauseCatchall     = Bool
False } in
  forall a.
QName
-> Int
-> Type
-> Maybe Substitution
-> SpineClause
-> TCM a
-> TCM a
atClause QName
name Int
0 Type
t forall a. Maybe a
Nothing SpineClause
clause forall a b. (a -> b) -> a -> b
$ do
  forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> TCMT IO Doc -> m ()
reportSDoc [Char]
"tc.def.alias" Int
10 forall a b. (a -> b) -> a -> b
$ TCMT IO Doc
"checkAlias" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
vcat
    [ forall (m :: * -> *). Applicative m => [Char] -> m Doc
text (forall a. Pretty a => a -> [Char]
prettyShow QName
name) forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall (m :: * -> *). Applicative m => m Doc
colon  forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Type
t
    , forall (m :: * -> *). Applicative m => [Char] -> m Doc
text (forall a. Pretty a => a -> [Char]
prettyShow QName
name) forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall (m :: * -> *). Applicative m => m Doc
equals forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Expr
e
    ]

  -- Infer the type of the rhs.
  -- Andreas, 2018-06-09, issue #2170.
  -- The context will only be resurrected if we have --irrelevant-projections.
  Term
v <- forall (tcm :: * -> *) r a.
(MonadTCM tcm, LensModality r) =>
r -> tcm a -> tcm a
applyModalityToContextFunBody ArgInfo
ai forall a b. (a -> b) -> a -> b
$ Comparison -> Expr -> Type -> TCMT IO Term
checkDontExpandLast Comparison
CmpLeq Expr
e Type
t

  forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> TCMT IO Doc -> m ()
reportSDoc [Char]
"tc.def.alias" Int
20 forall a b. (a -> b) -> a -> b
$ TCMT IO Doc
"checkAlias: finished checking"

  DefaultToInfty -> TCM ()
solveSizeConstraints DefaultToInfty
DontDefaultToInfty

  Term
v <- forall a (m :: * -> *).
(InstantiateFull a, MonadReduce m) =>
a -> m a
instantiateFull Term
v  -- if we omit this, we loop (stdlib: Relation.Binary.Sum)
    -- or the termination checker might stumble over levels in sorts
    -- that cannot be converted to expressions without the level built-ins
    -- (test/succeed/Issue655.agda)

  -- compute body modification for irrelevant definitions, see issue 610
  let bodyMod :: Term -> Term
bodyMod = case forall a. LensRelevance a => a -> Relevance
getRelevance ArgInfo
ai of
        Relevance
Irrelevant -> Term -> Term
dontCare
        Relevance
_          -> forall a. a -> a
id

  -- Add the definition
  FunctionData
fun <- forall (m :: * -> *). HasOptions m => m FunctionData
emptyFunctionData
  QName -> ArgInfo -> QName -> Type -> Defn -> TCM ()
addConstant' QName
name ArgInfo
ai QName
name Type
t forall a b. (a -> b) -> a -> b
$ forall o i. Lens' o i -> LensSet o i
set Lens' Defn Bool
funMacro (forall t. DefInfo' t -> IsMacro
Info.defMacro DefInfo
i forall a. Eq a => a -> a -> Bool
== IsMacro
MacroDef) forall a b. (a -> b) -> a -> b
$
      FunctionData -> Defn
FunctionDefn FunctionData
fun
          { _funClauses :: [Clause]
_funClauses   = [ Clause  -- trivial clause @name = v@
              { clauseLHSRange :: Range
clauseLHSRange    = forall a. HasRange a => a -> Range
getRange DefInfo
i
              , clauseFullRange :: Range
clauseFullRange   = forall a. HasRange a => a -> Range
getRange DefInfo
i
              , clauseTel :: Tele (Dom Type)
clauseTel         = forall a. Tele a
EmptyTel
              , namedClausePats :: [NamedArg DeBruijnPattern]
namedClausePats   = []
              , clauseBody :: Maybe Term
clauseBody        = forall a. a -> Maybe a
Just forall a b. (a -> b) -> a -> b
$ Term -> Term
bodyMod Term
v
              , clauseType :: Maybe (Arg Type)
clauseType        = forall a. a -> Maybe a
Just forall a b. (a -> b) -> a -> b
$ forall e. ArgInfo -> e -> Arg e
Arg ArgInfo
ai Type
t
              , clauseCatchall :: Bool
clauseCatchall    = Bool
False
              , clauseExact :: Maybe Bool
clauseExact       = forall a. a -> Maybe a
Just Bool
True
              , clauseRecursive :: Maybe Bool
clauseRecursive   = forall a. Maybe a
Nothing   -- we don't know yet
              , clauseUnreachable :: Maybe Bool
clauseUnreachable = forall a. a -> Maybe a
Just Bool
False
              , clauseEllipsis :: ExpandedEllipsis
clauseEllipsis    = ExpandedEllipsis
NoEllipsis
              , clauseWhereModule :: Maybe ModuleName
clauseWhereModule = forall a. Maybe a
Nothing
              } ]
          , _funCompiled :: Maybe CompiledClauses
_funCompiled  = forall a. a -> Maybe a
Just forall a b. (a -> b) -> a -> b
$ forall a. [Arg [Char]] -> a -> CompiledClauses' a
Done [] forall a b. (a -> b) -> a -> b
$ Term -> Term
bodyMod Term
v
          , _funSplitTree :: Maybe SplitTree
_funSplitTree = forall a. a -> Maybe a
Just forall a b. (a -> b) -> a -> b
$ forall a. Int -> SplitTree' a
SplittingDone Int
0
          , _funAbstr :: IsAbstract
_funAbstr     = forall t. DefInfo' t -> IsAbstract
Info.defAbstract DefInfo
i
          , _funOpaque :: IsOpaque
_funOpaque    = forall t. DefInfo' t -> IsOpaque
Info.defOpaque DefInfo
i
          }

  -- Andreas, 2017-01-01, issue #2372:
  -- Add the definition to the instance table, if needed, to update its type.
  case forall t. DefInfo' t -> IsInstance
Info.defInstance DefInfo
i of
    InstanceDef Range
_r -> forall (m :: * -> *) x a.
(MonadTrace m, HasRange x) =>
x -> m a -> m a
setCurrentRange QName
name forall a b. (a -> b) -> a -> b
$ QName -> Type -> TCM ()
addTypedInstance QName
name Type
t
      -- Put highlighting on the name only;
      -- @(getRange (r, name))@ does not give good results.
    IsInstance
NotInstanceDef -> forall (f :: * -> *) a. Applicative f => a -> f a
pure ()

  forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> TCMT IO Doc -> m ()
reportSDoc [Char]
"tc.def.alias" Int
20 forall a b. (a -> b) -> a -> b
$ TCMT IO Doc
"checkAlias: leaving"


-- | Type check a definition by pattern matching.
checkFunDef' :: Type             -- ^ the type we expect the function to have
             -> ArgInfo          -- ^ is it irrelevant (for instance)
             -> Maybe ExtLamInfo -- ^ does the definition come from an extended lambda
                                 --   (if so, we need to know some stuff about lambda-lifted args)
             -> Maybe QName      -- ^ is it a with function (if so, what's the name of the parent function)
             -> A.DefInfo        -- ^ range info
             -> QName            -- ^ the name of the function
             -> [A.Clause]       -- ^ the clauses to check
             -> TCM ()
checkFunDef' :: Type
-> ArgInfo
-> Maybe ExtLamInfo
-> Maybe QName
-> DefInfo
-> QName
-> [Clause' LHS]
-> TCM ()
checkFunDef' Type
t ArgInfo
ai Maybe ExtLamInfo
extlam Maybe QName
with DefInfo
i QName
name [Clause' LHS]
cs =
  Type
-> ArgInfo
-> Maybe ExtLamInfo
-> Maybe QName
-> DefInfo
-> QName
-> Maybe Substitution
-> [Clause' LHS]
-> TCM ()
checkFunDefS Type
t ArgInfo
ai Maybe ExtLamInfo
extlam Maybe QName
with DefInfo
i QName
name forall a. Maybe a
Nothing [Clause' LHS]
cs

-- | Type check a definition by pattern matching.
checkFunDefS :: Type             -- ^ the type we expect the function to have
             -> ArgInfo          -- ^ is it irrelevant (for instance)
             -> Maybe ExtLamInfo -- ^ does the definition come from an extended lambda
                                 --   (if so, we need to know some stuff about lambda-lifted args)
             -> Maybe QName      -- ^ is it a with function (if so, what's the name of the parent function)
             -> A.DefInfo        -- ^ range info
             -> QName            -- ^ the name of the function
             -> Maybe Substitution -- ^ substitution (from with abstraction) that needs to be applied to module parameters
             -> [A.Clause]       -- ^ the clauses to check
             -> TCM ()
checkFunDefS :: Type
-> ArgInfo
-> Maybe ExtLamInfo
-> Maybe QName
-> DefInfo
-> QName
-> Maybe Substitution
-> [Clause' LHS]
-> TCM ()
checkFunDefS Type
t ArgInfo
ai Maybe ExtLamInfo
extlam Maybe QName
with DefInfo
i QName
name Maybe Substitution
withSub [Clause' LHS]
cs = do

    forall (m :: * -> *) a. MonadTrace m => Call -> m a -> m a
traceCall (Range -> QName -> [Clause' LHS] -> Bool -> Call
CheckFunDefCall (forall a. HasRange a => a -> Range
getRange DefInfo
i) QName
name [Clause' LHS]
cs Bool
True) forall a b. (a -> b) -> a -> b
$ do
        forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> TCMT IO Doc -> m ()
reportSDoc [Char]
"tc.def.fun" Int
10 forall a b. (a -> b) -> a -> b
$
          forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
sep [ TCMT IO Doc
"checking body of" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM QName
name
              , forall (m :: * -> *). Functor m => Int -> m Doc -> m Doc
nest Int
2 forall a b. (a -> b) -> a -> b
$ TCMT IO Doc
":" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Type
t
              , forall (m :: * -> *). Functor m => Int -> m Doc -> m Doc
nest Int
2 forall a b. (a -> b) -> a -> b
$ TCMT IO Doc
"full type:" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> (forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM forall b c a. (b -> c) -> (a -> b) -> a -> c
. Definition -> Type
defType forall (m :: * -> *) a b. Monad m => (a -> m b) -> m a -> m b
=<< forall (m :: * -> *). HasConstInfo m => QName -> m Definition
getConstInfo QName
name)
              ]

        forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> TCMT IO Doc -> m ()
reportSDoc [Char]
"tc.def.fun" Int
70 forall a b. (a -> b) -> a -> b
$
          forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
sep forall a b. (a -> b) -> a -> b
$ TCMT IO Doc
"clauses:" forall a. a -> [a] -> [a]
: forall a b. (a -> b) -> [a] -> [b]
map (forall (m :: * -> *). Functor m => Int -> m Doc -> m Doc
nest Int
2 forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall (m :: * -> *). Applicative m => [Char] -> m Doc
text forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall a. Show a => a -> [Char]
show forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall a. ExprLike a => a -> a
A.deepUnscope) [Clause' LHS]
cs

        [SpineClause]
cs <- forall (m :: * -> *) a. Monad m => a -> m a
return forall a b. (a -> b) -> a -> b
$ forall a b. (a -> b) -> [a] -> [b]
map forall a b. LHSToSpine a b => a -> b
A.lhsToSpine [Clause' LHS]
cs

        forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> TCMT IO Doc -> m ()
reportSDoc [Char]
"tc.def.fun" Int
70 forall a b. (a -> b) -> a -> b
$
          forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
sep forall a b. (a -> b) -> a -> b
$ TCMT IO Doc
"spine clauses:" forall a. a -> [a] -> [a]
: forall a b. (a -> b) -> [a] -> [b]
map (forall (m :: * -> *). Functor m => Int -> m Doc -> m Doc
nest Int
2 forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall (m :: * -> *). Applicative m => [Char] -> m Doc
text forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall a. Show a => a -> [Char]
show forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall a. ExprLike a => a -> a
A.deepUnscope) [SpineClause]
cs

        -- Ensure that all clauses have the same number of trailing hidden patterns
        -- This is necessary since trailing implicits are no longer eagerly inserted.
        -- Andreas, 2013-10-13
        -- Since we have flexible function arity, it is no longer necessary
        -- to patch clauses to same arity
        -- cs <- trailingImplicits t cs

        -- Check the clauses
        [(Clause, ClausesPostChecks)]
cs <- forall (m :: * -> *) a. MonadTrace m => Call -> m a -> m a
traceCall Call
NoHighlighting forall a b. (a -> b) -> a -> b
$ do -- To avoid flicker.
          forall (t :: * -> *) (m :: * -> *) a b.
(Traversable t, Monad m) =>
t a -> (a -> m b) -> m (t b)
forM (forall a b. [a] -> [b] -> [(a, b)]
zip [SpineClause]
cs [Int
0..]) forall a b. (a -> b) -> a -> b
$ \ (SpineClause
c, Int
clauseNo) -> do
            forall a.
QName
-> Int
-> Type
-> Maybe Substitution
-> SpineClause
-> TCM a
-> TCM a
atClause QName
name Int
clauseNo Type
t Maybe Substitution
withSub SpineClause
c forall a b. (a -> b) -> a -> b
$ do
              (Clause
c,ClausesPostChecks
b) <- forall (tcm :: * -> *) r a.
(MonadTCM tcm, LensModality r) =>
r -> tcm a -> tcm a
applyModalityToContextFunBody ArgInfo
ai forall a b. (a -> b) -> a -> b
$ do
                Type
-> Maybe Substitution
-> SpineClause
-> TCMT IO (Clause, ClausesPostChecks)
checkClause Type
t Maybe Substitution
withSub SpineClause
c
              -- Andreas, 2013-11-23 do not solve size constraints here yet
              -- in case we are checking the body of an extended lambda.
              -- 2014-04-24: The size solver requires each clause to be
              -- checked individually, since otherwise we get constraints
              -- in typing contexts which are not prefixes of each other.
              forall m a. Monoid m => Maybe a -> m -> m
whenNothing Maybe ExtLamInfo
extlam forall a b. (a -> b) -> a -> b
$ DefaultToInfty -> TCM ()
solveSizeConstraints DefaultToInfty
DontDefaultToInfty
              -- Andreas, 2013-10-27 add clause as soon it is type-checked
              -- TODO: instantiateFull?
              forall (tcm :: * -> *) a.
(MonadTCEnv tcm, ReadTCState tcm) =>
tcm a -> tcm a
inTopContext forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *).
(MonadConstraint m, MonadTCState m) =>
QName -> [Clause] -> m ()
addClauses QName
name [Clause
c]
              forall (m :: * -> *) a. Monad m => a -> m a
return (Clause
c,ClausesPostChecks
b)

        ([Clause]
cs, CPC IntSet
isOneIxs) <- forall (m :: * -> *) a. Monad m => a -> m a
return forall a b. (a -> b) -> a -> b
$ (forall (p :: * -> * -> *) b c a.
Bifunctor p =>
(b -> c) -> p a b -> p a c
second forall a. Monoid a => [a] -> a
mconcat forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall a b. [(a, b)] -> ([a], [b])
unzip) [(Clause, ClausesPostChecks)]
cs

        let isSystem :: Bool
isSystem = Bool -> Bool
not forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall a. Null a => a -> Bool
null forall a b. (a -> b) -> a -> b
$ IntSet
isOneIxs

        Bool
canBeSystem <- do
          -- allow VarP and ConP i0/i1 fallThrough = yes, DotP
          let pss :: [[NamedArg DeBruijnPattern]]
pss = forall a b. (a -> b) -> [a] -> [b]
map Clause -> [NamedArg DeBruijnPattern]
namedClausePats [Clause]
cs
              allowed :: Pattern' x -> Bool
allowed = \case
                VarP{} -> Bool
True
                -- pattern inserted by splitPartial
                ConP ConHead
_ ConPatternInfo
cpi [] | ConPatternInfo -> Bool
conPFallThrough ConPatternInfo
cpi -> Bool
True
                DotP{} -> Bool
True
                Pattern' x
_      -> Bool
False
          forall (m :: * -> *) a. Monad m => a -> m a
return forall a b. (a -> b) -> a -> b
$! forall (t :: * -> *) a. Foldable t => (a -> Bool) -> t a -> Bool
all (forall {x}. Pattern' x -> Bool
allowed forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall a. NamedArg a -> a
namedArg) (forall (t :: * -> *) a. Foldable t => t [a] -> [a]
concat [[NamedArg DeBruijnPattern]]
pss)
        forall (f :: * -> *). Applicative f => Bool -> f () -> f ()
when Bool
isSystem forall a b. (a -> b) -> a -> b
$ forall (f :: * -> *). Applicative f => Bool -> f () -> f ()
unless Bool
canBeSystem forall a b. (a -> b) -> a -> b
$
          forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
TypeError -> m a
typeError forall a b. (a -> b) -> a -> b
$ [Char] -> TypeError
GenericError [Char]
"no pattern matching or path copatterns in systems!"


        forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> TCMT IO Doc -> m ()
reportSDoc [Char]
"tc.def.fun" Int
70 forall a b. (a -> b) -> a -> b
$ forall (tcm :: * -> *) a.
(MonadTCEnv tcm, ReadTCState tcm) =>
tcm a -> tcm a
inTopContext forall a b. (a -> b) -> a -> b
$ do
          forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
sep forall a b. (a -> b) -> a -> b
$ TCMT IO Doc
"checked clauses:" forall a. a -> [a] -> [a]
: forall a b. (a -> b) -> [a] -> [b]
map (forall (m :: * -> *). Functor m => Int -> m Doc -> m Doc
nest Int
2 forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall (m :: * -> *). Applicative m => [Char] -> m Doc
text forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall a. Show a => a -> [Char]
show) [Clause]
cs

        -- After checking, remove the clauses again.
        -- (Otherwise, @checkInjectivity@ loops for issue 801).
        QName -> ([Clause] -> [Clause]) -> TCM ()
modifyFunClauses QName
name (forall a b. a -> b -> a
const [])

        forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> TCMT IO Doc -> m ()
reportSDoc [Char]
"tc.cc" Int
25 forall a b. (a -> b) -> a -> b
$ forall (tcm :: * -> *) a.
(MonadTCEnv tcm, ReadTCState tcm) =>
tcm a -> tcm a
inTopContext forall a b. (a -> b) -> a -> b
$ do
          forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
sep [ TCMT IO Doc
"clauses before injectivity test"
              , forall (m :: * -> *). Functor m => Int -> m Doc -> m Doc
nest Int
2 forall a b. (a -> b) -> a -> b
$ forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM forall a b. (a -> b) -> a -> b
$ forall a b. (a -> b) -> [a] -> [b]
map (forall a. QName -> a -> QNamed a
QNamed QName
name) [Clause]
cs  -- broken, reify (QNamed n cl) expect cl to live at top level
              ]
        forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> TCMT IO Doc -> m ()
reportSDoc [Char]
"tc.cc" Int
60 forall a b. (a -> b) -> a -> b
$ forall (tcm :: * -> *) a.
(MonadTCEnv tcm, ReadTCState tcm) =>
tcm a -> tcm a
inTopContext forall a b. (a -> b) -> a -> b
$ do
          forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
sep [ TCMT IO Doc
"raw clauses: "
              , forall (m :: * -> *). Functor m => Int -> m Doc -> m Doc
nest Int
2 forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
sep forall a b. (a -> b) -> a -> b
$ forall a b. (a -> b) -> [a] -> [b]
map (forall (m :: * -> *). Applicative m => [Char] -> m Doc
text forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall a. Show a => a -> [Char]
show forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall a. QName -> a -> QNamed a
QNamed QName
name) [Clause]
cs
              ]

        -- Needed to calculate the proper fullType below.
        forall (tcm :: * -> *) m a.
(MonadTCEnv tcm, LensCohesion m) =>
m -> tcm a -> tcm a
applyCohesionToContext ArgInfo
ai forall a b. (a -> b) -> a -> b
$ do

        -- Systems have their own coverage and "coherence" check, we
        -- also add an absurd clause for the cases not needed.
        ([Clause]
cs,Maybe System
sys) <- if Bool -> Bool
not Bool
isSystem then forall (m :: * -> *) a. Monad m => a -> m a
return ([Clause]
cs, forall a. Null a => a
empty) else do
                 Type
fullType <- forall a b c. (a -> b -> c) -> b -> a -> c
flip forall t. Abstract t => Tele (Dom Type) -> t -> t
abstract Type
t forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> forall (m :: * -> *).
(Applicative m, MonadTCEnv m) =>
m (Tele (Dom Type))
getContextTelescope
                 System
sys <- forall (tcm :: * -> *) a.
(MonadTCEnv tcm, ReadTCState tcm) =>
tcm a -> tcm a
inTopContext forall a b. (a -> b) -> a -> b
$ QName -> [Int] -> Type -> [Clause] -> TCMT IO System
checkSystemCoverage QName
name (IntSet -> [Int]
IntSet.toList IntSet
isOneIxs) Type
fullType [Clause]
cs
                 Tele (Dom Type)
tel <- forall (m :: * -> *).
(Applicative m, MonadTCEnv m) =>
m (Tele (Dom Type))
getContextTelescope
                 let c :: Clause
c = Clause
                       { clauseFullRange :: Range
clauseFullRange = forall a. Range' a
noRange
                       , clauseLHSRange :: Range
clauseLHSRange  = forall a. Range' a
noRange
                       , clauseTel :: Tele (Dom Type)
clauseTel       = Tele (Dom Type)
tel
                       , namedClausePats :: [NamedArg DeBruijnPattern]
namedClausePats = forall a t. DeBruijn a => Tele (Dom t) -> [NamedArg a]
teleNamedArgs Tele (Dom Type)
tel
                       , clauseBody :: Maybe Term
clauseBody      = forall a. Maybe a
Nothing
                       , clauseType :: Maybe (Arg Type)
clauseType      = forall a. a -> Maybe a
Just (forall a. a -> Arg a
defaultArg Type
t)
                       , clauseCatchall :: Bool
clauseCatchall    = Bool
False
                       , clauseExact :: Maybe Bool
clauseExact       = forall a. a -> Maybe a
Just Bool
True
                       , clauseRecursive :: Maybe Bool
clauseRecursive   = forall a. a -> Maybe a
Just Bool
False
                       , clauseUnreachable :: Maybe Bool
clauseUnreachable = forall a. a -> Maybe a
Just Bool
False
                       , clauseEllipsis :: ExpandedEllipsis
clauseEllipsis    = ExpandedEllipsis
NoEllipsis
                       , clauseWhereModule :: Maybe ModuleName
clauseWhereModule = forall a. Maybe a
Nothing
                       }
                 forall (m :: * -> *) a. Monad m => a -> m a
return ([Clause]
cs forall a. [a] -> [a] -> [a]
++ [Clause
c], forall (f :: * -> *) a. Applicative f => a -> f a
pure System
sys)

        -- Annotate the clauses with which arguments are actually used.
        [Clause]
cs <- forall a (m :: * -> *).
(InstantiateFull a, MonadReduce m) =>
a -> m a
instantiateFull {- =<< mapM rebindClause -} [Clause]
cs
        -- Andreas, 2010-11-12
        -- rebindClause is the identity, and instantiateFull eta-contracts
        -- removing this eta-contraction fixes issue 361
        -- however, Data.Star.Decoration.gmapAll no longer type-checks
        -- possibly due to missing eta-contraction!?

        -- Inline copattern record constructors on demand.
        [Clause]
cs <- forall (t :: * -> *) a. Foldable t => t [a] -> [a]
concat forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> do
          forall (t :: * -> *) (m :: * -> *) a b.
(Traversable t, Monad m) =>
t a -> (a -> m b) -> m (t b)
forM [Clause]
cs forall a b. (a -> b) -> a -> b
$ \ Clause
cl -> do
            ([Clause]
cls, Bool
nonExactSplit) <- forall (m :: * -> *) a. Functor m => ChangeT m a -> m (a, Bool)
runChangeT forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *).
(MonadChange m, PureTCM m) =>
Clause -> m [Clause]
recordRHSToCopatterns Clause
cl
            forall (f :: * -> *). Applicative f => Bool -> f () -> f ()
when Bool
nonExactSplit do
              -- If we inlined a non-eta constructor,
              -- issue a warning that the clause does not hold as definitional equality.
              forall (m :: * -> *).
(HasCallStack, MonadWarning m) =>
Warning -> m ()
warning forall a b. (a -> b) -> a -> b
$ QName -> Clause -> Warning
InlineNoExactSplit QName
name Clause
cl
            forall (m :: * -> *) a. Monad m => a -> m a
return [Clause]
cls

        -- Check if the function is injective.
        -- Andreas, 2015-07-01 we do it here in order to resolve metas
        -- in mutual definitions, e.g. the U/El definition in succeed/Issue439.agda
        -- We do it again for the mutual block after termination checking, see Rules.Decl.
        forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> [Char] -> m ()
reportSLn [Char]
"tc.inj.def" Int
20 forall a b. (a -> b) -> a -> b
$ [Char]
"checkFunDef': checking injectivity..."
        FunctionInverse
inv <- forall (m :: * -> *) c.
MonadBench m =>
Account (BenchPhase m) -> m c -> m c
Bench.billTo [Phase
Bench.Injectivity] forall a b. (a -> b) -> a -> b
$
          QName -> [Clause] -> TCMT IO FunctionInverse
checkInjectivity QName
name [Clause]
cs

        forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> TCMT IO Doc -> m ()
reportSDoc [Char]
"tc.cc" Int
15 forall a b. (a -> b) -> a -> b
$ forall (tcm :: * -> *) a.
(MonadTCEnv tcm, ReadTCState tcm) =>
tcm a -> tcm a
inTopContext forall a b. (a -> b) -> a -> b
$ do
          forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
sep [ TCMT IO Doc
"clauses before compilation"
              , forall (m :: * -> *). Functor m => Int -> m Doc -> m Doc
nest Int
2 forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
sep forall a b. (a -> b) -> a -> b
$ forall a b. (a -> b) -> [a] -> [b]
map (forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall a. QName -> a -> QNamed a
QNamed QName
name) [Clause]
cs
              ]

        forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> TCMT IO Doc -> m ()
reportSDoc [Char]
"tc.cc.raw" Int
65 forall a b. (a -> b) -> a -> b
$ do
          forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
sep [ TCMT IO Doc
"clauses before compilation"
              , forall (m :: * -> *). Functor m => Int -> m Doc -> m Doc
nest Int
2 forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
sep forall a b. (a -> b) -> a -> b
$ forall a b. (a -> b) -> [a] -> [b]
map (forall (m :: * -> *). Applicative m => [Char] -> m Doc
text forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall a. Show a => a -> [Char]
show) [Clause]
cs
              ]

        -- add clauses for the coverage (& confluence) checker (needs to reduce)
        forall (tcm :: * -> *) a.
(MonadTCEnv tcm, ReadTCState tcm) =>
tcm a -> tcm a
inTopContext forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *).
(MonadConstraint m, MonadTCState m) =>
QName -> [Clause] -> m ()
addClauses QName
name [Clause]
cs

        forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> TCMT IO Doc -> m ()
reportSDoc [Char]
"tc.cc.type" Int
60 forall a b. (a -> b) -> a -> b
$ TCMT IO Doc
"  type   : " forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> (forall (m :: * -> *). Applicative m => [Char] -> m Doc
text forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall a. Pretty a => a -> [Char]
prettyShow) Type
t
        forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> TCMT IO Doc -> m ()
reportSDoc [Char]
"tc.cc.type" Int
60 forall a b. (a -> b) -> a -> b
$ TCMT IO Doc
"  context: " forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> (forall (m :: * -> *). Applicative m => [Char] -> m Doc
text forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall a. Pretty a => a -> [Char]
prettyShow forall (m :: * -> *) a b. Monad m => (a -> m b) -> m a -> m b
=<< forall (m :: * -> *).
(Applicative m, MonadTCEnv m) =>
m (Tele (Dom Type))
getContextTelescope)

        Type
fullType <- forall a b c. (a -> b -> c) -> b -> a -> c
flip Tele (Dom Type) -> Type -> Type
telePi Type
t forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> forall (m :: * -> *).
(Applicative m, MonadTCEnv m) =>
m (Tele (Dom Type))
getContextTelescope

        forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> [Char] -> m ()
reportSLn  [Char]
"tc.cc.type" Int
80 forall a b. (a -> b) -> a -> b
$ forall a. Show a => a -> [Char]
show Type
fullType

        -- Coverage check and compile the clauses
        (Maybe SplitTree
mst, Bool
_recordExpressionBecameCopatternLHS, CompiledClauses
cc) <- forall (m :: * -> *) c.
MonadBench m =>
Account (BenchPhase m) -> m c -> m c
Bench.billTo [Phase
Bench.Coverage] forall a b. (a -> b) -> a -> b
$
          forall (tcm :: * -> *) a.
(MonadTCEnv tcm, ReadTCState tcm) =>
tcm a -> tcm a
unsafeInTopContext forall a b. (a -> b) -> a -> b
$ Maybe (QName, Type)
-> [Clause] -> TCMT IO (Maybe SplitTree, Bool, CompiledClauses)
compileClauses (if Bool
isSystem then forall a. Maybe a
Nothing else (forall a. a -> Maybe a
Just (QName
name, Type
fullType)))
                                        [Clause]
cs
        -- Andreas, 2019-10-21 (see also issue #4142):
        -- We ignore whether the clause compilation turned some
        -- record expressions into copatterns
        -- (_recordExpressionsBecameCopatternLHS),
        -- since the defCopatternLHS flag is anyway set by traversing
        -- the compiled clauses looking for a copattern match
        -- (hasProjectionPatterns).

        -- Clause compilation runs the coverage checker, which might add
        -- some extra clauses.
        [Clause]
cs <- Definition -> [Clause]
defClauses forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> forall (m :: * -> *). HasConstInfo m => QName -> m Definition
getConstInfo QName
name

        forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> TCMT IO Doc -> m ()
reportSDoc [Char]
"tc.cc" Int
60 forall a b. (a -> b) -> a -> b
$ forall (tcm :: * -> *) a.
(MonadTCEnv tcm, ReadTCState tcm) =>
tcm a -> tcm a
inTopContext forall a b. (a -> b) -> a -> b
$ do
          forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
sep [ TCMT IO Doc
"compiled clauses of" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM QName
name
              , forall (m :: * -> *). Functor m => Int -> m Doc -> m Doc
nest Int
2 forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *) a. (Applicative m, Pretty a) => a -> m Doc
pretty CompiledClauses
cc
              ]

        -- The macro tag might be on the type signature
        Bool
ismacro <- Defn -> Bool
isMacro forall b c a. (b -> c) -> (a -> b) -> a -> c
. Definition -> Defn
theDef forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> forall (m :: * -> *). HasConstInfo m => QName -> m Definition
getConstInfo QName
name

        [Clause]
covering <- Defn -> [Clause]
funCovering forall b c a. (b -> c) -> (a -> b) -> a -> c
. Definition -> Defn
theDef forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> forall (m :: * -> *). HasConstInfo m => QName -> m Definition
getConstInfo QName
name

        -- Add the definition
        forall (tcm :: * -> *) a.
(MonadTCEnv tcm, ReadTCState tcm) =>
tcm a -> tcm a
inTopContext forall a b. (a -> b) -> a -> b
$ QName -> Definition -> TCM ()
addConstant QName
name forall (m :: * -> *) a b. Monad m => (a -> m b) -> m a -> m b
=<< do

          forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> TCMT IO Doc -> m ()
reportSDoc [Char]
"tc.def.fun.clauses" Int
15 forall a b. (a -> b) -> a -> b
$ forall (tcm :: * -> *) a.
(MonadTCEnv tcm, ReadTCState tcm) =>
tcm a -> tcm a
inTopContext forall a b. (a -> b) -> a -> b
$ do
            forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
vcat [ TCMT IO Doc
"final clauses for" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM QName
name forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+>  TCMT IO Doc
":"
                 , forall (m :: * -> *). Functor m => Int -> m Doc -> m Doc
nest Int
2 forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
vcat forall a b. (a -> b) -> a -> b
$ forall a b. (a -> b) -> [a] -> [b]
map (forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall a. QName -> a -> QNamed a
QNamed QName
name) [Clause]
cs
                 ]

          -- If there was a pragma for this definition, we can set the
          -- funTerminates field directly.
          FunctionData
fun  <- forall (m :: * -> *). HasOptions m => m FunctionData
emptyFunctionData
          Defn
defn <- Defn -> TCM Defn
autoInline forall a b. (a -> b) -> a -> b
$
             forall o i. Lens' o i -> LensSet o i
set Lens' Defn Bool
funMacro (Bool
ismacro Bool -> Bool -> Bool
|| forall t. DefInfo' t -> IsMacro
Info.defMacro DefInfo
i forall a. Eq a => a -> a -> Bool
== IsMacro
MacroDef) forall a b. (a -> b) -> a -> b
$
             FunctionData -> Defn
FunctionDefn FunctionData
fun
             { _funClauses :: [Clause]
_funClauses        = [Clause]
cs
             , _funCompiled :: Maybe CompiledClauses
_funCompiled       = forall a. a -> Maybe a
Just CompiledClauses
cc
             , _funSplitTree :: Maybe SplitTree
_funSplitTree      = Maybe SplitTree
mst
             , _funInv :: FunctionInverse
_funInv            = FunctionInverse
inv
             , _funAbstr :: IsAbstract
_funAbstr          = forall t. DefInfo' t -> IsAbstract
Info.defAbstract DefInfo
i
             , _funOpaque :: IsOpaque
_funOpaque         = forall t. DefInfo' t -> IsOpaque
Info.defOpaque DefInfo
i
             , _funExtLam :: Maybe ExtLamInfo
_funExtLam         = (\ ExtLamInfo
e -> ExtLamInfo
e { extLamSys :: Maybe System
extLamSys = Maybe System
sys }) forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> Maybe ExtLamInfo
extlam
             , _funWith :: Maybe QName
_funWith           = Maybe QName
with
             , _funCovering :: [Clause]
_funCovering       = [Clause]
covering
             }
          Language
lang <- forall (m :: * -> *). HasOptions m => m Language
getLanguage
          Definition -> TCMT IO Definition
useTerPragma forall a b. (a -> b) -> a -> b
$
            (Bool -> Bool) -> Definition -> Definition
updateDefCopatternLHS (forall a b. a -> b -> a
const forall a b. (a -> b) -> a -> b
$ CompiledClauses -> Bool
hasProjectionPatterns CompiledClauses
cc) forall a b. (a -> b) -> a -> b
$
            (ArgInfo -> QName -> Type -> Language -> Defn -> Definition
defaultDefn ArgInfo
ai QName
name Type
fullType Language
lang Defn
defn)

        forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> TCMT IO Doc -> m ()
reportSDoc [Char]
"tc.def.fun" Int
10 forall a b. (a -> b) -> a -> b
$ do
          forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
sep [ TCMT IO Doc
"added " forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM QName
name forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> TCMT IO Doc
":"
              , forall (m :: * -> *). Functor m => Int -> m Doc -> m Doc
nest Int
2 forall a b. (a -> b) -> a -> b
$ forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM forall b c a. (b -> c) -> (a -> b) -> a -> c
. Definition -> Type
defType forall (m :: * -> *) a b. Monad m => (a -> m b) -> m a -> m b
=<< forall (m :: * -> *). HasConstInfo m => QName -> m Definition
getConstInfo QName
name
              ]

        -- Jesper, 2019-05-30: if the constructors used in the
        -- lhs of a clause have rewrite rules, we need to check
        -- confluence here
        forall (m :: * -> *) a.
Monad m =>
m (Maybe a) -> (a -> m ()) -> m ()
whenJustM (PragmaOptions -> Maybe ConfluenceCheck
optConfluenceCheck forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> forall (m :: * -> *). HasOptions m => m PragmaOptions
pragmaOptions) forall a b. (a -> b) -> a -> b
$ \ConfluenceCheck
confChk -> forall (tcm :: * -> *) a.
(MonadTCEnv tcm, ReadTCState tcm) =>
tcm a -> tcm a
inTopContext forall a b. (a -> b) -> a -> b
$
          ConfluenceCheck -> QName -> TCM ()
checkConfluenceOfClauses ConfluenceCheck
confChk QName
name

-- | Set 'funTerminates' according to termination info in 'TCEnv',
--   which comes from a possible termination pragma.
useTerPragma :: Definition -> TCM Definition
useTerPragma :: Definition -> TCMT IO Definition
useTerPragma def :: Definition
def@Defn{ defName :: Definition -> QName
defName = QName
name, theDef :: Definition -> Defn
theDef = fun :: Defn
fun@Function{}} = do
  TerminationCheck ()
tc <- forall (m :: * -> *) a. MonadTCEnv m => Lens' TCEnv a -> m a
viewTC Lens' TCEnv (TerminationCheck ())
eTerminationCheck
  let terminates :: Maybe Bool
terminates = case TerminationCheck ()
tc of
        TerminationCheck ()
NonTerminating -> forall a. a -> Maybe a
Just Bool
False
        TerminationCheck ()
Terminating    -> forall a. a -> Maybe a
Just Bool
True
        TerminationCheck ()
_              -> forall a. Maybe a
Nothing
  forall a (m :: * -> *).
(ReportS a, MonadDebug m) =>
[Char] -> Int -> a -> m ()
reportS [Char]
"tc.fundef" Int
30 forall a b. (a -> b) -> a -> b
$
    [ [Char]
"funTerminates of " forall a. [a] -> [a] -> [a]
++ forall a. Pretty a => a -> [Char]
prettyShow QName
name forall a. [a] -> [a] -> [a]
++ [Char]
" set to " forall a. [a] -> [a] -> [a]
++ forall a. Show a => a -> [Char]
show Maybe Bool
terminates
    , [Char]
"  tc = " forall a. [a] -> [a] -> [a]
++ forall a. Show a => a -> [Char]
show TerminationCheck ()
tc
    ]
  forall (m :: * -> *) a. Monad m => a -> m a
return forall a b. (a -> b) -> a -> b
$ Definition
def { theDef :: Defn
theDef = Defn
fun { funTerminates :: Maybe Bool
funTerminates = Maybe Bool
terminates }}
useTerPragma Definition
def = forall (m :: * -> *) a. Monad m => a -> m a
return Definition
def

-- | Modify all the LHSCore of the given RHS.
-- (Used to insert patterns for @rewrite@ or the inspect idiom)
mapLHSCores :: (A.LHSCore -> A.LHSCore) -> (A.RHS -> A.RHS)
mapLHSCores :: (LHSCore' Expr -> LHSCore' Expr) -> RHS -> RHS
mapLHSCores LHSCore' Expr -> LHSCore' Expr
f = \case
  A.WithRHS QName
aux [WithExpr]
es List1 (Clause' LHS)
cs -> QName -> [WithExpr] -> List1 (Clause' LHS) -> RHS
A.WithRHS QName
aux [WithExpr]
es forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *) a b. Functor m => m a -> (a -> b) -> m b
for List1 (Clause' LHS)
cs forall a b. (a -> b) -> a -> b
$
    \ (A.Clause (A.LHS LHSInfo
info LHSCore' Expr
core)     [ProblemEq]
spats RHS
rhs                 WhereDeclarations
ds Bool
catchall) ->
       forall lhs.
lhs
-> [ProblemEq] -> RHS -> WhereDeclarations -> Bool -> Clause' lhs
A.Clause (LHSInfo -> LHSCore' Expr -> LHS
A.LHS LHSInfo
info (LHSCore' Expr -> LHSCore' Expr
f LHSCore' Expr
core)) [ProblemEq]
spats ((LHSCore' Expr -> LHSCore' Expr) -> RHS -> RHS
mapLHSCores LHSCore' Expr -> LHSCore' Expr
f RHS
rhs) WhereDeclarations
ds Bool
catchall
  A.RewriteRHS [RewriteEqn]
qes [ProblemEq]
spats RHS
rhs WhereDeclarations
wh -> [RewriteEqn] -> [ProblemEq] -> RHS -> WhereDeclarations -> RHS
A.RewriteRHS [RewriteEqn]
qes [ProblemEq]
spats ((LHSCore' Expr -> LHSCore' Expr) -> RHS -> RHS
mapLHSCores LHSCore' Expr -> LHSCore' Expr
f RHS
rhs) WhereDeclarations
wh
  rhs :: RHS
rhs@RHS
A.AbsurdRHS -> RHS
rhs
  rhs :: RHS
rhs@A.RHS{}     -> RHS
rhs

-- | Insert some names into the with-clauses LHS of the given RHS.
-- (Used for the inspect idiom)
insertNames :: [Arg (Maybe A.BindName)] -> A.RHS -> A.RHS
insertNames :: [Arg (Maybe BindName)] -> RHS -> RHS
insertNames = (LHSCore' Expr -> LHSCore' Expr) -> RHS -> RHS
mapLHSCores forall b c a. (b -> c) -> (a -> b) -> a -> c
. [Arg (Maybe BindName)] -> LHSCore' Expr -> LHSCore' Expr
insertInspects

insertInspects :: [Arg (Maybe A.BindName)] -> A.LHSCore -> A.LHSCore
insertInspects :: [Arg (Maybe BindName)] -> LHSCore' Expr -> LHSCore' Expr
insertInspects [Arg (Maybe BindName)]
ps = \case
  A.LHSWith LHSCore' Expr
core [Arg (Pattern' Expr)]
wps [] ->
    let ps' :: [Arg (Maybe (Arg (Pattern' Expr)))]
ps' = forall a b. (a -> b) -> [a] -> [b]
map (forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
fmap forall a b. (a -> b) -> a -> b
$ forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
fmap BindName -> Arg (Pattern' Expr)
patOfName) [Arg (Maybe BindName)]
ps in
    forall e.
LHSCore' e
-> [Arg (Pattern' e)] -> [NamedArg (Pattern' e)] -> LHSCore' e
A.LHSWith LHSCore' Expr
core (forall a. [Arg (Maybe (Arg a))] -> [Arg a] -> [Arg a]
insertIn [Arg (Maybe (Arg (Pattern' Expr)))]
ps' [Arg (Pattern' Expr)]
wps) []
  -- Andreas, AIM XXXV, 2022-05-09, issue #5728:
  -- Cases other than LHSWith actually do not make sense, but let them
  -- through to get a proper error later.
  LHSCore' Expr
lhs -> LHSCore' Expr
lhs

  where

    patOfName :: A.BindName -> Arg A.Pattern
    patOfName :: BindName -> Arg (Pattern' Expr)
patOfName = forall a. a -> Arg a
defaultArg forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall e. BindName -> Pattern' e
A.VarP

    insertIn :: [Arg (Maybe (Arg a))]
             -> [Arg a] -> [Arg a]
    insertIn :: forall a. [Arg (Maybe (Arg a))] -> [Arg a] -> [Arg a]
insertIn []                 [Arg a]
wps  = [Arg a]
wps
    insertIn (Arg ArgInfo
info Maybe (Arg a)
nm : [Arg (Maybe (Arg a))]
ps) (Arg a
w : [Arg a]
wps) | forall a. LensHiding a => a -> Bool
visible ArgInfo
info =
      Arg a
w forall a. a -> [a] -> [a]
: (forall b a. b -> (a -> b) -> Maybe a -> b
maybe [] forall (f :: * -> *) a. Applicative f => a -> f a
pure Maybe (Arg a)
nm) forall a. [a] -> [a] -> [a]
++ forall a. [Arg (Maybe (Arg a))] -> [Arg a] -> [Arg a]
insertIn [Arg (Maybe (Arg a))]
ps [Arg a]
wps
    insertIn (Arg ArgInfo
info Maybe (Arg a)
nm : [Arg (Maybe (Arg a))]
ps) [Arg a]
wps       | forall a. LensHiding a => a -> Bool
notVisible ArgInfo
info =
          (forall b a. b -> (a -> b) -> Maybe a -> b
maybe [] forall (f :: * -> *) a. Applicative f => a -> f a
pure Maybe (Arg a)
nm) forall a. [a] -> [a] -> [a]
++ forall a. [Arg (Maybe (Arg a))] -> [Arg a] -> [Arg a]
insertIn [Arg (Maybe (Arg a))]
ps [Arg a]
wps
    insertIn [Arg (Maybe (Arg a))]
_ [Arg a]
_ = forall a. HasCallStack => a
__IMPOSSIBLE__


-- | Insert some with-patterns into the with-clauses LHS of the given RHS.
-- (Used for @rewrite@)
insertPatterns :: [Arg A.Pattern] -> A.RHS -> A.RHS
insertPatterns :: [Arg (Pattern' Expr)] -> RHS -> RHS
insertPatterns [Arg (Pattern' Expr)]
pats = (LHSCore' Expr -> LHSCore' Expr) -> RHS -> RHS
mapLHSCores ([Arg (Pattern' Expr)] -> LHSCore' Expr -> LHSCore' Expr
insertPatternsLHSCore [Arg (Pattern' Expr)]
pats)

-- | Insert with-patterns before the trailing with patterns.
-- If there are none, append the with-patterns.
insertPatternsLHSCore :: [Arg A.Pattern] -> A.LHSCore -> A.LHSCore
insertPatternsLHSCore :: [Arg (Pattern' Expr)] -> LHSCore' Expr -> LHSCore' Expr
insertPatternsLHSCore [Arg (Pattern' Expr)]
pats = \case
  A.LHSWith LHSCore' Expr
core [Arg (Pattern' Expr)]
wps [] -> forall e.
LHSCore' e
-> [Arg (Pattern' e)] -> [NamedArg (Pattern' e)] -> LHSCore' e
A.LHSWith LHSCore' Expr
core ([Arg (Pattern' Expr)]
pats forall a. [a] -> [a] -> [a]
++ [Arg (Pattern' Expr)]
wps) []
  LHSCore' Expr
core                  -> forall e.
LHSCore' e
-> [Arg (Pattern' e)] -> [NamedArg (Pattern' e)] -> LHSCore' e
A.LHSWith LHSCore' Expr
core [Arg (Pattern' Expr)]
pats []

-- | Parameters for creating a @with@-function.
data WithFunctionProblem
  = NoWithFunction
  | WithFunction
    { WithFunctionProblem -> QName
wfParentName :: QName                             -- ^ Parent function name.
    , WithFunctionProblem -> QName
wfName       :: QName                             -- ^ With function name.
    , WithFunctionProblem -> Type
wfParentType :: Type                              -- ^ Type of the parent function.
    , WithFunctionProblem -> Tele (Dom Type)
wfParentTel  :: Telescope                         -- ^ Context of the parent patterns.
    , WithFunctionProblem -> Tele (Dom Type)
wfBeforeTel  :: Telescope                         -- ^ Types of arguments to the with function before the with expressions (needed vars).
    , WithFunctionProblem -> Tele (Dom Type)
wfAfterTel   :: Telescope                         -- ^ Types of arguments to the with function after the with expressions (unneeded vars).
    , WithFunctionProblem -> [Arg (Term, EqualityView)]
wfExprs      :: [Arg (Term, EqualityView)]        -- ^ With and rewrite expressions and their types.
    , WithFunctionProblem -> Type
wfRHSType    :: Type                              -- ^ Type of the right hand side.
    , WithFunctionProblem -> [NamedArg DeBruijnPattern]
wfParentPats :: [NamedArg DeBruijnPattern]        -- ^ Parent patterns.
    , WithFunctionProblem -> Int
wfParentParams :: Nat                             -- ^ Number of module parameters in parent patterns
    , WithFunctionProblem -> Permutation
wfPermSplit  :: Permutation                       -- ^ Permutation resulting from splitting the telescope into needed and unneeded vars.
    , WithFunctionProblem -> Permutation
wfPermParent :: Permutation                       -- ^ Permutation reordering the variables in the parent pattern.
    , WithFunctionProblem -> Permutation
wfPermFinal  :: Permutation                       -- ^ Final permutation (including permutation for the parent clause).
    , WithFunctionProblem -> List1 (Clause' LHS)
wfClauses    :: List1 A.Clause                    -- ^ The given clauses for the with function
    , WithFunctionProblem -> Substitution
wfCallSubst :: Substitution                       -- ^ Subtsitution to generate call for the parent.
    }

checkSystemCoverage
  :: QName
  -> [Int]
  -> Type
  -> [Clause]
  -> TCM System
checkSystemCoverage :: QName -> [Int] -> Type -> [Clause] -> TCMT IO System
checkSystemCoverage QName
f [Int
n] Type
t [Clause]
cs = do
  forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> TCMT IO Doc -> m ()
reportSDoc [Char]
"tc.sys.cover" Int
10 forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *). Applicative m => [Char] -> m Doc
text (forall a. Show a => a -> [Char]
show (Int
n , forall (t :: * -> *) a. Foldable t => t a -> Int
length [Clause]
cs)) forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Type
t
  TelV Tele (Dom Type)
gamma Type
t <- forall (m :: * -> *).
(MonadReduce m, MonadAddContext m) =>
Int -> Type -> m (TelV Type)
telViewUpTo Int
n Type
t
  forall b (m :: * -> *) a.
(AddContext b, MonadAddContext m) =>
b -> m a -> m a
addContext Tele (Dom Type)
gamma forall a b. (a -> b) -> a -> b
$ do
  TelV (ExtendTel Dom Type
a Abs (Tele (Dom Type))
_) Type
_ <- forall (m :: * -> *).
(MonadReduce m, MonadAddContext m) =>
Int -> Type -> m (TelV Type)
telViewUpTo Int
1 Type
t
  Term
a <- forall a (m :: * -> *). (Reduce a, MonadReduce m) => a -> m a
reduce forall a b. (a -> b) -> a -> b
$ forall t a. Type'' t a -> a
unEl forall a b. (a -> b) -> a -> b
$ forall t e. Dom' t e -> e
unDom Dom Type
a

  case Term
a of
    Def QName
q [Apply Arg Term
phi] -> do
      [Maybe QName
iz,Maybe QName
io] <- forall (t :: * -> *) (m :: * -> *) a b.
(Traversable t, Monad m) =>
(a -> m b) -> t a -> m (t b)
mapM forall (m :: * -> *). HasBuiltins m => BuiltinId -> m (Maybe QName)
getBuiltinName' [BuiltinId
builtinIZero, BuiltinId
builtinIOne]
      Term
ineg <- forall (m :: * -> *).
(HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) =>
m Term
primINeg
      Term
imin <- forall (m :: * -> *).
(HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) =>
m Term
primIMin
      Term
imax <- forall (m :: * -> *).
(HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) =>
m Term
primIMax
      Term
i0 <- forall (m :: * -> *).
(HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) =>
m Term
primIZero
      Term
i1 <- forall (m :: * -> *).
(HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) =>
m Term
primIOne
      let
        isDir :: DeBruijnPattern -> Maybe Bool
isDir (ConP ConHead
q ConPatternInfo
_ []) | forall a. a -> Maybe a
Just (ConHead -> QName
conName ConHead
q) forall a. Eq a => a -> a -> Bool
== Maybe QName
iz = forall a. a -> Maybe a
Just Bool
False
        isDir (ConP ConHead
q ConPatternInfo
_ []) | forall a. a -> Maybe a
Just (ConHead -> QName
conName ConHead
q) forall a. Eq a => a -> a -> Bool
== Maybe QName
io = forall a. a -> Maybe a
Just Bool
True
        isDir DeBruijnPattern
_ = forall a. Maybe a
Nothing

        collectDirs :: [Int] -> [DeBruijnPattern] -> [(Int,Bool)]
        collectDirs :: [Int] -> [DeBruijnPattern] -> [(Int, Bool)]
collectDirs [] [] = []
        collectDirs (Int
i : [Int]
is) (DeBruijnPattern
p : [DeBruijnPattern]
ps) | Just Bool
d <- DeBruijnPattern -> Maybe Bool
isDir DeBruijnPattern
p = (Int
i,Bool
d) forall a. a -> [a] -> [a]
: [Int] -> [DeBruijnPattern] -> [(Int, Bool)]
collectDirs [Int]
is [DeBruijnPattern]
ps
                                      | Bool
otherwise         = [Int] -> [DeBruijnPattern] -> [(Int, Bool)]
collectDirs [Int]
is [DeBruijnPattern]
ps
        collectDirs [Int]
_ [DeBruijnPattern]
_ = forall a. HasCallStack => a
__IMPOSSIBLE__

        dir :: (Int,Bool) -> Term
        dir :: (Int, Bool) -> Term
dir (Int
i,Bool
False) = Term
ineg forall t. Apply t => t -> [Arg Term] -> t
`apply` [forall a. a -> Arg a
argN forall a b. (a -> b) -> a -> b
$ Int -> Term
var Int
i]
        dir (Int
i,Bool
True) = Int -> Term
var Int
i

        -- andI and orI have cases for singletons to improve error messages.
        andI, orI :: [Term] -> Term
        andI :: [Term] -> Term
andI [] = Term
i1
        andI [Term
t] = Term
t
        andI (Term
t:[Term]
ts) = (\ Term
x -> Term
imin forall t. Apply t => t -> [Arg Term] -> t
`apply` [forall a. a -> Arg a
argN Term
t, forall a. a -> Arg a
argN Term
x]) forall a b. (a -> b) -> a -> b
$ [Term] -> Term
andI [Term]
ts

        orI :: [Term] -> Term
orI [] = Term
i0
        orI [Term
t] = Term
t
        orI (Term
t:[Term]
ts) = Term
imax forall t. Apply t => t -> [Arg Term] -> t
`apply` [forall a. a -> Arg a
argN Term
t, forall a. a -> Arg a
argN ([Term] -> Term
orI [Term]
ts)]

      let
        pats :: [[DeBruijnPattern]]
pats = forall a b. (a -> b) -> [a] -> [b]
map (forall a. Int -> [a] -> [a]
take Int
n forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall a b. (a -> b) -> [a] -> [b]
map (forall name a. Named name a -> a
namedThing forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall e. Arg e -> e
unArg) forall b c a. (b -> c) -> (a -> b) -> a -> c
. Clause -> [NamedArg DeBruijnPattern]
namedClausePats) [Clause]
cs
        alphas :: [[(Int,Bool)]] -- the face maps corresponding to each clause
        alphas :: [[(Int, Bool)]]
alphas = forall a b. (a -> b) -> [a] -> [b]
map ([Int] -> [DeBruijnPattern] -> [(Int, Bool)]
collectDirs (forall a. Integral a => a -> [a]
downFrom Int
n)) [[DeBruijnPattern]]
pats
        phis :: [Term] -- the φ terms for each clause (i.e. the alphas as terms)
        phis :: [Term]
phis = forall a b. (a -> b) -> [a] -> [b]
map ([Term] -> Term
andI forall b c a. (b -> c) -> (a -> b) -> a -> c
. (forall a b. (a -> b) -> [a] -> [b]
map (Int, Bool) -> Term
dir)) [[(Int, Bool)]]
alphas
        psi :: Term
psi = [Term] -> Term
orI forall a b. (a -> b) -> a -> b
$ [Term]
phis
        pcs :: [(Term, Clause)]
pcs = forall a b. [a] -> [b] -> [(a, b)]
zip [Term]
phis [Clause]
cs

      forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> TCMT IO Doc -> m ()
reportSDoc [Char]
"tc.sys.cover" Int
20 forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
fsep forall a b. (a -> b) -> a -> b
$ forall a b. (a -> b) -> [a] -> [b]
map forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM [[DeBruijnPattern]]
pats
      Type
interval <- forall (m :: * -> *).
(HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) =>
m Type
primIntervalType
      forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> TCMT IO Doc -> m ()
reportSDoc [Char]
"tc.sys.cover" Int
10 forall a b. (a -> b) -> a -> b
$ TCMT IO Doc
"equalTerm " forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM (forall e. Arg e -> e
unArg Arg Term
phi) forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Term
psi
      forall (m :: * -> *).
MonadConversion m =>
Type -> Term -> Term -> m ()
equalTerm Type
interval (forall e. Arg e -> e
unArg Arg Term
phi) Term
psi

      forall (t :: * -> *) (m :: * -> *) a b.
(Foldable t, Monad m) =>
t a -> (a -> m b) -> m ()
forM_ (forall a. [a] -> [a] -> [a]
initWithDefault forall a. HasCallStack => a
__IMPOSSIBLE__ forall a b. (a -> b) -> a -> b
$
             forall a. [a] -> [a] -> [a]
initWithDefault forall a. HasCallStack => a
__IMPOSSIBLE__ forall a b. (a -> b) -> a -> b
$ forall a. [a] -> [[a]]
List.tails [(Term, Clause)]
pcs) forall a b. (a -> b) -> a -> b
$ \ ((Term
phi1,Clause
cl1):[(Term, Clause)]
pcs') -> do
        forall (t :: * -> *) (m :: * -> *) a b.
(Foldable t, Monad m) =>
t a -> (a -> m b) -> m ()
forM_ [(Term, Clause)]
pcs' forall a b. (a -> b) -> a -> b
$ \ (Term
phi2,Clause
cl2) -> do
          Term
phi12 <- forall a (m :: * -> *). (Reduce a, MonadReduce m) => a -> m a
reduce (Term
imin forall t. Apply t => t -> [Arg Term] -> t
`apply` [forall a. a -> Arg a
argN Term
phi1, forall a. a -> Arg a
argN Term
phi2])
          forall (m :: * -> *) a.
MonadConversion m =>
Term
-> (IntMap Bool -> Blocker -> Term -> m a)
-> (IntMap Bool -> Substitution -> m a)
-> m [a]
forallFaceMaps Term
phi12 (\ IntMap Bool
_ Blocker
_ -> forall a. HasCallStack => a
__IMPOSSIBLE__) forall a b. (a -> b) -> a -> b
$ \IntMap Bool
_ Substitution
sigma -> do
            let args :: [Arg Term]
args = Substitution
sigma forall a. Subst a => Substitution' (SubstArg a) -> a -> a
`applySubst` forall a t. DeBruijn a => Tele (Dom t) -> [Arg a]
teleArgs Tele (Dom Type)
gamma
                t' :: Type
t' = Substitution
sigma forall a. Subst a => Substitution' (SubstArg a) -> a -> a
`applySubst` Type
t
                fromReduced :: Reduced (Blocked' t yes) yes -> yes
fromReduced (YesReduction Simplification
_ yes
x) = yes
x
                fromReduced (NoReduction Blocked' t yes
x) = forall t a. Blocked' t a -> a
ignoreBlocking Blocked' t yes
x
                body :: Clause -> TCMT IO Term
body Clause
cl = do
                  let extra :: Int
extra = forall (t :: * -> *) a. Foldable t => t a -> Int
length (forall a. Int -> [a] -> [a]
drop Int
n forall a b. (a -> b) -> a -> b
$ Clause -> [NamedArg DeBruijnPattern]
namedClausePats Clause
cl)
                  TelV Tele (Dom Type)
delta Type
_ <- forall (m :: * -> *).
(MonadReduce m, MonadAddContext m) =>
Int -> Type -> m (TelV Type)
telViewUpTo Int
extra Type
t'
                  forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
fmap (forall t. Abstract t => Tele (Dom Type) -> t -> t
abstract Tele (Dom Type)
delta) forall a b. (a -> b) -> a -> b
$ forall b (m :: * -> *) a.
(AddContext b, MonadAddContext m) =>
b -> m a -> m a
addContext Tele (Dom Type)
delta forall a b. (a -> b) -> a -> b
$ do
                    forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
fmap forall {t} {yes}. Reduced (Blocked' t yes) yes -> yes
fromReduced forall a b. (a -> b) -> a -> b
$ forall a. ReduceM a -> TCM a
runReduceM forall a b. (a -> b) -> a -> b
$
                      QName
-> Term
-> [Clause]
-> RewriteRules
-> MaybeReducedArgs
-> ReduceM (Reduced (Blocked Term) Term)
appDef' QName
f (QName -> Elims -> Term
Def QName
f []) [Clause
cl] [] (forall a b. (a -> b) -> [a] -> [b]
map forall a. a -> MaybeReduced a
notReduced forall a b. (a -> b) -> a -> b
$ forall a. Subst a => Int -> a -> a
raise (forall a. Sized a => a -> Int
size Tele (Dom Type)
delta) [Arg Term]
args forall a. [a] -> [a] -> [a]
++ forall a t. DeBruijn a => Tele (Dom t) -> [Arg a]
teleArgs Tele (Dom Type)
delta)
            Term
v1 <- Clause -> TCMT IO Term
body Clause
cl1
            Term
v2 <- Clause -> TCMT IO Term
body Clause
cl2
            forall (m :: * -> *).
MonadConversion m =>
Type -> Term -> Term -> m ()
equalTerm Type
t' Term
v1 Term
v2

      [([(Term, Bool)], Term)]
sys <- forall (t :: * -> *) (m :: * -> *) a b.
(Traversable t, Monad m) =>
t a -> (a -> m b) -> m (t b)
forM (forall a b. [a] -> [b] -> [(a, b)]
zip [[(Int, Bool)]]
alphas [Clause]
cs) forall a b. (a -> b) -> a -> b
$ \ ([(Int, Bool)]
alpha,Clause
cl) -> do

            let
                -- Δ = Γ_α , Δ'α
                delta :: Tele (Dom Type)
delta = Clause -> Tele (Dom Type)
clauseTel Clause
cl
                -- Δ ⊢ b
                Just Term
b = Clause -> Maybe Term
clauseBody Clause
cl
                -- Δ ⊢ ps : Γ , o : [φ] , Δ'
                -- we assume that there's no pattern matching other
                -- than from the system
                ps :: [NamedArg DeBruijnPattern]
ps = Clause -> [NamedArg DeBruijnPattern]
namedClausePats Clause
cl
                extra :: Int
extra = forall (t :: * -> *) a. Foldable t => t a -> Int
length (forall a. Int -> [a] -> [a]
drop (forall a. Sized a => a -> Int
size Tele (Dom Type)
gamma forall a. Num a => a -> a -> a
+ Int
1) [NamedArg DeBruijnPattern]
ps)
                -- size Δ'α = size Δ' = extra
                -- Γ , α ⊢ u
                takeLast :: Int -> [a] -> [a]
takeLast Int
n [a]
xs = forall a. Int -> [a] -> [a]
drop (forall (t :: * -> *) a. Foldable t => t a -> Int
length [a]
xs forall a. Num a => a -> a -> a
- Int
n) [a]
xs
                weak :: [Int] -> Substitution' a
weak [] = forall a. Substitution' a
idS
                weak (Int
i:[Int]
is) = [Int] -> Substitution' a
weak [Int]
is forall a.
EndoSubst a =>
Substitution' a -> Substitution' a -> Substitution' a
`composeS` forall a. Int -> Substitution' a -> Substitution' a
liftS Int
i (forall a. Int -> Substitution' a
raiseS Int
1)
                tel :: Tele (Dom Type)
tel = [Dom ([Char], Type)] -> Tele (Dom Type)
telFromList (forall a. Int -> [a] -> [a]
takeLast Int
extra (forall t. Tele (Dom t) -> [Dom ([Char], t)]
telToList Tele (Dom Type)
delta))
                u :: Term
u = forall t. Abstract t => Tele (Dom Type) -> t -> t
abstract Tele (Dom Type)
tel (forall a. Int -> Substitution' a -> Substitution' a
liftS Int
extra (forall {a}. (SubstArg a ~ a, Subst a) => [Int] -> Substitution' a
weak forall a b. (a -> b) -> a -> b
$ forall a. Ord a => [a] -> [a]
List.sort forall a b. (a -> b) -> a -> b
$ forall a b. (a -> b) -> [a] -> [b]
map forall a b. (a, b) -> a
fst [(Int, Bool)]
alpha) forall a. Subst a => Substitution' (SubstArg a) -> a -> a
`applySubst` Term
b)
            forall (m :: * -> *) a. Monad m => a -> m a
return (forall a b. (a -> b) -> [a] -> [b]
map (forall (p :: * -> * -> *) a b c.
Bifunctor p =>
(a -> b) -> p a c -> p b c
first Int -> Term
var) [(Int, Bool)]
alpha,Term
u)

      forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> TCMT IO Doc -> m ()
reportSDoc [Char]
"tc.sys.cover.sys" Int
20 forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
fsep forall a b. (a -> b) -> a -> b
$ forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Tele (Dom Type)
gamma forall a. a -> [a] -> [a]
: forall a b. (a -> b) -> [a] -> [b]
map forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM [([(Term, Bool)], Term)]
sys
      forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> TCMT IO Doc -> m ()
reportSDoc [Char]
"tc.sys.cover.sys" Int
40 forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
fsep forall a b. (a -> b) -> a -> b
$ (forall (m :: * -> *). Applicative m => [Char] -> m Doc
text forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall a. Show a => a -> [Char]
show) Tele (Dom Type)
gamma forall a. a -> [a] -> [a]
: forall a b. (a -> b) -> [a] -> [b]
map (forall (m :: * -> *). Applicative m => [Char] -> m Doc
text forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall a. Show a => a -> [Char]
show) [([(Term, Bool)], Term)]
sys
      forall (m :: * -> *) a. Monad m => a -> m a
return (Tele (Dom Type) -> [([(Term, Bool)], Term)] -> System
System Tele (Dom Type)
gamma [([(Term, Bool)], Term)]
sys) -- gamma uses names from the type, not the patterns, could we do better?
    Term
_ -> forall a. HasCallStack => a
__IMPOSSIBLE__
checkSystemCoverage QName
_ [Int]
_ Type
t [Clause]
cs = forall a. HasCallStack => a
__IMPOSSIBLE__


-- * Info that is needed after all clauses have been processed.

data ClausesPostChecks = CPC
    { ClausesPostChecks -> IntSet
cpcPartialSplits :: IntSet
      -- ^ Which argument indexes have a partial split.
    }

instance Semigroup ClausesPostChecks where
  CPC IntSet
xs <> :: ClausesPostChecks -> ClausesPostChecks -> ClausesPostChecks
<> CPC IntSet
xs' = IntSet -> ClausesPostChecks
CPC (IntSet -> IntSet -> IntSet
IntSet.union IntSet
xs IntSet
xs')

instance Monoid ClausesPostChecks where
  mempty :: ClausesPostChecks
mempty  = IntSet -> ClausesPostChecks
CPC forall a. Null a => a
empty
  mappend :: ClausesPostChecks -> ClausesPostChecks -> ClausesPostChecks
mappend = forall a. Semigroup a => a -> a -> a
(<>)

-- | The LHS part of checkClause.
checkClauseLHS :: Type -> Maybe Substitution -> A.SpineClause -> (LHSResult -> TCM a) -> TCM a
checkClauseLHS :: forall a.
Type
-> Maybe Substitution
-> SpineClause
-> (LHSResult -> TCM a)
-> TCM a
checkClauseLHS Type
t Maybe Substitution
withSub c :: SpineClause
c@(A.Clause lhs :: SpineLHS
lhs@(A.SpineLHS LHSInfo
i QName
x [NamedArg (Pattern' Expr)]
aps) [ProblemEq]
strippedPats RHS
rhs0 WhereDeclarations
wh Bool
catchall) LHSResult -> TCM a
ret = do
    forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> TCMT IO Doc -> m ()
reportSDoc [Char]
"tc.lhs.top" Int
30 forall a b. (a -> b) -> a -> b
$ TCMT IO Doc
"Checking clause" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
$$ forall a (m :: * -> *).
(ToConcrete a, Pretty (ConOfAbs a), MonadAbsToCon m) =>
a -> m Doc
prettyA SpineClause
c
    forall (m :: * -> *) a.
(Monad m, Null a) =>
a -> (a -> m ()) -> m ()
unlessNull ([NamedArg (Pattern' Expr)] -> [NamedArg (Pattern' Expr)]
trailingWithPatterns [NamedArg (Pattern' Expr)]
aps) forall a b. (a -> b) -> a -> b
$ \ [NamedArg (Pattern' Expr)]
withPats -> do
      forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
TypeError -> m a
typeError forall a b. (a -> b) -> a -> b
$ [Pattern' Expr] -> TypeError
UnexpectedWithPatterns forall a b. (a -> b) -> a -> b
$ forall a b. (a -> b) -> [a] -> [b]
map forall a. NamedArg a -> a
namedArg [NamedArg (Pattern' Expr)]
withPats
    forall (m :: * -> *) a. MonadTrace m => Call -> m a -> m a
traceCall (Type -> SpineClause -> Call
CheckClause Type
t SpineClause
c) forall a b. (a -> b) -> a -> b
$ do
      [NamedArg (Pattern' Expr)]
aps <- forall a. ExpandPatternSynonyms a => a -> TCM a
expandPatternSynonyms [NamedArg (Pattern' Expr)]
aps
      forall (f :: * -> *). Applicative f => Bool -> f () -> f ()
unless (forall a. Null a => a -> Bool
null [ProblemEq]
strippedPats) forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> TCMT IO Doc -> m ()
reportSDoc [Char]
"tc.lhs.top" Int
50 forall a b. (a -> b) -> a -> b
$
        TCMT IO Doc
"strippedPats:" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
vcat [ forall a (m :: * -> *).
(ToConcrete a, Pretty (ConOfAbs a), MonadAbsToCon m) =>
a -> m Doc
prettyA Pattern' Expr
p forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> TCMT IO Doc
"=" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Term
v forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> TCMT IO Doc
":" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Dom Type
a | A.ProblemEq Pattern' Expr
p Term
v Dom Type
a <- [ProblemEq]
strippedPats ]
      Type
closed_t <- forall a b c. (a -> b -> c) -> b -> a -> c
flip forall t. Abstract t => Tele (Dom Type) -> t -> t
abstract Type
t forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> forall (m :: * -> *).
(Applicative m, MonadTCEnv m) =>
m (Tele (Dom Type))
getContextTelescope
      forall a.
Call
-> Maybe QName
-> [NamedArg (Pattern' Expr)]
-> Type
-> Maybe Substitution
-> [ProblemEq]
-> (LHSResult -> TCM a)
-> TCM a
checkLeftHandSide (SpineLHS -> Call
CheckLHS SpineLHS
lhs) (forall a. a -> Maybe a
Just QName
x) [NamedArg (Pattern' Expr)]
aps Type
t Maybe Substitution
withSub [ProblemEq]
strippedPats LHSResult -> TCM a
ret

-- | Type check a function clause.

checkClause
  :: Type          -- ^ Type of function defined by this clause.
  -> Maybe Substitution  -- ^ Module parameter substitution arising from with-abstraction.
  -> A.SpineClause -- ^ Clause.
  -> TCM (Clause,ClausesPostChecks)  -- ^ Type-checked clause

checkClause :: Type
-> Maybe Substitution
-> SpineClause
-> TCMT IO (Clause, ClausesPostChecks)
checkClause Type
t Maybe Substitution
withSub c :: SpineClause
c@(A.Clause lhs :: SpineLHS
lhs@(A.SpineLHS LHSInfo
i QName
x [NamedArg (Pattern' Expr)]
aps) [ProblemEq]
strippedPats RHS
rhs0 WhereDeclarations
wh Bool
catchall) = do
  [Name]
cxtNames <- forall a. [a] -> [a]
reverse forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall a b. (a -> b) -> [a] -> [b]
map (forall a b. (a, b) -> a
fst forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall t e. Dom' t e -> e
unDom) forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> forall (m :: * -> *). MonadTCEnv m => m [Dom' Term (Name, Type)]
getContext
  forall a.
Type
-> Maybe Substitution
-> SpineClause
-> (LHSResult -> TCM a)
-> TCM a
checkClauseLHS Type
t Maybe Substitution
withSub SpineClause
c forall a b. (a -> b) -> a -> b
$ \ lhsResult :: LHSResult
lhsResult@(LHSResult Int
npars Tele (Dom Type)
delta [NamedArg DeBruijnPattern]
ps Bool
absurdPat Arg Type
trhs Substitution
patSubst [AsBinding]
asb IntSet
psplit Bool
ixsplit) -> do
        -- Note that we might now be in irrelevant context,
        -- in case checkLeftHandSide walked over an irrelevant projection pattern.

        -- Subtle: checkRHS expects the function type to be the lambda lifted
        -- type. If we're checking a with-function that's already the case,
        -- otherwise we need to abstract over the module telescope.
        Type
t' <- case Maybe Substitution
withSub of
                Just{}  -> forall (m :: * -> *) a. Monad m => a -> m a
return Type
t
                Maybe Substitution
Nothing -> do
                  Tele (Dom Type)
theta <- forall (m :: * -> *).
(Functor m, ReadTCState m) =>
ModuleName -> m (Tele (Dom Type))
lookupSection (QName -> ModuleName
qnameModule QName
x)
                  forall (m :: * -> *) a. Monad m => a -> m a
return forall a b. (a -> b) -> a -> b
$ forall t. Abstract t => Tele (Dom Type) -> t -> t
abstract Tele (Dom Type)
theta Type
t

        -- At this point we should update the named dots potential with-clauses
        -- in the right-hand side. When checking a clause we expect the named
        -- dots to live in the context of the closest parent lhs, but the named
        -- dots added by buildWithFunction live in the context of the
        -- with-function arguments before pattern matching. That's what we need
        -- patSubst for.
        let rhs :: RHS
rhs = RHS -> RHS
updateRHS RHS
rhs0
            updateRHS :: RHS -> RHS
updateRHS rhs :: RHS
rhs@A.RHS{}               = RHS
rhs
            updateRHS rhs :: RHS
rhs@A.AbsurdRHS{}         = RHS
rhs
            updateRHS (A.WithRHS QName
q [WithExpr]
es List1 (Clause' LHS)
cs)       = QName -> [WithExpr] -> List1 (Clause' LHS) -> RHS
A.WithRHS QName
q [WithExpr]
es forall a b. (a -> b) -> a -> b
$ forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
fmap Clause' LHS -> Clause' LHS
updateClause List1 (Clause' LHS)
cs
            updateRHS (A.RewriteRHS [RewriteEqn]
qes [ProblemEq]
spats RHS
rhs WhereDeclarations
wh) =
              [RewriteEqn] -> [ProblemEq] -> RHS -> WhereDeclarations -> RHS
A.RewriteRHS [RewriteEqn]
qes (forall a. Subst a => Substitution' (SubstArg a) -> a -> a
applySubst Substitution
patSubst [ProblemEq]
spats) (RHS -> RHS
updateRHS RHS
rhs) WhereDeclarations
wh

            updateClause :: Clause' LHS -> Clause' LHS
updateClause (A.Clause LHS
f [ProblemEq]
spats RHS
rhs WhereDeclarations
wh Bool
ca) =
              forall lhs.
lhs
-> [ProblemEq] -> RHS -> WhereDeclarations -> Bool -> Clause' lhs
A.Clause LHS
f (forall a. Subst a => Substitution' (SubstArg a) -> a -> a
applySubst Substitution
patSubst [ProblemEq]
spats) (RHS -> RHS
updateRHS RHS
rhs) WhereDeclarations
wh Bool
ca

        (Maybe Term
body, WithFunctionProblem
with) <- forall a. [AsBinding] -> TCM a -> TCM a
bindAsPatterns [AsBinding]
asb forall a b. (a -> b) -> a -> b
$ forall a. WhereDeclarations -> TCM a -> TCM a
checkWhere WhereDeclarations
wh forall a b. (a -> b) -> a -> b
$ LHSInfo
-> QName
-> [NamedArg (Pattern' Expr)]
-> Type
-> LHSResult
-> RHS
-> TCM (Maybe Term, WithFunctionProblem)
checkRHS LHSInfo
i QName
x [NamedArg (Pattern' Expr)]
aps Type
t' LHSResult
lhsResult RHS
rhs

        -- Note that the with function doesn't necessarily share any part of
        -- the context with the parent (but withSub will take you from parent
        -- to child).

        Maybe Term
wbody <- forall (tcm :: * -> *) a.
(MonadTCEnv tcm, ReadTCState tcm) =>
tcm a -> tcm a
unsafeInTopContext forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *) c.
MonadBench m =>
Account (BenchPhase m) -> m c -> m c
Bench.billTo [Phase
Bench.Typing, Phase
Bench.With] forall a b. (a -> b) -> a -> b
$ [Name] -> WithFunctionProblem -> TCMT IO (Maybe Term)
checkWithFunction [Name]
cxtNames WithFunctionProblem
with

        Maybe Term
body <- forall (m :: * -> *) a. Monad m => a -> m a
return forall a b. (a -> b) -> a -> b
$ Maybe Term
body forall (m :: * -> *) a. MonadPlus m => m a -> m a -> m a
`mplus` Maybe Term
wbody

        forall (m :: * -> *). Monad m => m Bool -> m () -> m ()
whenM (PragmaOptions -> Bool
optDoubleCheck forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> forall (m :: * -> *). HasOptions m => m PragmaOptions
pragmaOptions) forall a b. (a -> b) -> a -> b
$ case Maybe Term
body of
          Just Term
v  -> do
            forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> TCMT IO Doc -> m ()
reportSDoc [Char]
"tc.lhs.top" Int
30 forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
vcat
              [ TCMT IO Doc
"double checking rhs"
              , forall (m :: * -> *). Functor m => Int -> m Doc -> m Doc
nest Int
2 (forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Term
v forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> TCMT IO Doc
" : " forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM (forall e. Arg e -> e
unArg Arg Type
trhs))
              ]
            forall (m :: * -> *) a.
(MonadConstraint m, MonadWarning m, MonadError TCErr m,
 MonadFresh ProblemId m) =>
m a -> m a
noConstraints forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *) a.
(MonadMetaSolver m, MonadTCState m) =>
m a -> m a
withFrozenMetas forall a b. (a -> b) -> a -> b
$ forall a (m :: * -> *).
(CheckInternal a, MonadCheckInternal m) =>
a -> Comparison -> TypeOf a -> m ()
checkInternal Term
v Comparison
CmpLeq forall a b. (a -> b) -> a -> b
$ forall e. Arg e -> e
unArg Arg Type
trhs
          Maybe Term
Nothing -> forall (m :: * -> *) a. Monad m => a -> m a
return ()

        forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> TCMT IO Doc -> m ()
reportSDoc [Char]
"tc.lhs.top" Int
10 forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
vcat
          [ TCMT IO Doc
"Clause before translation:"
          , forall (m :: * -> *). Functor m => Int -> m Doc -> m Doc
nest Int
2 forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
vcat
            [ TCMT IO Doc
"delta =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> do forall (m :: * -> *) a.
MonadAddContext m =>
Impossible -> Int -> m a -> m a
escapeContext HasCallStack => Impossible
impossible (forall a. Sized a => a -> Int
size Tele (Dom Type)
delta) forall a b. (a -> b) -> a -> b
$ forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Tele (Dom Type)
delta
            , TCMT IO Doc
"ps    =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> do forall (t :: * -> *). Foldable t => t Doc -> Doc
P.fsep forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> forall (m :: * -> *).
MonadPretty m =>
[NamedArg DeBruijnPattern] -> m [Doc]
prettyTCMPatterns [NamedArg DeBruijnPattern]
ps
            , TCMT IO Doc
"body  =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall b a. b -> (a -> b) -> Maybe a -> b
maybe TCMT IO Doc
"_|_" forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Maybe Term
body
            , TCMT IO Doc
"type  =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Type
t
            ]
          ]

        forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> TCMT IO Doc -> m ()
reportSDoc [Char]
"tc.lhs.top" Int
60 forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *) a.
MonadAddContext m =>
Impossible -> Int -> m a -> m a
escapeContext HasCallStack => Impossible
impossible (forall a. Sized a => a -> Int
size Tele (Dom Type)
delta) forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
vcat
          [ TCMT IO Doc
"Clause before translation (raw):"
          , forall (m :: * -> *). Functor m => Int -> m Doc -> m Doc
nest Int
2 forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
vcat
            [ TCMT IO Doc
"ps    =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall (m :: * -> *). Applicative m => [Char] -> m Doc
text (forall a. Show a => a -> [Char]
show [NamedArg DeBruijnPattern]
ps)
            , TCMT IO Doc
"body  =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall (m :: * -> *). Applicative m => [Char] -> m Doc
text (forall a. Show a => a -> [Char]
show Maybe Term
body)
            , TCMT IO Doc
"type  =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall (m :: * -> *). Applicative m => [Char] -> m Doc
text (forall a. Show a => a -> [Char]
show Type
t)
            ]
          ]

        -- compute body modification for irrelevant definitions, see issue 610
        Relevance
rel <- forall (m :: * -> *) a. MonadTCEnv m => Lens' TCEnv a -> m a
viewTC Lens' TCEnv Relevance
eRelevance
        let bodyMod :: Maybe Term -> Maybe Term
bodyMod Maybe Term
body = case Relevance
rel of
              Relevance
Irrelevant -> Term -> Term
dontCare forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> Maybe Term
body
              Relevance
_          -> Maybe Term
body

        -- absurd clauses don't define computational behaviour, so it's fine to
        -- treat them as catchalls.
        let catchall' :: Bool
catchall' = Bool
catchall Bool -> Bool -> Bool
|| forall a. Maybe a -> Bool
isNothing Maybe Term
body

        -- absurd clauses are not exact
        let exact :: Maybe Bool
exact = if forall a. Maybe a -> Bool
isNothing Maybe Term
body then forall a. a -> Maybe a
Just Bool
False else forall a. Maybe a
Nothing -- we don't know yet

        forall (m :: * -> *) a. Monad m => a -> m a
return forall a b. (a -> b) -> a -> b
$ (, IntSet -> ClausesPostChecks
CPC IntSet
psplit)
          Clause { clauseLHSRange :: Range
clauseLHSRange  = forall a. HasRange a => a -> Range
getRange LHSInfo
i
                 , clauseFullRange :: Range
clauseFullRange = forall a. HasRange a => a -> Range
getRange SpineClause
c
                 , clauseTel :: Tele (Dom Type)
clauseTel       = forall a. KillRange a => KillRangeT a
killRange Tele (Dom Type)
delta
                 , namedClausePats :: [NamedArg DeBruijnPattern]
namedClausePats = [NamedArg DeBruijnPattern]
ps
                 , clauseBody :: Maybe Term
clauseBody      = Maybe Term -> Maybe Term
bodyMod Maybe Term
body
                 , clauseType :: Maybe (Arg Type)
clauseType      = forall a. a -> Maybe a
Just Arg Type
trhs
                 , clauseCatchall :: Bool
clauseCatchall  = Bool
catchall'
                 , clauseExact :: Maybe Bool
clauseExact       = Maybe Bool
exact
                 , clauseRecursive :: Maybe Bool
clauseRecursive   = forall a. Maybe a
Nothing -- we don't know yet
                 , clauseUnreachable :: Maybe Bool
clauseUnreachable = forall a. Maybe a
Nothing -- we don't know yet
                 , clauseEllipsis :: ExpandedEllipsis
clauseEllipsis    = LHSInfo -> ExpandedEllipsis
lhsEllipsis LHSInfo
i
                 , clauseWhereModule :: Maybe ModuleName
clauseWhereModule = WhereDeclarations -> Maybe ModuleName
A.whereModule WhereDeclarations
wh
                 }



-- | Generate the abstract pattern corresponding to Refl
getReflPattern :: TCM A.Pattern
getReflPattern :: TCM (Pattern' Expr)
getReflPattern = do
  -- Get the name of builtin REFL.
  Con ConHead
reflCon ConInfo
_ [] <- forall (m :: * -> *).
(HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) =>
m Term
primRefl

  Maybe ArgInfo
reflInfo <- forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
fmap (forall a. LensOrigin a => Origin -> a -> a
setOrigin Origin
Inserted) forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> ConHead -> TCMT IO (Maybe ArgInfo)
getReflArgInfo ConHead
reflCon
  let patInfo :: ConPatInfo
patInfo = ConInfo -> PatInfo -> ConPatLazy -> ConPatInfo
ConPatInfo ConInfo
ConOCon PatInfo
patNoRange ConPatLazy
ConPatEager
  -- The REFL constructor might have an argument
  let reflArg :: [NamedArg (Pattern' Expr)]
reflArg = forall a. Maybe a -> [a]
maybeToList forall a b. (a -> b) -> a -> b
$ forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
fmap (\ ArgInfo
ai -> forall e. ArgInfo -> e -> Arg e
Arg ArgInfo
ai forall a b. (a -> b) -> a -> b
$ forall a name. a -> Named name a
unnamed forall a b. (a -> b) -> a -> b
$ forall e. PatInfo -> Pattern' e
A.WildP PatInfo
patNoRange) Maybe ArgInfo
reflInfo

  forall (f :: * -> *) a. Applicative f => a -> f a
pure forall a b. (a -> b) -> a -> b
$ forall e. ConPatInfo -> AmbiguousQName -> NAPs e -> Pattern' e
A.ConP ConPatInfo
patInfo (QName -> AmbiguousQName
unambiguous forall a b. (a -> b) -> a -> b
$ ConHead -> QName
conName ConHead
reflCon) [NamedArg (Pattern' Expr)]
reflArg

-- | Type check the @with@ and @rewrite@ lhss and/or the rhs.
checkRHS
  :: LHSInfo                 -- ^ Range of lhs.
  -> QName                   -- ^ Name of function.
  -> [NamedArg A.Pattern]    -- ^ Patterns in lhs.
  -> Type                    -- ^ Top-level type of function.
  -> LHSResult               -- ^ Result of type-checking patterns
  -> A.RHS                   -- ^ Rhs to check.
  -> TCM (Maybe Term, WithFunctionProblem)
                                              -- Note: the as-bindings are already bound (in checkClause)
checkRHS :: LHSInfo
-> QName
-> [NamedArg (Pattern' Expr)]
-> Type
-> LHSResult
-> RHS
-> TCM (Maybe Term, WithFunctionProblem)
checkRHS LHSInfo
i QName
x [NamedArg (Pattern' Expr)]
aps Type
t lhsResult :: LHSResult
lhsResult@(LHSResult Int
_ Tele (Dom Type)
delta [NamedArg DeBruijnPattern]
ps Bool
absurdPat Arg Type
trhs Substitution
_ [AsBinding]
_asb IntSet
_ Bool
_) RHS
rhs0 =
  RHS -> TCM (Maybe Term, WithFunctionProblem)
handleRHS RHS
rhs0 where

  handleRHS :: A.RHS -> TCM (Maybe Term, WithFunctionProblem)
  handleRHS :: RHS -> TCM (Maybe Term, WithFunctionProblem)
handleRHS RHS
rhs = case RHS
rhs of
    A.RHS Expr
e Maybe Expr
_                  -> Expr -> TCM (Maybe Term, WithFunctionProblem)
ordinaryRHS Expr
e
    RHS
A.AbsurdRHS                -> TCM (Maybe Term, WithFunctionProblem)
noRHS
    A.RewriteRHS [RewriteEqn]
eqs [ProblemEq]
ps RHS
rhs WhereDeclarations
wh -> [RewriteEqn]
-> [ProblemEq]
-> RHS
-> WhereDeclarations
-> TCM (Maybe Term, WithFunctionProblem)
rewriteEqnsRHS [RewriteEqn]
eqs [ProblemEq]
ps RHS
rhs WhereDeclarations
wh
    A.WithRHS QName
aux [WithExpr]
es List1 (Clause' LHS)
cs        -> QName
-> [WithExpr]
-> List1 (Clause' LHS)
-> TCM (Maybe Term, WithFunctionProblem)
withRHS QName
aux [WithExpr]
es List1 (Clause' LHS)
cs

  -- Ordinary case: f xs = e
  ordinaryRHS :: A.Expr -> TCM (Maybe Term, WithFunctionProblem)
  ordinaryRHS :: Expr -> TCM (Maybe Term, WithFunctionProblem)
ordinaryRHS Expr
e = forall (m :: * -> *) c.
MonadBench m =>
Account (BenchPhase m) -> m c -> m c
Bench.billTo [Phase
Bench.Typing, Phase
Bench.CheckRHS] forall a b. (a -> b) -> a -> b
$ do
    -- If there is an absurd pattern, we do not need a RHS. If we have
    -- one we complain, ignore it and return the same @(Nothing, NoWithFunction)@
    -- as the case dealing with @A.AbsurdRHS@.
    Maybe Term
mv <- if Bool
absurdPat
          then do
            [NamedArg DeBruijnPattern]
ps <- forall a (m :: * -> *).
(InstantiateFull a, MonadReduce m) =>
a -> m a
instantiateFull [NamedArg DeBruijnPattern]
ps
            forall a. Maybe a
Nothing forall (f :: * -> *) a b. Functor f => a -> f b -> f a
<$ forall (m :: * -> *) x a.
(MonadTrace m, HasRange x) =>
x -> m a -> m a
setCurrentRange Expr
e (forall (m :: * -> *).
(HasCallStack, MonadWarning m) =>
Warning -> m ()
warning forall a b. (a -> b) -> a -> b
$ [NamedArg DeBruijnPattern] -> Warning
AbsurdPatternRequiresNoRHS [NamedArg DeBruijnPattern]
ps)
          else forall a. a -> Maybe a
Just forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> Expr -> Type -> TCMT IO Term
checkExpr Expr
e (forall e. Arg e -> e
unArg Arg Type
trhs)
    forall (m :: * -> *) a. Monad m => a -> m a
return (Maybe Term
mv, WithFunctionProblem
NoWithFunction)

  -- Absurd case: no right hand side
  noRHS :: TCM (Maybe Term, WithFunctionProblem)
  noRHS :: TCM (Maybe Term, WithFunctionProblem)
noRHS = do
    forall (f :: * -> *). Applicative f => Bool -> f () -> f ()
unless Bool
absurdPat forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
TypeError -> m a
typeError forall a b. (a -> b) -> a -> b
$ [NamedArg (Pattern' Expr)] -> TypeError
NoRHSRequiresAbsurdPattern [NamedArg (Pattern' Expr)]
aps
    forall (m :: * -> *) a. Monad m => a -> m a
return (forall a. Maybe a
Nothing, WithFunctionProblem
NoWithFunction)

  -- With case: @f xs with {a} in eqa | b in eqb | {{c}} | ...; ... | ps1 = rhs1; ... | ps2 = rhs2; ...@
  -- We need to modify the patterns `ps1, ps2, ...` in the user-provided clauses
  -- to insert the {eqb} names so that the equality proofs are available on the various RHS.
  withRHS ::
       QName             -- name of the with-function
    -> [A.WithExpr]      -- @[{a} in eqa, b in eqb, {{c}}, ...]@
    -> List1 A.Clause    -- @[(ps1 = rhs1), (ps2 = rhs), ...]@
    -> TCM (Maybe Term, WithFunctionProblem)
  withRHS :: QName
-> [WithExpr]
-> List1 (Clause' LHS)
-> TCM (Maybe Term, WithFunctionProblem)
withRHS QName
aux [WithExpr]
es List1 (Clause' LHS)
cs = do

    forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> TCMT IO Doc -> m ()
reportSDoc [Char]
"tc.with.top" Int
15 forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
vcat
      [ TCMT IO Doc
"TC.Rules.Def.checkclause reached A.WithRHS"
      , forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
sep forall a b. (a -> b) -> a -> b
$ forall a (m :: * -> *).
(ToConcrete a, Pretty (ConOfAbs a), MonadAbsToCon m) =>
a -> m Doc
prettyA QName
aux forall a. a -> [a] -> [a]
: forall a b. (a -> b) -> [a] -> [b]
map (forall (m :: * -> *). Functor m => m Doc -> m Doc
parens forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall a (m :: * -> *).
(ToConcrete a, Pretty (ConOfAbs a), MonadAbsToCon m) =>
a -> m Doc
prettyA forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall name a. Named name a -> a
namedThing) [WithExpr]
es
      ]
    forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> TCMT IO Doc -> m ()
reportSDoc [Char]
"tc.with.top" Int
20 forall a b. (a -> b) -> a -> b
$ do
      Int
nfv <- TCM Int
getCurrentModuleFreeVars
      ModuleName
m   <- forall (m :: * -> *). MonadTCEnv m => m ModuleName
currentModule
      forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
sep [ TCMT IO Doc
"with function module:" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+>
             forall (m :: * -> *) (t :: * -> *).
(Applicative m, Semigroup (m Doc), Foldable t) =>
t (m Doc) -> m Doc
prettyList (forall a b. (a -> b) -> [a] -> [b]
map forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM forall a b. (a -> b) -> a -> b
$ ModuleName -> [Name]
mnameToList ModuleName
m)
          ,  forall (m :: * -> *). Applicative m => [Char] -> m Doc
text forall a b. (a -> b) -> a -> b
$ [Char]
"free variables: " forall a. [a] -> [a] -> [a]
++ forall a. Show a => a -> [Char]
show Int
nfv
          ]

    -- Infer the types of the with expressions

    [Arg (Term, EqualityView)]
vtys <- forall (t :: * -> *) (m :: * -> *) a b.
(Traversable t, Monad m) =>
t a -> (a -> m b) -> m (t b)
forM [WithExpr]
es forall a b. (a -> b) -> a -> b
$ \ (Named Maybe BindName
nm Arg Expr
we) -> do
              (Term
e, Type
ty) <- Arg Expr -> TCM (Term, Type)
inferExprForWith Arg Expr
we
              forall (f :: * -> *) a. Applicative f => a -> f a
pure forall a b. (a -> b) -> a -> b
$ (forall (f :: * -> *) a b. Functor f => a -> f b -> f a
<$ Arg Expr
we) forall b c a. (b -> c) -> (a -> b) -> a -> c
. (Term
e,) forall a b. (a -> b) -> a -> b
$ case Maybe BindName
nm of
                Maybe BindName
Nothing -> Type -> EqualityView
OtherType Type
ty
                Just{}  -> Type -> EqualityView
IdiomType Type
ty

    let names :: [Arg (Maybe BindName)]
names = forall a b. (a -> b) -> [a] -> [b]
map (\ (Named Maybe BindName
nm Arg Expr
e) -> Maybe BindName
nm forall (f :: * -> *) a b. Functor f => a -> f b -> f a
<$ Arg Expr
e) [WithExpr]
es
    List1 (Clause' LHS)
cs <- forall (t :: * -> *) (m :: * -> *) a b.
(Traversable t, Monad m) =>
t a -> (a -> m b) -> m (t b)
forM List1 (Clause' LHS)
cs forall a b. (a -> b) -> a -> b
$ \ c :: Clause' LHS
c@(A.Clause (A.LHS LHSInfo
i LHSCore' Expr
core) [ProblemEq]
eqs RHS
rhs WhereDeclarations
wh Bool
b) -> do
      let rhs' :: RHS
rhs'  = [Arg (Maybe BindName)] -> RHS -> RHS
insertNames    [Arg (Maybe BindName)]
names RHS
rhs
      let core' :: LHSCore' Expr
core' = [Arg (Maybe BindName)] -> LHSCore' Expr -> LHSCore' Expr
insertInspects [Arg (Maybe BindName)]
names LHSCore' Expr
core
      forall (f :: * -> *) a. Applicative f => a -> f a
pure forall a b. (a -> b) -> a -> b
$ forall lhs.
lhs
-> [ProblemEq] -> RHS -> WhereDeclarations -> Bool -> Clause' lhs
A.Clause (LHSInfo -> LHSCore' Expr -> LHS
A.LHS LHSInfo
i LHSCore' Expr
core') [ProblemEq]
eqs RHS
rhs' WhereDeclarations
wh Bool
b

    -- Andreas, 2016-01-23, Issue #1796
    -- Run the size constraint solver to improve with-abstraction
    -- in case the with-expression contains size metas.
    DefaultToInfty -> TCM ()
solveSizeConstraints DefaultToInfty
DefaultToInfty

    QName
-> QName
-> Type
-> LHSResult
-> [Arg (Term, EqualityView)]
-> List1 (Clause' LHS)
-> TCM (Maybe Term, WithFunctionProblem)
checkWithRHS QName
x QName
aux Type
t LHSResult
lhsResult [Arg (Term, EqualityView)]
vtys List1 (Clause' LHS)
cs

  -- Rewrite case: f xs (rewrite / invert) a | b | c | ...
  rewriteEqnsRHS
    :: [A.RewriteEqn]
    -> [A.ProblemEq]
    -> A.RHS
    -> A.WhereDeclarations
    -> TCM (Maybe Term, WithFunctionProblem)

  rewriteEqnsRHS :: [RewriteEqn]
-> [ProblemEq]
-> RHS
-> WhereDeclarations
-> TCM (Maybe Term, WithFunctionProblem)
rewriteEqnsRHS [] [ProblemEq]
strippedPats RHS
rhs WhereDeclarations
wh = forall a. WhereDeclarations -> TCM a -> TCM a
checkWhere WhereDeclarations
wh forall a b. (a -> b) -> a -> b
$ RHS -> TCM (Maybe Term, WithFunctionProblem)
handleRHS RHS
rhs
      -- Case: @rewrite@
      -- Andreas, 2014-01-17, Issue 1402:
      -- If the rewrites are discarded since lhs=rhs, then
      -- we can actually have where clauses.
  rewriteEqnsRHS (RewriteEqn
r:[RewriteEqn]
rs) [ProblemEq]
strippedPats RHS
rhs WhereDeclarations
wh = case RewriteEqn
r of
    Rewrite ((QName
qname, Expr
eq) :| [(QName, Expr)]
qes) ->
      QName
-> Expr -> [RewriteEqn] -> TCM (Maybe Term, WithFunctionProblem)
rewriteEqnRHS QName
qname Expr
eq forall a b. (a -> b) -> a -> b
$
        forall a b. [a] -> b -> (List1 a -> b) -> b
List1.ifNull [(QName, Expr)]
qes {-then-} [RewriteEqn]
rs {-else-} forall a b. (a -> b) -> a -> b
$ \ NonEmpty (QName, Expr)
qes -> forall qn nm p e. List1 (qn, e) -> RewriteEqn' qn nm p e
Rewrite NonEmpty (QName, Expr)
qes forall a. a -> [a] -> [a]
: [RewriteEqn]
rs
    Invert QName
qname List1 (Named BindName (Pattern' Expr, Expr))
pes -> QName
-> [Named BindName (Pattern' Expr, Expr)]
-> [RewriteEqn]
-> TCM (Maybe Term, WithFunctionProblem)
invertEqnRHS QName
qname (forall l. IsList l => l -> [Item l]
List1.toList List1 (Named BindName (Pattern' Expr, Expr))
pes) [RewriteEqn]
rs

    where

    -- @invert@ clauses
    invertEqnRHS :: QName -> [Named A.BindName (A.Pattern,A.Expr)] -> [A.RewriteEqn] -> TCM (Maybe Term, WithFunctionProblem)
    invertEqnRHS :: QName
-> [Named BindName (Pattern' Expr, Expr)]
-> [RewriteEqn]
-> TCM (Maybe Term, WithFunctionProblem)
invertEqnRHS QName
qname [Named BindName (Pattern' Expr, Expr)]
pes [RewriteEqn]
rs = do

      let ([Named BindName (Pattern' Expr)]
npats, [Named BindName Expr]
es) = forall a b c. (a -> (b, c)) -> [a] -> ([b], [c])
unzipWith (\ (Named Maybe BindName
nm (Pattern' Expr
p , Expr
e)) -> (forall name a. Maybe name -> a -> Named name a
Named Maybe BindName
nm Pattern' Expr
p, forall name a. Maybe name -> a -> Named name a
Named Maybe BindName
nm Expr
e)) [Named BindName (Pattern' Expr, Expr)]
pes
      -- Infer the types of the with expressions
      [Arg (Term, EqualityView)]
vtys <- forall (t :: * -> *) (m :: * -> *) a b.
(Traversable t, Monad m) =>
t a -> (a -> m b) -> m (t b)
forM [Named BindName Expr]
es forall a b. (a -> b) -> a -> b
$ \ (Named Maybe BindName
nm Expr
we) -> do
        (Term
e, Type
ty) <- Arg Expr -> TCM (Term, Type)
inferExprForWith (forall a. a -> Arg a
defaultArg Expr
we)
        forall (f :: * -> *) a. Applicative f => a -> f a
pure forall a b. (a -> b) -> a -> b
$ forall a. a -> Arg a
defaultArg forall b c a. (b -> c) -> (a -> b) -> a -> c
. (Term
e,) forall a b. (a -> b) -> a -> b
$ case Maybe BindName
nm of
          Maybe BindName
Nothing -> Type -> EqualityView
OtherType Type
ty
          Just{}  -> Type -> EqualityView
IdiomType Type
ty

      let pats :: [Arg (Pattern' Expr)]
pats = forall (t :: * -> *) a b. Foldable t => (a -> [b]) -> t a -> [b]
concatMap (forall a b. (a -> b) -> [a] -> [b]
map forall a. a -> Arg a
defaultArg) forall a b. (a -> b) -> a -> b
$
            forall (m :: * -> *) a b. Functor m => m a -> (a -> b) -> m b
for [Named BindName (Pattern' Expr)]
npats forall a b. (a -> b) -> a -> b
$ \ (Named Maybe BindName
nm Pattern' Expr
p) -> case Maybe BindName
nm of
              Maybe BindName
Nothing -> [Pattern' Expr
p]
              Just BindName
n  -> [Pattern' Expr
p, forall e. BindName -> Pattern' e
A.VarP BindName
n]

      -- Andreas, 2016-04-14, see also Issue #1796
      -- Run the size constraint solver to improve with-abstraction
      -- in case the with-expression contains size metas.
      DefaultToInfty -> TCM ()
solveSizeConstraints DefaultToInfty
DefaultToInfty

      let rhs' :: RHS
rhs' = [Arg (Pattern' Expr)] -> RHS -> RHS
insertPatterns [Arg (Pattern' Expr)]
pats RHS
rhs
          (RHS
rhs'', WhereDeclarations
outerWhere) -- the where clauses should go on the inner-most with
            | forall a. Null a => a -> Bool
null [RewriteEqn]
rs  = (RHS
rhs', WhereDeclarations
wh)
            | Bool
otherwise = ([RewriteEqn] -> [ProblemEq] -> RHS -> WhereDeclarations -> RHS
A.RewriteRHS [RewriteEqn]
rs [ProblemEq]
strippedPats RHS
rhs' WhereDeclarations
wh, WhereDeclarations
A.noWhereDecls)
          -- Andreas, 2014-03-05 kill range of copied patterns
          -- since they really do not have a source location.
          cl :: Clause' LHS
cl = forall lhs.
lhs
-> [ProblemEq] -> RHS -> WhereDeclarations -> Bool -> Clause' lhs
A.Clause (LHSInfo -> LHSCore' Expr -> LHS
A.LHS LHSInfo
i forall a b. (a -> b) -> a -> b
$ [Arg (Pattern' Expr)] -> LHSCore' Expr -> LHSCore' Expr
insertPatternsLHSCore [Arg (Pattern' Expr)]
pats forall a b. (a -> b) -> a -> b
$ forall e. QName -> [NamedArg (Pattern' e)] -> LHSCore' e
A.LHSHead QName
x forall a b. (a -> b) -> a -> b
$ forall a. KillRange a => KillRangeT a
killRange [NamedArg (Pattern' Expr)]
aps)
                 [ProblemEq]
strippedPats RHS
rhs'' WhereDeclarations
outerWhere Bool
False

      forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> TCMT IO Doc -> m ()
reportSDoc [Char]
"tc.invert" Int
60 forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
vcat
        [ forall (m :: * -> *). Applicative m => [Char] -> m Doc
text [Char]
"invert"
        , TCMT IO Doc
"  rhs' = " forall a. Semigroup a => a -> a -> a
<> (forall (m :: * -> *). Applicative m => [Char] -> m Doc
text forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall a. Show a => a -> [Char]
show) RHS
rhs'
        ]
      QName
-> QName
-> Type
-> LHSResult
-> [Arg (Term, EqualityView)]
-> List1 (Clause' LHS)
-> TCM (Maybe Term, WithFunctionProblem)
checkWithRHS QName
x QName
qname Type
t LHSResult
lhsResult [Arg (Term, EqualityView)]
vtys forall a b. (a -> b) -> a -> b
$ forall el coll. Singleton el coll => el -> coll
singleton Clause' LHS
cl

    -- @rewrite@ clauses
    rewriteEqnRHS
      :: QName
      -> A.Expr
      -> [A.RewriteEqn]
      -> TCM (Maybe Term, WithFunctionProblem)
    rewriteEqnRHS :: QName
-> Expr -> [RewriteEqn] -> TCM (Maybe Term, WithFunctionProblem)
rewriteEqnRHS QName
qname Expr
eq [RewriteEqn]
rs = do

      -- Action for skipping this rewrite.
      -- We do not want to create unsolved metas in case of
      -- a futile rewrite with a reflexive equation.
      -- Thus, we restore the state in this case,
      -- unless the rewrite expression contains questionmarks.
      TCState
st <- forall (m :: * -> *). MonadTCState m => m TCState
getTC
      -- TODO:: recurse defined but not used
      let recurse :: TCM (Maybe Term, WithFunctionProblem)
recurse = do
           TCState
st' <- forall (m :: * -> *). MonadTCState m => m TCState
getTC
           -- Comparing the whole stInteractionPoints maps is a bit
           -- wasteful, but we assume
           -- 1. rewriting with a reflexive equality to happen rarely,
           -- 2. especially with ?-holes in the rewrite expression
           -- 3. and a large overall number of ?s.
           let sameIP :: TCState -> TCState -> Bool
sameIP = forall a. Eq a => a -> a -> Bool
(==) forall b c a. (b -> b -> c) -> (a -> b) -> a -> a -> c
`on` (forall o i. o -> Lens' o i -> i
^.Lens' TCState InteractionPoints
stInteractionPoints)
           forall (f :: * -> *). Applicative f => Bool -> f () -> f ()
when (TCState -> TCState -> Bool
sameIP TCState
st TCState
st') forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *). MonadTCState m => TCState -> m ()
putTC TCState
st
           RHS -> TCM (Maybe Term, WithFunctionProblem)
handleRHS forall a b. (a -> b) -> a -> b
$ [RewriteEqn] -> [ProblemEq] -> RHS -> WhereDeclarations -> RHS
A.RewriteRHS [RewriteEqn]
rs [ProblemEq]
strippedPats RHS
rhs WhereDeclarations
wh

      -- Get value and type of rewrite-expression.

      (Term
proof, Type
eqt) <- Expr -> TCM (Term, Type)
inferExpr Expr
eq

      -- Andreas, 2016-04-14, see also Issue #1796
      -- Run the size constraint solver to improve with-abstraction
      -- in case the with-expression contains size metas.
      DefaultToInfty -> TCM ()
solveSizeConstraints DefaultToInfty
DefaultToInfty

      -- Check that the type is actually an equality (lhs ≡ rhs)
      -- and extract lhs, rhs, and their type.

      Type
t' <- forall a (m :: * -> *). (Reduce a, MonadReduce m) => a -> m a
reduce forall (m :: * -> *) a b. Monad m => (a -> m b) -> m a -> m b
=<< forall a (m :: * -> *).
(InstantiateFull a, MonadReduce m) =>
a -> m a
instantiateFull Type
eqt
      (EqualityView
eqt,Type
rewriteType,Term
rewriteFrom,Term
rewriteTo) <- Type -> TCM EqualityView
equalityView Type
t' forall (m :: * -> *) a b. Monad m => m a -> (a -> m b) -> m b
>>= \case
        eqt :: EqualityView
eqt@(EqualityType Sort
_s QName
_eq [Arg Term]
_params (Arg ArgInfo
_ Term
dom) Arg Term
a Arg Term
b) -> do
          Sort
s <- forall (m :: * -> *).
(PureTCM m, MonadBlock m, MonadConstraint m) =>
Term -> m Sort
sortOf Term
dom
          forall (m :: * -> *) a. Monad m => a -> m a
return (EqualityView
eqt, forall t a. Sort' t -> a -> Type'' t a
El Sort
s Term
dom, forall e. Arg e -> e
unArg Arg Term
a, forall e. Arg e -> e
unArg Arg Term
b)
          -- Note: the sort _s of the equality need not be the sort of the type @dom@!
        OtherType{} -> forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
TypeError -> m a
typeError forall a b. (a -> b) -> a -> b
$ Type -> TypeError
CannotRewriteByNonEquation Type
t'
        IdiomType{} -> forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
TypeError -> m a
typeError forall a b. (a -> b) -> a -> b
$ Type -> TypeError
CannotRewriteByNonEquation Type
t'

      Pattern' Expr
reflPat <- TCM (Pattern' Expr)
getReflPattern

      -- Andreas, 2015-12-25  Issue #1740:
      -- After the fix of #520, rewriting with a reflexive equation
      -- has to be desugared as matching against refl.
      let isReflexive :: TCMT IO Bool
isReflexive = forall (m :: * -> *).
(MonadConstraint m, MonadWarning m, MonadError TCErr m,
 MonadFresh ProblemId m) =>
m () -> m Bool
tryConversion forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *) a.
(MonadTCEnv m, HasOptions m, MonadDebug m) =>
m a -> m a
dontAssignMetas forall a b. (a -> b) -> a -> b
$
           forall (m :: * -> *).
MonadConversion m =>
Type -> Term -> Term -> m ()
equalTerm Type
rewriteType Term
rewriteFrom Term
rewriteTo

      ([Pattern' Expr]
pats', Term
withExpr, EqualityView
withType) <- do
        forall (m :: * -> *) a. Monad m => m Bool -> m a -> m a -> m a
ifM TCMT IO Bool
isReflexive
          {-then-} (forall (m :: * -> *) a. Monad m => a -> m a
return ([ Pattern' Expr
reflPat ]                    , Term
proof, Type -> EqualityView
OtherType Type
t'))
          {-else-} (forall (m :: * -> *) a. Monad m => a -> m a
return ([ forall e. PatInfo -> Pattern' e
A.WildP PatInfo
patNoRange, Pattern' Expr
reflPat ], Term
proof, EqualityView
eqt))
      let pats :: [Arg (Pattern' Expr)]
pats = forall a. a -> Arg a
defaultArg forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> [Pattern' Expr]
pats'

      let rhs' :: RHS
rhs' = [Arg (Pattern' Expr)] -> RHS -> RHS
insertPatterns [Arg (Pattern' Expr)]
pats RHS
rhs
          (RHS
rhs'', WhereDeclarations
outerWhere) -- the where clauses should go on the inner-most with
            | forall a. Null a => a -> Bool
null [RewriteEqn]
rs  = (RHS
rhs', WhereDeclarations
wh)
            | Bool
otherwise = ([RewriteEqn] -> [ProblemEq] -> RHS -> WhereDeclarations -> RHS
A.RewriteRHS [RewriteEqn]
rs [ProblemEq]
strippedPats RHS
rhs' WhereDeclarations
wh, WhereDeclarations
A.noWhereDecls)
          -- Andreas, 2014-03-05 kill range of copied patterns
          -- since they really do not have a source location.
          cl :: Clause' LHS
cl = forall lhs.
lhs
-> [ProblemEq] -> RHS -> WhereDeclarations -> Bool -> Clause' lhs
A.Clause (LHSInfo -> LHSCore' Expr -> LHS
A.LHS LHSInfo
i forall a b. (a -> b) -> a -> b
$ [Arg (Pattern' Expr)] -> LHSCore' Expr -> LHSCore' Expr
insertPatternsLHSCore [Arg (Pattern' Expr)]
pats forall a b. (a -> b) -> a -> b
$ forall e. QName -> [NamedArg (Pattern' e)] -> LHSCore' e
A.LHSHead QName
x forall a b. (a -> b) -> a -> b
$ forall a. KillRange a => KillRangeT a
killRange [NamedArg (Pattern' Expr)]
aps)
                 [ProblemEq]
strippedPats RHS
rhs'' WhereDeclarations
outerWhere Bool
False

      forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> TCMT IO Doc -> m ()
reportSDoc [Char]
"tc.rewrite" Int
60 forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
vcat
        [ forall (m :: * -> *). Applicative m => [Char] -> m Doc
text [Char]
"rewrite"
        , TCMT IO Doc
"  rhs' = " forall a. Semigroup a => a -> a -> a
<> (forall (m :: * -> *). Applicative m => [Char] -> m Doc
text forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall a. Show a => a -> [Char]
show) RHS
rhs'
        ]
      QName
-> QName
-> Type
-> LHSResult
-> [Arg (Term, EqualityView)]
-> List1 (Clause' LHS)
-> TCM (Maybe Term, WithFunctionProblem)
checkWithRHS QName
x QName
qname Type
t LHSResult
lhsResult [forall a. a -> Arg a
defaultArg (Term
withExpr, EqualityView
withType)] forall a b. (a -> b) -> a -> b
$ forall el coll. Singleton el coll => el -> coll
singleton Clause' LHS
cl

checkWithRHS
  :: QName                             -- ^ Name of function.
  -> QName                             -- ^ Name of the with-function.
  -> Type                              -- ^ Type of function.
  -> LHSResult                         -- ^ Result of type-checking patterns
  -> [Arg (Term, EqualityView)]        -- ^ Expressions and types of with-expressions.
  -> List1 A.Clause                    -- ^ With-clauses to check.
  -> TCM (Maybe Term, WithFunctionProblem)
                                -- Note: as-bindings already bound (in checkClause)
checkWithRHS :: QName
-> QName
-> Type
-> LHSResult
-> [Arg (Term, EqualityView)]
-> List1 (Clause' LHS)
-> TCM (Maybe Term, WithFunctionProblem)
checkWithRHS QName
x QName
aux Type
t (LHSResult Int
npars Tele (Dom Type)
delta [NamedArg DeBruijnPattern]
ps Bool
_absurdPat Arg Type
trhs Substitution
_ [AsBinding]
_asb IntSet
_ Bool
_) [Arg (Term, EqualityView)]
vtys0 List1 (Clause' LHS)
cs =
  forall (m :: * -> *) a.
MonadDebug m =>
[Char] -> Int -> [Char] -> m a -> m a
verboseBracket [Char]
"tc.with.top" Int
25 [Char]
"checkWithRHS" forall a b. (a -> b) -> a -> b
$ do
    forall (m :: * -> *) c.
MonadBench m =>
Account (BenchPhase m) -> m c -> m c
Bench.billTo [Phase
Bench.Typing, Phase
Bench.With] forall a b. (a -> b) -> a -> b
$ do
        [Arg Term]
withArgs <- [Arg (Term, EqualityView)] -> TCM [Arg Term]
withArguments [Arg (Term, EqualityView)]
vtys0
        let perm :: Permutation
perm = forall a. a -> Maybe a -> a
fromMaybe forall a. HasCallStack => a
__IMPOSSIBLE__ forall a b. (a -> b) -> a -> b
$ [NamedArg DeBruijnPattern] -> Maybe Permutation
dbPatPerm [NamedArg DeBruijnPattern]
ps

        forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> TCMT IO Doc -> m ()
reportSDoc [Char]
"tc.with.top" Int
30 forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
vcat forall a b. (a -> b) -> a -> b
$
          -- declared locally because we do not want to use the unzip'd thing!
          let ([Term]
vs, [EqualityView]
as) = forall a b c. (a -> (b, c)) -> [a] -> ([b], [c])
unzipWith forall e. Arg e -> e
unArg [Arg (Term, EqualityView)]
vtys0 in
          [ TCMT IO Doc
"vs (before normalization) =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM [Term]
vs
          , TCMT IO Doc
"as (before normalization) =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM [EqualityView]
as
          ]
        forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> TCMT IO Doc -> m ()
reportSDoc [Char]
"tc.with.top" Int
45 forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
vcat forall a b. (a -> b) -> a -> b
$
          -- declared locally because we do not want to use the unzip'd thing!
          let ([Term]
vs, [EqualityView]
as) = forall a b c. (a -> (b, c)) -> [a] -> ([b], [c])
unzipWith forall e. Arg e -> e
unArg [Arg (Term, EqualityView)]
vtys0 in
          [ TCMT IO Doc
"vs (before norm., raw) =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall (m :: * -> *) a. (Applicative m, Pretty a) => a -> m Doc
pretty [Term]
vs
          ]
        [Arg (Term, EqualityView)]
vtys0 <- forall a (m :: * -> *). (Normalise a, MonadReduce m) => a -> m a
normalise [Arg (Term, EqualityView)]
vtys0

        -- Andreas, 2012-09-17: for printing delta,
        -- we should remove it from the context first
        forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> TCMT IO Doc -> m ()
reportSDoc [Char]
"tc.with.top" Int
25 forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *) a.
MonadAddContext m =>
Impossible -> Int -> m a -> m a
escapeContext HasCallStack => Impossible
impossible (forall a. Sized a => a -> Int
size Tele (Dom Type)
delta) forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
vcat
          [ TCMT IO Doc
"delta  =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Tele (Dom Type)
delta
          ]
        forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> TCMT IO Doc -> m ()
reportSDoc [Char]
"tc.with.top" Int
25 forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
vcat forall a b. (a -> b) -> a -> b
$
          -- declared locally because we do not want to use the unzip'd thing!
          let ([Term]
vs, [EqualityView]
as) = forall a b c. (a -> (b, c)) -> [a] -> ([b], [c])
unzipWith forall e. Arg e -> e
unArg [Arg (Term, EqualityView)]
vtys0 in
          [ TCMT IO Doc
"vs     =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM [Term]
vs
          , TCMT IO Doc
"as     =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM [EqualityView]
as
          , TCMT IO Doc
"perm   =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall (m :: * -> *). Applicative m => [Char] -> m Doc
text (forall a. Show a => a -> [Char]
show Permutation
perm)
          ]

        -- Split the telescope into the part needed to type the with arguments
        -- and all the other stuff
        let (Tele (Dom Type)
delta1, Tele (Dom Type)
delta2, Permutation
perm', Type
t', [Arg (Term, EqualityView)]
vtys) = Tele (Dom Type)
-> Type
-> [Arg (Term, EqualityView)]
-> (Tele (Dom Type), Tele (Dom Type), Permutation, Type,
    [Arg (Term, EqualityView)])
splitTelForWith Tele (Dom Type)
delta (forall e. Arg e -> e
unArg Arg Type
trhs) [Arg (Term, EqualityView)]
vtys0
        let finalPerm :: Permutation
finalPerm = Permutation -> Permutation -> Permutation
composeP Permutation
perm' Permutation
perm

        forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> [Char] -> m ()
reportSLn [Char]
"tc.with.top" Int
75 forall a b. (a -> b) -> a -> b
$ [Char]
"delta  = " forall a. [a] -> [a] -> [a]
++ forall a. Show a => a -> [Char]
show Tele (Dom Type)
delta

        -- Andreas, 2012-09-17: for printing delta,
        -- we should remove it from the context first
        forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> TCMT IO Doc -> m ()
reportSDoc [Char]
"tc.with.top" Int
25 forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *) a.
MonadAddContext m =>
Impossible -> Int -> m a -> m a
escapeContext HasCallStack => Impossible
impossible (forall a. Sized a => a -> Int
size Tele (Dom Type)
delta) forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
vcat
          [ TCMT IO Doc
"delta1 =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Tele (Dom Type)
delta1
          , TCMT IO Doc
"delta2 =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall b (m :: * -> *) a.
(AddContext b, MonadAddContext m) =>
b -> m a -> m a
addContext Tele (Dom Type)
delta1 (forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Tele (Dom Type)
delta2)
          ]
        forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> TCMT IO Doc -> m ()
reportSDoc [Char]
"tc.with.top" Int
25 forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
vcat
          [ TCMT IO Doc
"perm'  =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall (m :: * -> *). Applicative m => [Char] -> m Doc
text (forall a. Show a => a -> [Char]
show Permutation
perm')
          , TCMT IO Doc
"fPerm  =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall (m :: * -> *). Applicative m => [Char] -> m Doc
text (forall a. Show a => a -> [Char]
show Permutation
finalPerm)
          ]

        -- Create the body of the original function

        -- All the context variables
        [Term]
us <- forall (m :: * -> *). (Applicative m, MonadTCEnv m) => m [Term]
getContextTerms
        let n :: Int
n = forall a. Sized a => a -> Int
size [Term]
us
            m :: Int
m = forall a. Sized a => a -> Int
size Tele (Dom Type)
delta
            -- First the variables bound outside this definition
            ([Term]
us0, [Term]
us1') = forall a. Int -> [a] -> ([a], [a])
splitAt (Int
n forall a. Num a => a -> a -> a
- Int
m) [Term]
us
            -- Then permute the rest and grab those needed to for the with arguments
            ([Term]
us1, [Term]
us2)  = forall a. Int -> [a] -> ([a], [a])
splitAt (forall a. Sized a => a -> Int
size Tele (Dom Type)
delta1) forall a b. (a -> b) -> a -> b
$ forall a. Permutation -> [a] -> [a]
permute Permutation
perm' [Term]
us1'
            -- Now stuff the with arguments in between and finish with the remaining variables
            argsS :: Substitution
argsS = forall a. DeBruijn a => [a] -> Substitution' a
parallelS forall a b. (a -> b) -> a -> b
$ forall a. [a] -> [a]
reverse forall a b. (a -> b) -> a -> b
$ [Term]
us0 forall a. [a] -> [a] -> [a]
++ [Term]
us1 forall a. [a] -> [a] -> [a]
++ forall a b. (a -> b) -> [a] -> [b]
map forall e. Arg e -> e
unArg [Arg Term]
withArgs forall a. [a] -> [a] -> [a]
++ [Term]
us2
            v :: Maybe a
v         = forall a. Maybe a
Nothing -- generated by checkWithFunction
        -- Andreas, 2013-02-26 add with-name to signature for printing purposes
        QName -> Definition -> TCM ()
addConstant QName
aux forall (m :: * -> *) a b. Monad m => (a -> m b) -> m a -> m b
=<< do
          Language
lang <- forall (m :: * -> *). HasOptions m => m Language
getLanguage
          Definition -> TCMT IO Definition
useTerPragma forall (m :: * -> *) a b. Monad m => (a -> m b) -> m a -> m b
=<<
            ArgInfo -> QName -> Type -> Language -> Defn -> Definition
defaultDefn ArgInfo
defaultArgInfo QName
aux HasCallStack => Type
__DUMMY_TYPE__ Language
lang forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$>
              forall (m :: * -> *). HasOptions m => m Defn
emptyFunction

        forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> TCMT IO Doc -> m ()
reportSDoc [Char]
"tc.with.top" Int
20 forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
vcat forall a b. (a -> b) -> a -> b
$
          let ([Term]
vs, [EqualityView]
as) = forall a b c. (a -> (b, c)) -> [a] -> ([b], [c])
unzipWith forall e. Arg e -> e
unArg [Arg (Term, EqualityView)]
vtys in
          [ TCMT IO Doc
"    with arguments" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> do forall (m :: * -> *) a.
MonadAddContext m =>
Impossible -> Int -> m a -> m a
escapeContext HasCallStack => Impossible
impossible (forall a. Sized a => a -> Int
size Tele (Dom Type)
delta) forall a b. (a -> b) -> a -> b
$ forall b (m :: * -> *) a.
(AddContext b, MonadAddContext m) =>
b -> m a -> m a
addContext Tele (Dom Type)
delta1 forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *) (t :: * -> *).
(Applicative m, Semigroup (m Doc), Foldable t) =>
t (m Doc) -> m Doc
prettyList (forall a b. (a -> b) -> [a] -> [b]
map forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM [Term]
vs)
          , TCMT IO Doc
"             types" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> do forall (m :: * -> *) a.
MonadAddContext m =>
Impossible -> Int -> m a -> m a
escapeContext HasCallStack => Impossible
impossible (forall a. Sized a => a -> Int
size Tele (Dom Type)
delta) forall a b. (a -> b) -> a -> b
$ forall b (m :: * -> *) a.
(AddContext b, MonadAddContext m) =>
b -> m a -> m a
addContext Tele (Dom Type)
delta1 forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *) (t :: * -> *).
(Applicative m, Semigroup (m Doc), Foldable t) =>
t (m Doc) -> m Doc
prettyList (forall a b. (a -> b) -> [a] -> [b]
map forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM [EqualityView]
as)
          , TCMT IO Doc
"           context" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> (forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM forall (m :: * -> *) a b. Monad m => (a -> m b) -> m a -> m b
=<< forall (m :: * -> *).
(Applicative m, MonadTCEnv m) =>
m (Tele (Dom Type))
getContextTelescope)
          , TCMT IO Doc
"             delta" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> do forall (m :: * -> *) a.
MonadAddContext m =>
Impossible -> Int -> m a -> m a
escapeContext HasCallStack => Impossible
impossible (forall a. Sized a => a -> Int
size Tele (Dom Type)
delta) forall a b. (a -> b) -> a -> b
$ forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Tele (Dom Type)
delta
          , TCMT IO Doc
"            delta1" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> do forall (m :: * -> *) a.
MonadAddContext m =>
Impossible -> Int -> m a -> m a
escapeContext HasCallStack => Impossible
impossible (forall a. Sized a => a -> Int
size Tele (Dom Type)
delta) forall a b. (a -> b) -> a -> b
$ forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Tele (Dom Type)
delta1
          , TCMT IO Doc
"            delta2" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> do forall (m :: * -> *) a.
MonadAddContext m =>
Impossible -> Int -> m a -> m a
escapeContext HasCallStack => Impossible
impossible (forall a. Sized a => a -> Int
size Tele (Dom Type)
delta) forall a b. (a -> b) -> a -> b
$ forall b (m :: * -> *) a.
(AddContext b, MonadAddContext m) =>
b -> m a -> m a
addContext Tele (Dom Type)
delta1 forall a b. (a -> b) -> a -> b
$ forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Tele (Dom Type)
delta2
          ]

        forall (m :: * -> *) a. Monad m => a -> m a
return (forall a. Maybe a
v, QName
-> QName
-> Type
-> Tele (Dom Type)
-> Tele (Dom Type)
-> Tele (Dom Type)
-> [Arg (Term, EqualityView)]
-> Type
-> [NamedArg DeBruijnPattern]
-> Int
-> Permutation
-> Permutation
-> Permutation
-> List1 (Clause' LHS)
-> Substitution
-> WithFunctionProblem
WithFunction QName
x QName
aux Type
t Tele (Dom Type)
delta Tele (Dom Type)
delta1 Tele (Dom Type)
delta2 [Arg (Term, EqualityView)]
vtys Type
t' [NamedArg DeBruijnPattern]
ps Int
npars Permutation
perm' Permutation
perm Permutation
finalPerm List1 (Clause' LHS)
cs Substitution
argsS)

-- | Invoked in empty context.
checkWithFunction :: [Name] -> WithFunctionProblem -> TCM (Maybe Term)
checkWithFunction :: [Name] -> WithFunctionProblem -> TCMT IO (Maybe Term)
checkWithFunction [Name]
_ WithFunctionProblem
NoWithFunction = forall (m :: * -> *) a. Monad m => a -> m a
return forall a. Maybe a
Nothing
checkWithFunction [Name]
cxtNames (WithFunction QName
f QName
aux Type
t Tele (Dom Type)
delta Tele (Dom Type)
delta1 Tele (Dom Type)
delta2 [Arg (Term, EqualityView)]
vtys Type
b [NamedArg DeBruijnPattern]
qs Int
npars Permutation
perm' Permutation
perm Permutation
finalPerm List1 (Clause' LHS)
cs Substitution
argsS) = do

  let -- Δ₁ ws Δ₂ ⊢ withSub : Δ′    (where Δ′ is the context of the parent lhs)
      withSub :: Substitution
      withSub :: Substitution
withSub = let as :: [EqualityView]
as = forall a b. (a -> b) -> [a] -> [b]
map (forall a b. (a, b) -> b
snd forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall e. Arg e -> e
unArg) [Arg (Term, EqualityView)]
vtys in
                forall a. Int -> Substitution' a -> Substitution' a
liftS (forall a. Sized a => a -> Int
size Tele (Dom Type)
delta2) (forall a. Int -> Substitution' a -> Substitution' a
wkS ([EqualityView] -> Int
countWithArgs [EqualityView]
as) forall a. Substitution' a
idS)
                forall a.
EndoSubst a =>
Substitution' a -> Substitution' a -> Substitution' a
`composeS` forall a.
DeBruijn a =>
Impossible -> Permutation -> Substitution' a
renaming HasCallStack => Impossible
impossible (Permutation -> Permutation
reverseP Permutation
perm')

  forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> TCMT IO Doc -> m ()
reportSDoc [Char]
"tc.with.top" Int
10 forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
vcat
    [ TCMT IO Doc
"checkWithFunction"
    , forall (m :: * -> *). Functor m => Int -> m Doc -> m Doc
nest Int
2 forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
vcat forall a b. (a -> b) -> a -> b
$
      let ([Term]
vs, [EqualityView]
as) = forall a b c. (a -> (b, c)) -> [a] -> ([b], [c])
unzipWith forall e. Arg e -> e
unArg [Arg (Term, EqualityView)]
vtys in
      [ TCMT IO Doc
"delta1 =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Tele (Dom Type)
delta1
      , TCMT IO Doc
"delta2 =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall b (m :: * -> *) a.
(AddContext b, MonadAddContext m) =>
b -> m a -> m a
addContext Tele (Dom Type)
delta1 (forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Tele (Dom Type)
delta2)
      , TCMT IO Doc
"t      =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Type
t
      , TCMT IO Doc
"as     =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall b (m :: * -> *) a.
(AddContext b, MonadAddContext m) =>
b -> m a -> m a
addContext Tele (Dom Type)
delta1 (forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM [EqualityView]
as)
      , TCMT IO Doc
"vs     =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> do forall b (m :: * -> *) a.
(AddContext b, MonadAddContext m) =>
b -> m a -> m a
addContext Tele (Dom Type)
delta1 forall a b. (a -> b) -> a -> b
$ forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM [Term]
vs
      , TCMT IO Doc
"b      =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> do forall b (m :: * -> *) a.
(AddContext b, MonadAddContext m) =>
b -> m a -> m a
addContext Tele (Dom Type)
delta1 forall a b. (a -> b) -> a -> b
$ forall b (m :: * -> *) a.
(AddContext b, MonadAddContext m) =>
b -> m a -> m a
addContext Tele (Dom Type)
delta2 forall a b. (a -> b) -> a -> b
$ forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Type
b
      , TCMT IO Doc
"qs     =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> do forall b (m :: * -> *) a.
(AddContext b, MonadAddContext m) =>
b -> m a -> m a
addContext Tele (Dom Type)
delta forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *).
MonadPretty m =>
[NamedArg DeBruijnPattern] -> m Doc
prettyTCMPatternList [NamedArg DeBruijnPattern]
qs
      , TCMT IO Doc
"perm'  =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall (m :: * -> *). Applicative m => [Char] -> m Doc
text (forall a. Show a => a -> [Char]
show Permutation
perm')
      , TCMT IO Doc
"perm   =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall (m :: * -> *). Applicative m => [Char] -> m Doc
text (forall a. Show a => a -> [Char]
show Permutation
perm)
      , TCMT IO Doc
"fperm  =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall (m :: * -> *). Applicative m => [Char] -> m Doc
text (forall a. Show a => a -> [Char]
show Permutation
finalPerm)
      , TCMT IO Doc
"withSub=" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall (m :: * -> *). Applicative m => [Char] -> m Doc
text (forall a. Show a => a -> [Char]
show Substitution
withSub)
      ]
    ]

  -- Add the type of the auxiliary function to the signature

  -- Jesper, 2020-04-05: Currently variable generalization inserts
  -- dummy terms, we have to reduce projections to get rid of them.
  -- (see also #1332).
  let reds :: SmallSet AllowedReduction
reds = forall a. SmallSetElement a => [a] -> SmallSet a
SmallSet.fromList [AllowedReduction
ProjectionReductions]
  Tele (Dom Type)
delta1 <- forall (m :: * -> *) a.
MonadTCEnv m =>
(SmallSet AllowedReduction -> SmallSet AllowedReduction)
-> m a -> m a
modifyAllowedReductions (forall a b. a -> b -> a
const SmallSet AllowedReduction
reds) forall a b. (a -> b) -> a -> b
$ forall a (m :: * -> *). (Normalise a, MonadReduce m) => a -> m a
normalise Tele (Dom Type)
delta1

  -- Generate the type of the with function
  (Type
withFunType, Int
n) <- do
    let ps :: [NamedArg DeBruijnPattern]
ps = forall a.
DeBruijn a =>
Impossible -> Permutation -> Substitution' a
renaming HasCallStack => Impossible
impossible (Permutation -> Permutation
reverseP Permutation
perm') forall a. Subst a => Substitution' (SubstArg a) -> a -> a
`applySubst` [NamedArg DeBruijnPattern]
qs
    forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> TCMT IO Doc -> m ()
reportSDoc [Char]
"tc.with.bndry" Int
40 forall a b. (a -> b) -> a -> b
$ forall b (m :: * -> *) a.
(AddContext b, MonadAddContext m) =>
b -> m a -> m a
addContext Tele (Dom Type)
delta1 forall a b. (a -> b) -> a -> b
$ forall b (m :: * -> *) a.
(AddContext b, MonadAddContext m) =>
b -> m a -> m a
addContext Tele (Dom Type)
delta2
                                  forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *). Applicative m => [Char] -> m Doc
text [Char]
"ps =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall (m :: * -> *) a. (Applicative m, Pretty a) => a -> m Doc
pretty [NamedArg DeBruijnPattern]
ps
    let vs :: [Int]
vs = forall p. IApplyVars p => p -> [Int]
iApplyVars [NamedArg DeBruijnPattern]
ps
    [(Int, (Term, Term))]
bndry <- if forall a. Null a => a -> Bool
null [Int]
vs then forall (m :: * -> *) a. Monad m => a -> m a
return [] else do
      Term
iz <- forall (m :: * -> *).
(HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) =>
m Term
primIZero
      Term
io <- forall (m :: * -> *).
(HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) =>
m Term
primIOne
      let tm :: Term
tm = QName -> Elims -> Term
Def QName
f ([NamedArg DeBruijnPattern] -> Elims
patternsToElims [NamedArg DeBruijnPattern]
ps)
      forall (m :: * -> *) a. Monad m => a -> m a
return [(Int
i,(forall a. EndoSubst a => Int -> a -> Substitution' a
inplaceS Int
i Term
iz forall a. Subst a => Substitution' (SubstArg a) -> a -> a
`applySubst` Term
tm, forall a. EndoSubst a => Int -> a -> Substitution' a
inplaceS Int
i Term
io forall a. Subst a => Substitution' (SubstArg a) -> a -> a
`applySubst` Term
tm)) | Int
i <- [Int]
vs]
    forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> TCMT IO Doc -> m ()
reportSDoc [Char]
"tc.with.bndry" Int
40 forall a b. (a -> b) -> a -> b
$ forall b (m :: * -> *) a.
(AddContext b, MonadAddContext m) =>
b -> m a -> m a
addContext Tele (Dom Type)
delta1 forall a b. (a -> b) -> a -> b
$ forall b (m :: * -> *) a.
(AddContext b, MonadAddContext m) =>
b -> m a -> m a
addContext Tele (Dom Type)
delta2
                                  forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *). Applicative m => [Char] -> m Doc
text [Char]
"bndry =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall (m :: * -> *) a. (Applicative m, Pretty a) => a -> m Doc
pretty [(Int, (Term, Term))]
bndry
    Tele (Dom Type)
-> [Arg (Term, EqualityView)]
-> Tele (Dom Type)
-> Type
-> [(Int, (Term, Term))]
-> TCMT IO (Type, Int)
withFunctionType Tele (Dom Type)
delta1 [Arg (Term, EqualityView)]
vtys Tele (Dom Type)
delta2 Type
b [(Int, (Term, Term))]
bndry
  forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> TCMT IO Doc -> m ()
reportSDoc [Char]
"tc.with.type" Int
10 forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
sep [ TCMT IO Doc
"with-function type:", forall (m :: * -> *). Functor m => Int -> m Doc -> m Doc
nest Int
2 forall a b. (a -> b) -> a -> b
$ forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Type
withFunType ]
  forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> TCMT IO Doc -> m ()
reportSDoc [Char]
"tc.with.type" Int
50 forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
sep [ TCMT IO Doc
"with-function type:", forall (m :: * -> *). Functor m => Int -> m Doc -> m Doc
nest Int
2 forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *) a. (Applicative m, Pretty a) => a -> m Doc
pretty Type
withFunType ]

  Term
call_in_parent <- do
    (TelV Tele (Dom Type)
tel Type
_,Boundary
bs) <- forall (m :: * -> *).
PureTCM m =>
Int -> Type -> m (TelV Type, Boundary)
telViewUpToPathBoundaryP (Int
n forall a. Num a => a -> a -> a
+ forall a. Sized a => a -> Int
size Tele (Dom Type)
delta) Type
withFunType
    forall (m :: * -> *) a. Monad m => a -> m a
return forall a b. (a -> b) -> a -> b
$ Substitution
argsS forall a. Subst a => Substitution' (SubstArg a) -> a -> a
`applySubst` QName -> Elims -> Term
Def QName
aux (forall a.
DeBruijn a =>
Tele (Dom Type) -> Boundary' (a, a) -> [Elim' a]
teleElims Tele (Dom Type)
tel Boundary
bs)

  forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> TCMT IO Doc -> m ()
reportSDoc [Char]
"tc.with.top" Int
20 forall a b. (a -> b) -> a -> b
$ forall b (m :: * -> *) a.
(AddContext b, MonadAddContext m) =>
b -> m a -> m a
addContext Tele (Dom Type)
delta forall a b. (a -> b) -> a -> b
$ TCMT IO Doc
"with function call" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Term
call_in_parent

  -- Andreas, 2013-10-21
  -- Check generated type directly in internal syntax.
  forall (m :: * -> *) x a.
(MonadTrace m, HasRange x) =>
x -> m a -> m a
setCurrentRange List1 (Clause' LHS)
cs forall a b. (a -> b) -> a -> b
$
    forall (m :: * -> *) a. MonadTrace m => Call -> m a -> m a
traceCall Call
NoHighlighting forall a b. (a -> b) -> a -> b
$   -- To avoid flicker.
    forall (m :: * -> *) a. MonadTrace m => Call -> m a -> m a
traceCall (Type -> Call
CheckWithFunctionType Type
withFunType) forall a b. (a -> b) -> a -> b
$
    forall (m :: * -> *). MonadCheckInternal m => Type -> m ()
checkType Type
withFunType

  -- With display forms are closed
  Open DisplayForm
df <- forall (tcm :: * -> *) a.
(MonadTCEnv tcm, ReadTCState tcm) =>
tcm a -> tcm a
inTopContext forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *) a.
(ReadTCState m, MonadTCEnv m) =>
a -> m (Open a)
makeOpen forall (m :: * -> *) a b. Monad m => (a -> m b) -> m a -> m b
=<< QName
-> QName
-> Tele (Dom Type)
-> Tele (Dom Type)
-> Int
-> [NamedArg DeBruijnPattern]
-> Permutation
-> Permutation
-> TCM DisplayForm
withDisplayForm QName
f QName
aux Tele (Dom Type)
delta1 Tele (Dom Type)
delta2 Int
n [NamedArg DeBruijnPattern]
qs Permutation
perm' Permutation
perm

  forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> [Char] -> m ()
reportSLn [Char]
"tc.with.top" Int
20 [Char]
"created with display form"

  case forall (t :: * -> *) a. Decoration t => t a -> a
dget Open DisplayForm
df of
    Display Int
n Elims
ts DisplayTerm
dt ->
      forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> TCMT IO Doc -> m ()
reportSDoc [Char]
"tc.with.top" Int
20 forall a b. (a -> b) -> a -> b
$ TCMT IO Doc
"Display" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
fsep
        [ forall (m :: * -> *). Applicative m => [Char] -> m Doc
text (forall a. Show a => a -> [Char]
show Int
n)
        , forall (m :: * -> *) (t :: * -> *).
(Applicative m, Semigroup (m Doc), Foldable t) =>
t (m Doc) -> m Doc
prettyList forall a b. (a -> b) -> a -> b
$ forall a b. (a -> b) -> [a] -> [b]
map forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Elims
ts
        , forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM DisplayTerm
dt
        ]
  QName -> Definition -> TCM ()
addConstant QName
aux forall (m :: * -> *) a b. Monad m => (a -> m b) -> m a -> m b
=<< do
    Language
lang <- forall (m :: * -> *). HasOptions m => m Language
getLanguage
    Defn
fun  <- forall (m :: * -> *). HasOptions m => m Defn
emptyFunction
    Definition -> TCMT IO Definition
useTerPragma forall a b. (a -> b) -> a -> b
$
      (ArgInfo -> QName -> Type -> Language -> Defn -> Definition
defaultDefn ArgInfo
defaultArgInfo QName
aux Type
withFunType Language
lang Defn
fun)
        { defDisplay :: [Open DisplayForm]
defDisplay = [Open DisplayForm
df] }
  -- solveSizeConstraints -- Andreas, 2012-10-16 does not seem necessary

  forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> TCMT IO Doc -> m ()
reportSDoc [Char]
"tc.with.top" Int
10 forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
sep
    [ TCMT IO Doc
"added with function" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> (forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM QName
aux) forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> TCMT IO Doc
"of type"
    , forall (m :: * -> *). Functor m => Int -> m Doc -> m Doc
nest Int
2 forall a b. (a -> b) -> a -> b
$ forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Type
withFunType
    , forall (m :: * -> *). Functor m => Int -> m Doc -> m Doc
nest Int
2 forall a b. (a -> b) -> a -> b
$ TCMT IO Doc
"-|" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> (forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM forall (m :: * -> *) a b. Monad m => (a -> m b) -> m a -> m b
=<< forall (m :: * -> *).
(Applicative m, MonadTCEnv m) =>
m (Tele (Dom Type))
getContextTelescope)
    ]
  forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> TCMT IO Doc -> m ()
reportSDoc [Char]
"tc.with.top" Int
70 forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
vcat
    [ forall (m :: * -> *). Functor m => Int -> m Doc -> m Doc
nest Int
2 forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *). Applicative m => [Char] -> m Doc
text forall a b. (a -> b) -> a -> b
$ [Char]
"raw with func. type = " forall a. [a] -> [a] -> [a]
++ forall a. Show a => a -> [Char]
show Type
withFunType
    ]


  -- Construct the body for the with function
  NonEmpty SpineClause
cs <- forall (m :: * -> *) a. Monad m => a -> m a
return forall a b. (a -> b) -> a -> b
$ forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
fmap (forall a b. LHSToSpine a b => a -> b
A.lhsToSpine) List1 (Clause' LHS)
cs
  NonEmpty SpineClause
cs <- [Name]
-> QName
-> QName
-> Type
-> Tele (Dom Type)
-> [NamedArg DeBruijnPattern]
-> Int
-> Substitution
-> Permutation
-> Int
-> Int
-> NonEmpty SpineClause
-> TCMT IO (NonEmpty SpineClause)
buildWithFunction [Name]
cxtNames QName
f QName
aux Type
t Tele (Dom Type)
delta [NamedArg DeBruijnPattern]
qs Int
npars Substitution
withSub Permutation
finalPerm (forall a. Sized a => a -> Int
size Tele (Dom Type)
delta1) Int
n NonEmpty SpineClause
cs
  List1 (Clause' LHS)
cs <- forall (m :: * -> *) a. Monad m => a -> m a
return forall a b. (a -> b) -> a -> b
$ forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
fmap (forall a b. LHSToSpine a b => b -> a
A.spineToLhs) NonEmpty SpineClause
cs

  -- #4833: inherit abstract mode from parent
  IsAbstract
abstr <- Definition -> IsAbstract
defAbstract forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> forall (m :: * -> *) a. MonadTCEnv m => m a -> m a
ignoreAbstractMode (forall (m :: * -> *). HasConstInfo m => QName -> m Definition
getConstInfo QName
f)

  -- Check the with function
  let info :: DefInfo
info = forall t.
Name -> Fixity' -> Access -> IsAbstract -> Range -> DefInfo' t
Info.mkDefInfo (Name -> Name
nameConcrete forall a b. (a -> b) -> a -> b
$ QName -> Name
qnameName QName
aux) Fixity'
noFixity' Access
PublicAccess IsAbstract
abstr (forall a. HasRange a => a -> Range
getRange List1 (Clause' LHS)
cs)
  ArgInfo
ai <- Definition -> ArgInfo
defArgInfo forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> forall (m :: * -> *). HasConstInfo m => QName -> m Definition
getConstInfo QName
f
  Type
-> ArgInfo
-> Maybe ExtLamInfo
-> Maybe QName
-> DefInfo
-> QName
-> Maybe Substitution
-> [Clause' LHS]
-> TCM ()
checkFunDefS Type
withFunType ArgInfo
ai forall a. Maybe a
Nothing (forall a. a -> Maybe a
Just QName
f) DefInfo
info QName
aux (forall a. a -> Maybe a
Just Substitution
withSub) forall a b. (a -> b) -> a -> b
$ forall l. IsList l => l -> [Item l]
List1.toList List1 (Clause' LHS)
cs
  forall (m :: * -> *) a. Monad m => a -> m a
return forall a b. (a -> b) -> a -> b
$ forall a. a -> Maybe a
Just forall a b. (a -> b) -> a -> b
$ Term
call_in_parent

-- | Type check a where clause.
checkWhere
  :: A.WhereDeclarations -- ^ Where-declarations to check.
  -> TCM a               -- ^ Continuation.
  -> TCM a
checkWhere :: forall a. WhereDeclarations -> TCM a -> TCM a
checkWhere wh :: WhereDeclarations
wh@(A.WhereDecls Maybe ModuleName
whmod Bool
whNamed Maybe Declaration
ds) TCM a
ret = do
  forall (f :: * -> *). Applicative f => Bool -> f () -> f ()
when (Bool -> Bool
not Bool
whNamed) forall a b. (a -> b) -> a -> b
$ forall {m :: * -> *}.
(MonadTrace m, HasOptions m, MonadDebug m, MonadError TCErr m) =>
Maybe ModuleName -> m ()
ensureNoNamedWhereInRefinedContext Maybe ModuleName
whmod
  Maybe Declaration -> TCM a
loop Maybe Declaration
ds
  where
    loop :: Maybe Declaration -> TCM a
loop = \case
      Maybe Declaration
Nothing -> TCM a
ret
      -- [A.ScopedDecl scope ds] -> withScope_ scope $ loop ds  -- IMPOSSIBLE
      Just (A.Section Range
_ Erased
e ModuleName
m GeneralizeTelescope
tel [Declaration]
ds) -> forall a.
Erased -> ModuleName -> GeneralizeTelescope -> TCM a -> TCM a
newSection Erased
e ModuleName
m GeneralizeTelescope
tel forall a b. (a -> b) -> a -> b
$ do
          forall (m :: * -> *) a.
MonadTCEnv m =>
(TCEnv -> TCEnv) -> m a -> m a
localTC (\ TCEnv
e -> TCEnv
e { envCheckingWhere :: Bool
envCheckingWhere = Bool
True }) forall a b. (a -> b) -> a -> b
$ do
            [Declaration] -> TCM ()
checkDecls [Declaration]
ds
            TCM a
ret
      Maybe Declaration
_ -> forall a. HasCallStack => a
__IMPOSSIBLE__

    -- #2897: We can't handle named where-modules in refined contexts.
    ensureNoNamedWhereInRefinedContext :: Maybe ModuleName -> m ()
ensureNoNamedWhereInRefinedContext Maybe ModuleName
Nothing = forall (m :: * -> *) a. Monad m => a -> m a
return ()
    ensureNoNamedWhereInRefinedContext (Just ModuleName
m) = forall (m :: * -> *) a. MonadTrace m => Call -> m a -> m a
traceCall (ModuleName -> Call
CheckNamedWhere ModuleName
m) forall a b. (a -> b) -> a -> b
$ do
      [Term]
args <- forall a b. (a -> b) -> [a] -> [b]
map forall e. Arg e -> e
unArg forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> (forall (m :: * -> *).
(Functor m, Applicative m, HasOptions m, MonadTCEnv m,
 ReadTCState m, MonadDebug m) =>
ModuleName -> m [Arg Term]
moduleParamsToApply forall (m :: * -> *) a b. Monad m => (a -> m b) -> m a -> m b
=<< forall (m :: * -> *). MonadTCEnv m => m ModuleName
currentModule)
      forall (f :: * -> *). Applicative f => Bool -> f () -> f ()
unless ([Term] -> Bool
isWeakening [Term]
args) forall a b. (a -> b) -> a -> b
$ do -- weakened contexts are fine
        [[Char]]
names <- forall a b. (a -> b) -> [a] -> [b]
map ([Char] -> [Char]
argNameToString forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall a b. (a, b) -> a
fst forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall t e. Dom' t e -> e
unDom) forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall t. Tele (Dom t) -> [Dom ([Char], t)]
telToList forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$>
                (forall (m :: * -> *).
(Functor m, ReadTCState m) =>
ModuleName -> m (Tele (Dom Type))
lookupSection forall (m :: * -> *) a b. Monad m => (a -> m b) -> m a -> m b
=<< forall (m :: * -> *). MonadTCEnv m => m ModuleName
currentModule)
        forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
TypeError -> m a
typeError forall a b. (a -> b) -> a -> b
$ [Term] -> [[Char]] -> TypeError
NamedWhereModuleInRefinedContext [Term]
args [[Char]]
names
      where
        isWeakening :: [Term] -> Bool
isWeakening [] = Bool
True
        isWeakening (Var Int
i [] : [Term]
args) = Int -> [Term] -> Bool
isWk (Int
i forall a. Num a => a -> a -> a
- Int
1) [Term]
args
          where
            isWk :: Int -> [Term] -> Bool
isWk Int
i []                = Bool
True
            isWk Int
i (Var Int
j [] : [Term]
args) = Int
i forall a. Eq a => a -> a -> Bool
== Int
j Bool -> Bool -> Bool
&& Int -> [Term] -> Bool
isWk (Int
i forall a. Num a => a -> a -> a
- Int
1) [Term]
args
            isWk Int
_ [Term]
_ = Bool
False
        isWeakening [Term]
_ = Bool
False


-- | Enter a new section during type-checking.

newSection ::
  Erased -> ModuleName -> A.GeneralizeTelescope -> TCM a -> TCM a
newSection :: forall a.
Erased -> ModuleName -> GeneralizeTelescope -> TCM a -> TCM a
newSection Erased
e ModuleName
m gtel :: GeneralizeTelescope
gtel@(A.GeneralizeTel Map QName Name
_ Telescope
tel) TCM a
cont = do
  -- If the section is erased, then hard compile-time mode is entered.
  Erased -> TCM ()
warnForPlentyInHardCompileTimeMode Erased
e
  forall a. Erased -> TCM a -> TCM a
setHardCompileTimeModeIfErased Erased
e forall a b. (a -> b) -> a -> b
$ do
  forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> TCMT IO Doc -> m ()
reportSDoc [Char]
"tc.section" Int
10 forall a b. (a -> b) -> a -> b
$
    TCMT IO Doc
"checking section" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> (Erased -> Doc -> Doc
C.prettyErased Erased
e forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM ModuleName
m) forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+>
    forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
fsep (forall a b. (a -> b) -> [a] -> [b]
map forall a (m :: * -> *).
(ToConcrete a, Pretty (ConOfAbs a), MonadAbsToCon m) =>
a -> m Doc
prettyA Telescope
tel)

  forall a.
Maybe ModuleName
-> GeneralizeTelescope
-> ([Maybe Name] -> Tele (Dom Type) -> TCM a)
-> TCM a
checkGeneralizeTelescope (forall a. a -> Maybe a
Just ModuleName
m) GeneralizeTelescope
gtel forall a b. (a -> b) -> a -> b
$ \ [Maybe Name]
_ Tele (Dom Type)
tel' -> do
    forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> TCMT IO Doc -> m ()
reportSDoc [Char]
"tc.section" Int
10 forall a b. (a -> b) -> a -> b
$
      TCMT IO Doc
"adding section:" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM ModuleName
m forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall (m :: * -> *). Applicative m => [Char] -> m Doc
text (forall a. Show a => a -> [Char]
show (forall a. Sized a => a -> Int
size Tele (Dom Type)
tel'))

    ModuleName -> TCM ()
addSection ModuleName
m

    forall (m :: * -> *).
MonadDebug m =>
[Char] -> Int -> TCMT IO Doc -> m ()
reportSDoc [Char]
"tc.section" Int
10 forall a b. (a -> b) -> a -> b
$ forall (tcm :: * -> *) a.
(MonadTCEnv tcm, ReadTCState tcm) =>
tcm a -> tcm a
inTopContext forall a b. (a -> b) -> a -> b
$
      forall (m :: * -> *). Functor m => Int -> m Doc -> m Doc
nest Int
4 forall a b. (a -> b) -> a -> b
$ TCMT IO Doc
"actual tele:" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> do forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM forall (m :: * -> *) a b. Monad m => (a -> m b) -> m a -> m b
=<< forall (m :: * -> *).
(Functor m, ReadTCState m) =>
ModuleName -> m (Tele (Dom Type))
lookupSection ModuleName
m

    forall (m :: * -> *) a. MonadTCEnv m => ModuleName -> m a -> m a
withCurrentModule ModuleName
m TCM a
cont

-- | Set the current clause number.

atClause :: QName -> Int -> Type -> Maybe Substitution -> A.SpineClause -> TCM a -> TCM a
atClause :: forall a.
QName
-> Int
-> Type
-> Maybe Substitution
-> SpineClause
-> TCM a
-> TCM a
atClause QName
name Int
i Type
t Maybe Substitution
sub SpineClause
cl TCM a
ret = do
  Closure ()
clo <- forall (m :: * -> *) a.
(MonadTCEnv m, ReadTCState m) =>
a -> m (Closure a)
buildClosure ()
  forall (m :: * -> *) a.
MonadTCEnv m =>
(TCEnv -> TCEnv) -> m a -> m a
localTC (\ TCEnv
e -> TCEnv
e { envClause :: IPClause
envClause = QName
-> Int
-> Type
-> Maybe Substitution
-> SpineClause
-> Closure ()
-> IPClause
IPClause QName
name Int
i Type
t Maybe Substitution
sub SpineClause
cl Closure ()
clo }) TCM a
ret