{-# LANGUAGE NondecreasingIndentation #-}

module Agda.TypeChecking.Conversion where

import Control.Arrow (second)
import Control.Monad
import Control.Monad.Except
-- Control.Monad.Fail import is redundant since GHC 8.8.1
import Control.Monad.Fail (MonadFail)

import Data.Function
import Data.Semigroup ((<>))
import Data.IntMap (IntMap)

import qualified Data.List   as List
import qualified Data.IntMap as IntMap
import qualified Data.IntSet as IntSet
import qualified Data.Set    as Set

import Agda.Syntax.Common
import Agda.Syntax.Internal
import Agda.Syntax.Internal.MetaVars
import Agda.Syntax.Translation.InternalToAbstract (reify)

import Agda.TypeChecking.Monad
import Agda.TypeChecking.MetaVars
import Agda.TypeChecking.MetaVars.Occurs (killArgs,PruneResult(..),rigidVarsNotContainedIn)
import Agda.TypeChecking.Names
import Agda.TypeChecking.Reduce
import Agda.TypeChecking.Substitute
import qualified Agda.TypeChecking.SyntacticEquality as SynEq
import Agda.TypeChecking.Telescope
import Agda.TypeChecking.Constraints
import Agda.TypeChecking.Conversion.Pure (pureCompareAs, runPureConversion)
import {-# SOURCE #-} Agda.TypeChecking.CheckInternal (infer)
import Agda.TypeChecking.Forcing (isForced, nextIsForced)
import Agda.TypeChecking.Free
import Agda.TypeChecking.Datatypes (getConType, getFullyAppliedConType)
import Agda.TypeChecking.Records
import Agda.TypeChecking.Pretty
import Agda.TypeChecking.Injectivity
import Agda.TypeChecking.Polarity
import Agda.TypeChecking.SizedTypes
import Agda.TypeChecking.Level
import Agda.TypeChecking.Implicit (implicitArgs)
import Agda.TypeChecking.Irrelevance
import Agda.TypeChecking.Primitive
import Agda.TypeChecking.Warnings (MonadWarning)
import Agda.Interaction.Options

import Agda.Utils.Functor
import Agda.Utils.List1 (List1, pattern (:|))
import qualified Agda.Utils.List1 as List1
import Agda.Utils.Monad
import Agda.Utils.Maybe
import Agda.Utils.Permutation
import Agda.Utils.Pretty (prettyShow)
import qualified Agda.Utils.ProfileOptions as Profile
import Agda.Utils.BoolSet (BoolSet)
import qualified Agda.Utils.BoolSet as BoolSet
import Agda.Utils.Size
import Agda.Utils.Tuple
import Agda.Utils.WithDefault

import Agda.Utils.Impossible

type MonadConversion m =
  ( PureTCM m
  , MonadConstraint m
  , MonadMetaSolver m
  , MonadError TCErr m
  , MonadWarning m
  , MonadStatistics m
  , MonadFresh ProblemId m
  , MonadFresh Int m
  , MonadFail m
  )

-- | Try whether a computation runs without errors or new constraints
--   (may create new metas, though).
--   Restores state upon failure.
tryConversion
  :: (MonadConstraint m, MonadWarning m, MonadError TCErr m, MonadFresh ProblemId m)
  => m () -> m Bool
tryConversion :: forall (m :: * -> *).
(MonadConstraint m, MonadWarning m, MonadError TCErr m,
 MonadFresh ProblemId m) =>
m () -> m Bool
tryConversion = forall a. Maybe a -> Bool
isJust forall (m :: * -> *) b c a.
Functor m =>
(b -> c) -> (a -> m b) -> a -> m c
<.> forall (m :: * -> *) a.
(MonadConstraint m, MonadWarning m, MonadError TCErr m,
 MonadFresh ProblemId m) =>
m a -> m (Maybe a)
tryConversion'

-- | Try whether a computation runs without errors or new constraints
--   (may create new metas, though).
--   Return 'Just' the result upon success.
--   Return 'Nothing' and restore state upon failure.
tryConversion'
  :: (MonadConstraint m, MonadWarning m, MonadError TCErr m, MonadFresh ProblemId m)
  => m a -> m (Maybe a)
tryConversion' :: forall (m :: * -> *) a.
(MonadConstraint m, MonadWarning m, MonadError TCErr m,
 MonadFresh ProblemId m) =>
m a -> m (Maybe a)
tryConversion' m a
m = forall e (m :: * -> *) a.
(MonadError e m, Functor m) =>
m a -> m (Maybe a)
tryMaybe forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *) a.
(MonadConstraint m, MonadWarning m, MonadError TCErr m,
 MonadFresh ProblemId m) =>
m a -> m a
noConstraints m a
m

-- | Check if to lists of arguments are the same (and all variables).
--   Precondition: the lists have the same length.
sameVars :: Elims -> Elims -> Bool
sameVars :: Elims -> Elims -> Bool
sameVars Elims
xs Elims
ys = forall (t :: * -> *). Foldable t => t Bool -> Bool
and forall a b. (a -> b) -> a -> b
$ forall a b c. (a -> b -> c) -> [a] -> [b] -> [c]
zipWith Elim' Term -> Elim' Term -> Bool
same Elims
xs Elims
ys
    where
        same :: Elim' Term -> Elim' Term -> Bool
same (Apply (Arg ArgInfo
_ (Var Int
n []))) (Apply (Arg ArgInfo
_ (Var Int
m []))) = Int
n forall a. Eq a => a -> a -> Bool
== Int
m
        same Elim' Term
_ Elim' Term
_ = Bool
False

-- | @intersectVars us vs@ checks whether all relevant elements in @us@ and @vs@
--   are variables, and if yes, returns a prune list which says @True@ for
--   arguments which are different and can be pruned.
intersectVars :: Elims -> Elims -> Maybe [Bool]
intersectVars :: Elims -> Elims -> Maybe [Bool]
intersectVars = forall (m :: * -> *) a b c.
Applicative m =>
(a -> b -> m c) -> [a] -> [b] -> m [c]
zipWithM Elim' Term -> Elim' Term -> Maybe Bool
areVars where
    -- ignore irrelevant args
    areVars :: Elim' Term -> Elim' Term -> Maybe Bool
areVars (Apply Arg Term
u) Elim' Term
v | forall a. LensRelevance a => a -> Bool
isIrrelevant Arg Term
u = forall a. a -> Maybe a
Just Bool
False -- do not prune
    areVars (Apply (Arg ArgInfo
_ (Var Int
n []))) (Apply (Arg ArgInfo
_ (Var Int
m []))) = forall a. a -> Maybe a
Just forall a b. (a -> b) -> a -> b
$ Int
n forall a. Eq a => a -> a -> Bool
/= Int
m -- prune different vars
    areVars Elim' Term
_ Elim' Term
_                                   = forall a. Maybe a
Nothing

-- | Run the given computation but turn any errors into blocked computations with the given blocker
blockOnError :: MonadError TCErr m => Blocker -> m a -> m a
blockOnError :: forall (m :: * -> *) a. MonadError TCErr m => Blocker -> m a -> m a
blockOnError Blocker
blocker m a
f
  | Blocker
blocker forall a. Eq a => a -> a -> Bool
== Blocker
neverUnblock = m a
f
  | Bool
otherwise               = m a
f forall e (m :: * -> *) a.
MonadError e m =>
m a -> (e -> m a) -> m a
`catchError` \case
    TypeError{}         -> forall e (m :: * -> *) a. MonadError e m => e -> m a
throwError forall a b. (a -> b) -> a -> b
$ Blocker -> TCErr
PatternErr Blocker
blocker
    PatternErr Blocker
blocker' -> forall e (m :: * -> *) a. MonadError e m => e -> m a
throwError forall a b. (a -> b) -> a -> b
$ Blocker -> TCErr
PatternErr forall a b. (a -> b) -> a -> b
$ Blocker -> Blocker -> Blocker
unblockOnEither Blocker
blocker Blocker
blocker'
    err :: TCErr
err@Exception{}     -> forall e (m :: * -> *) a. MonadError e m => e -> m a
throwError TCErr
err
    err :: TCErr
err@IOException{}   -> forall e (m :: * -> *) a. MonadError e m => e -> m a
throwError TCErr
err

equalTerm :: MonadConversion m => Type -> Term -> Term -> m ()
equalTerm :: forall (m :: * -> *).
MonadConversion m =>
Type -> Term -> Term -> m ()
equalTerm = forall (m :: * -> *).
MonadConversion m =>
Comparison -> Type -> Term -> Term -> m ()
compareTerm Comparison
CmpEq

equalAtom :: MonadConversion m => CompareAs -> Term -> Term -> m ()
equalAtom :: forall (m :: * -> *).
MonadConversion m =>
CompareAs -> Term -> Term -> m ()
equalAtom = forall (m :: * -> *).
MonadConversion m =>
Comparison -> CompareAs -> Term -> Term -> m ()
compareAtom Comparison
CmpEq

equalType :: MonadConversion m => Type -> Type -> m ()
equalType :: forall (m :: * -> *). MonadConversion m => Type -> Type -> m ()
equalType = forall (m :: * -> *).
MonadConversion m =>
Comparison -> Type -> Type -> m ()
compareType Comparison
CmpEq

{- Comparing in irrelevant context always succeeds.

   However, we might want to dig for solutions of irrelevant metas.

   To this end, we can just ignore errors during conversion checking.
 -}

-- convError ::  MonadTCM tcm => TypeError -> tcm a
-- | Ignore errors in irrelevant context.
convError :: TypeError -> TCM ()
convError :: TypeError -> TCM ()
convError TypeError
err = forall (m :: * -> *) a. Monad m => m Bool -> m a -> m a -> m a
ifM (forall a. Eq a => a -> a -> Bool
(==) Relevance
Irrelevant forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> forall (m :: * -> *) a. MonadTCEnv m => (TCEnv -> a) -> m a
asksTC forall a. LensRelevance a => a -> Relevance
getRelevance) (forall (m :: * -> *) a. Monad m => a -> m a
return ()) forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
TypeError -> m a
typeError TypeError
err

-- | Type directed equality on values.
--
compareTerm :: forall m. MonadConversion m => Comparison -> Type -> Term -> Term -> m ()
compareTerm :: forall (m :: * -> *).
MonadConversion m =>
Comparison -> Type -> Term -> Term -> m ()
compareTerm Comparison
cmp Type
a Term
u Term
v = forall (m :: * -> *).
MonadConversion m =>
Comparison -> CompareAs -> Term -> Term -> m ()
compareAs Comparison
cmp (Type -> CompareAs
AsTermsOf Type
a) Term
u Term
v

-- | Type directed equality on terms or types.
compareAs :: forall m. MonadConversion m => Comparison -> CompareAs -> Term -> Term -> m ()
  -- If one term is a meta, try to instantiate right away. This avoids unnecessary unfolding.
  -- Andreas, 2012-02-14: This is UNSOUND for subtyping!
compareAs :: forall (m :: * -> *).
MonadConversion m =>
Comparison -> CompareAs -> Term -> Term -> m ()
compareAs Comparison
cmp CompareAs
a Term
u Term
v = do
  forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"tc.conv.term" Int
20 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
"compareTerm"
    , 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 Term
u forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Comparison
cmp 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 :: * -> *). 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 CompareAs
a
    ]
  forall (m :: * -> *). MonadDebug m => ProfileOption -> m () -> m ()
whenProfile ProfileOption
Profile.Conversion forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *). MonadStatistics m => VerboseKey -> m ()
tick VerboseKey
"compare"

  -- OLD CODE, traverses the *full* terms u v at each step, even if they
  -- are different somewhere.  Leads to infeasibility in issue 854.
  -- (u, v) <- instantiateFull (u, v)
  -- let equal = u == v

  -- Check syntactic equality. This actually saves us quite a bit of work.
  forall a (m :: * -> *) b.
(Instantiate a, SynEq a, MonadReduce m) =>
a -> a -> (a -> a -> m b) -> (a -> a -> m b) -> m b
SynEq.checkSyntacticEquality Term
u Term
v
    (\Term
_ Term
_ -> forall (m :: * -> *). MonadDebug m => ProfileOption -> m () -> m ()
whenProfile ProfileOption
Profile.Conversion forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *). MonadStatistics m => VerboseKey -> m ()
tick VerboseKey
"compare equal") forall a b. (a -> b) -> a -> b
$
    \Term
u Term
v -> do
      forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"tc.conv.term" Int
15 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
"compareTerm (not syntactically equal)"
        , 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 Term
u forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Comparison
cmp 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 :: * -> *). 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 CompareAs
a
        ]
      -- If we are at type Size, we cannot short-cut comparison
      -- against metas by assignment.
      -- Andreas, 2014-04-12: this looks incomplete.
      -- It seems to assume we are never comparing
      -- at function types into Size.
      let fallback :: m ()
fallback = forall (m :: * -> *).
MonadConversion m =>
Comparison -> CompareAs -> Term -> Term -> m ()
compareAs' Comparison
cmp CompareAs
a Term
u Term
v
          unlessSubtyping :: m () -> m ()
          unlessSubtyping :: m () -> m ()
unlessSubtyping m ()
cont =
              if Comparison
cmp forall a. Eq a => a -> a -> Bool
== Comparison
CmpEq then m ()
cont else do
                -- Andreas, 2014-04-12 do not short cut if type is blocked.
                forall t (m :: * -> *) a.
(Reduce t, IsMeta t, MonadReduce m) =>
t -> (Blocker -> t -> m a) -> (NotBlocked -> t -> m a) -> m a
ifBlocked CompareAs
a (\ Blocker
_ CompareAs
_ -> m ()
fallback) {-else-} forall a b. (a -> b) -> a -> b
$ \ NotBlocked
_ CompareAs
a -> do
                  -- do not short circuit size comparison!
                  forall (m :: * -> *) a b.
Monad m =>
m (Maybe a) -> m b -> (a -> m b) -> m b
caseMaybeM (forall a (m :: * -> *).
(IsSizeType a, HasOptions m, HasBuiltins m) =>
a -> m (Maybe BoundedSize)
isSizeType CompareAs
a) m ()
cont (\ BoundedSize
_ -> m ()
fallback)

          dir :: CompareDirection
dir = Comparison -> CompareDirection
fromCmp Comparison
cmp
          rid :: CompareDirection
rid = CompareDirection -> CompareDirection
flipCmp CompareDirection
dir     -- The reverse direction.  Bad name, I know.
      case (Term
u, Term
v) of
        (MetaV MetaId
x Elims
us, MetaV MetaId
y Elims
vs)
          | MetaId
x forall a. Eq a => a -> a -> Bool
/= MetaId
y    -> m () -> m ()
unlessSubtyping forall a b. (a -> b) -> a -> b
$ m ()
solve1 m () -> m () -> m ()
`orelse` m ()
solve2 m () -> m () -> m ()
`orelse` m ()
fallback
          | Bool
otherwise -> m ()
fallback
          where
            (m ()
solve1, m ()
solve2) | MetaId
x forall a. Ord a => a -> a -> Bool
> MetaId
y     = (CompareDirection -> MetaId -> Elims -> Term -> m ()
assign CompareDirection
dir MetaId
x Elims
us Term
v, CompareDirection -> MetaId -> Elims -> Term -> m ()
assign CompareDirection
rid MetaId
y Elims
vs Term
u)
                             | Bool
otherwise = (CompareDirection -> MetaId -> Elims -> Term -> m ()
assign CompareDirection
rid MetaId
y Elims
vs Term
u, CompareDirection -> MetaId -> Elims -> Term -> m ()
assign CompareDirection
dir MetaId
x Elims
us Term
v)
        (MetaV MetaId
x Elims
us, Term
_) -> m () -> m ()
unlessSubtyping forall a b. (a -> b) -> a -> b
$ CompareDirection -> MetaId -> Elims -> Term -> m ()
assign CompareDirection
dir MetaId
x Elims
us Term
v m () -> m () -> m ()
`orelse` m ()
fallback
        (Term
_, MetaV MetaId
y Elims
vs) -> m () -> m ()
unlessSubtyping forall a b. (a -> b) -> a -> b
$ CompareDirection -> MetaId -> Elims -> Term -> m ()
assign CompareDirection
rid MetaId
y Elims
vs Term
u m () -> m () -> m ()
`orelse` m ()
fallback
        (Def QName
f Elims
es, Def QName
f' Elims
es') | QName
f forall a. Eq a => a -> a -> Bool
== QName
f' ->
          forall (m :: * -> *) a. Monad m => m Bool -> m a -> m a -> m a
ifNotM (PragmaOptions -> Bool
optFirstOrder forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> forall (m :: * -> *). HasOptions m => m PragmaOptions
pragmaOptions) m ()
fallback forall a b. (a -> b) -> a -> b
$ {- else -} m () -> m ()
unlessSubtyping forall a b. (a -> b) -> a -> b
$ do
          Definition
def <- forall (m :: * -> *). HasConstInfo m => QName -> m Definition
getConstInfo QName
f
          -- We do not shortcut projection-likes,
          -- Andreas, 2022-03-07, issue #5809:
          -- but irrelevant projections since they are applied to their parameters.
          if forall a. Maybe a -> Bool
isJust forall a b. (a -> b) -> a -> b
$ Definition -> Maybe Projection
isRelevantProjection_ Definition
def then m ()
fallback else do
          [Polarity]
pol <- forall (m :: * -> *).
HasConstInfo m =>
Comparison -> QName -> m [Polarity]
getPolarity' Comparison
cmp QName
f
          forall (m :: * -> *). MonadDebug m => ProfileOption -> m () -> m ()
whenProfile ProfileOption
Profile.Conversion forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *). MonadStatistics m => VerboseKey -> m ()
tick VerboseKey
"compare first-order shortcut"
          forall (m :: * -> *).
MonadConversion m =>
[Polarity] -> [IsForced] -> Type -> Term -> Elims -> Elims -> m ()
compareElims [Polarity]
pol [] (Definition -> Type
defType Definition
def) (QName -> Elims -> Term
Def QName
f []) Elims
es Elims
es' m () -> m () -> m ()
`orelse` m ()
fallback
        (Term, Term)
_               -> m ()
fallback
  where
    assign :: CompareDirection -> MetaId -> Elims -> Term -> m ()
    assign :: CompareDirection -> MetaId -> Elims -> Term -> m ()
assign CompareDirection
dir MetaId
x Elims
es Term
v = do
      -- Andreas, 2013-10-19 can only solve if no projections
      forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"tc.conv.term.shortcut" Int
20 forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
sep
        [ TCMT IO Doc
"attempting shortcut"
        , 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 (MetaId -> Elims -> Term
MetaV MetaId
x Elims
es) 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 :: * -> *). Monad m => m Bool -> m () -> m ()
whenM (forall a (m :: * -> *).
(IsInstantiatedMeta a, MonadFail m, ReadTCState m) =>
a -> m Bool
isInstantiatedMeta MetaId
x) (forall (m :: * -> *) a. MonadBlock m => Blocker -> m a
patternViolation Blocker
alwaysUnblock) -- Already instantiated, retry right away
      forall (m :: * -> *). MonadDebug m => ProfileOption -> m () -> m ()
whenProfile ProfileOption
Profile.Conversion forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *). MonadStatistics m => VerboseKey -> m ()
tick VerboseKey
"compare meta shortcut"
      forall (m :: * -> *).
MonadConversion m =>
CompareDirection
-> MetaId
-> Elims
-> Term
-> CompareAs
-> (Term -> Term -> m ())
-> m ()
assignE CompareDirection
dir MetaId
x Elims
es Term
v CompareAs
a forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *).
MonadConversion m =>
CompareDirection -> CompareAs -> Term -> Term -> m ()
compareAsDir CompareDirection
dir CompareAs
a
      forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"tc.conv.term.shortcut" Int
50 forall a b. (a -> b) -> a -> b
$
        TCMT IO Doc
"shortcut successful" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
$$ forall (m :: * -> *). Functor m => Int -> m Doc -> m Doc
nest Int
2 (TCMT IO Doc
"result:" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> (forall (m :: * -> *) a. (Applicative m, Pretty a) => a -> m Doc
pretty forall (m :: * -> *) a b. Monad m => (a -> m b) -> m a -> m b
=<< forall a (m :: * -> *). (Instantiate a, MonadReduce m) => a -> m a
instantiate (MetaId -> Elims -> Term
MetaV MetaId
x Elims
es)))
      forall (m :: * -> *). MonadDebug m => ProfileOption -> m () -> m ()
whenProfile ProfileOption
Profile.Conversion forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *). MonadStatistics m => VerboseKey -> m ()
tick VerboseKey
"compare meta shortcut successful"
    -- Should be ok with catchError_ but catchError is much safer since we don't
    -- rethrow errors.
    orelse :: m () -> m () -> m ()
    orelse :: m () -> m () -> m ()
orelse m ()
m m ()
h = forall e (m :: * -> *) a.
MonadError e m =>
m a -> (e -> m a) -> m a
catchError m ()
m (\TCErr
_ -> m ()
h)

-- | Try to assign meta.  If meta is projected, try to eta-expand
--   and run conversion check again.
assignE :: (MonadConversion m)
        => CompareDirection -> MetaId -> Elims -> Term -> CompareAs -> (Term -> Term -> m ()) -> m ()
assignE :: forall (m :: * -> *).
MonadConversion m =>
CompareDirection
-> MetaId
-> Elims
-> Term
-> CompareAs
-> (Term -> Term -> m ())
-> m ()
assignE CompareDirection
dir MetaId
x Elims
es Term
v CompareAs
a Term -> Term -> m ()
comp = do
  forall (m :: * -> *). MonadDebug m => ProfileOption -> m () -> m ()
whenProfile ProfileOption
Profile.Conversion forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *). MonadStatistics m => VerboseKey -> m ()
tick VerboseKey
"compare meta"
  case forall a. [Elim' a] -> Maybe [Arg a]
allApplyElims Elims
es of
    Just [Arg Term]
vs -> forall (m :: * -> *).
MonadMetaSolver m =>
CompareDirection
-> MetaId -> [Arg Term] -> Term -> CompareAs -> m ()
assignV CompareDirection
dir MetaId
x [Arg Term]
vs Term
v CompareAs
a
    Maybe [Arg Term]
Nothing -> do
      forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"tc.conv.assign" Int
30 forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
sep
        [ TCMT IO Doc
"assigning to projected meta "
        , forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM MetaId
x forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
sep (forall a b. (a -> b) -> [a] -> [b]
map forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Elims
es) forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall (m :: * -> *). Applicative m => VerboseKey -> m Doc
text (VerboseKey
":" forall a. [a] -> [a] -> [a]
++ forall a. Show a => a -> VerboseKey
show CompareDirection
dir) 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 :: * -> *).
MonadMetaSolver m =>
[MetaKind] -> MetaId -> m ()
etaExpandMeta [MetaKind
Records] MetaId
x
      Maybe Term
res <- forall (m :: * -> *).
(MonadFail m, ReadTCState m) =>
MetaId -> m (Maybe Term)
isInstantiatedMeta' MetaId
x
      case Maybe Term
res of
        Just Term
u  -> do
          forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"tc.conv.assign" Int
30 forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
sep
            [ TCMT IO Doc
"seems like eta expansion instantiated meta "
            , forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM MetaId
x forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall (m :: * -> *). Applicative m => VerboseKey -> m Doc
text  (VerboseKey
":" forall a. [a] -> [a] -> [a]
++ forall a. Show a => a -> VerboseKey
show CompareDirection
dir) forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Term
u
            ]
          let w :: Term
w = Term
u forall t. Apply t => t -> Elims -> t
`applyE` Elims
es
          Term -> Term -> m ()
comp Term
w Term
v
        Maybe Term
Nothing ->  do
          forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> VerboseKey -> m ()
reportSLn VerboseKey
"tc.conv.assign" Int
30 VerboseKey
"eta expansion did not instantiate meta"
          forall (m :: * -> *) a. MonadBlock m => Blocker -> m a
patternViolation forall a b. (a -> b) -> a -> b
$ MetaId -> Blocker
unblockOnMeta MetaId
x -- nothing happened, give up

compareAsDir :: MonadConversion m => CompareDirection -> CompareAs -> Term -> Term -> m ()
compareAsDir :: forall (m :: * -> *).
MonadConversion m =>
CompareDirection -> CompareAs -> Term -> Term -> m ()
compareAsDir CompareDirection
dir CompareAs
a = forall a c.
(Comparison -> a -> a -> c) -> CompareDirection -> a -> a -> c
dirToCmp (forall (m :: * -> *).
MonadConversion m =>
Comparison -> CompareAs -> Term -> Term -> m ()
`compareAs'` CompareAs
a) CompareDirection
dir

compareAs' :: forall m. MonadConversion m => Comparison -> CompareAs -> Term -> Term -> m ()
compareAs' :: forall (m :: * -> *).
MonadConversion m =>
Comparison -> CompareAs -> Term -> Term -> m ()
compareAs' Comparison
cmp CompareAs
tt Term
m Term
n = case CompareAs
tt of
  AsTermsOf Type
a -> forall (m :: * -> *).
MonadConversion m =>
Comparison -> Type -> Term -> Term -> m ()
compareTerm' Comparison
cmp Type
a Term
m Term
n
  CompareAs
AsSizes     -> forall (m :: * -> *).
MonadConversion m =>
Comparison -> Term -> Term -> m ()
compareSizes Comparison
cmp Term
m Term
n
  CompareAs
AsTypes     -> forall (m :: * -> *).
MonadConversion m =>
Comparison -> CompareAs -> Term -> Term -> m ()
compareAtom Comparison
cmp CompareAs
AsTypes Term
m Term
n

compareTerm' :: forall m. MonadConversion m => Comparison -> Type -> Term -> Term -> m ()
compareTerm' :: forall (m :: * -> *).
MonadConversion m =>
Comparison -> Type -> Term -> Term -> m ()
compareTerm' Comparison
cmp Type
a Term
m Term
n =
  forall (m :: * -> *) a.
MonadDebug m =>
VerboseKey -> Int -> VerboseKey -> m a -> m a
verboseBracket VerboseKey
"tc.conv.term" Int
20 VerboseKey
"compareTerm" forall a b. (a -> b) -> a -> b
$ do
  (Blocker
ba, Type
a') <- forall a (m :: * -> *).
(Reduce a, IsMeta a, MonadReduce m) =>
a -> m (Blocker, a)
reduceWithBlocker Type
a
  (forall (m :: * -> *).
MonadConstraint m =>
Constraint -> m () -> m ()
catchConstraint (Comparison -> CompareAs -> Term -> Term -> Constraint
ValueCmp Comparison
cmp (Type -> CompareAs
AsTermsOf Type
a') Term
m Term
n) :: m () -> m ()) forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *) a. MonadError TCErr m => Blocker -> m a -> m a
blockOnError Blocker
ba forall a b. (a -> b) -> a -> b
$ do
    forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"tc.conv.term" Int
30 forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
fsep
      [ TCMT IO Doc
"compareTerm", forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Term
m, forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Comparison
cmp, forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Term
n, TCMT IO Doc
":", forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Type
a' ]
    Bool
propIrr  <- forall (m :: * -> *). HasOptions m => m Bool
isPropEnabled
    Bool
isSize   <- forall a. Maybe a -> Bool
isJust forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> forall a (m :: * -> *).
(IsSizeType a, HasOptions m, HasBuiltins m) =>
a -> m (Maybe BoundedSize)
isSizeType Type
a'
    (Blocker
bs, Sort
s)  <- forall a (m :: * -> *).
(Reduce a, IsMeta a, MonadReduce m) =>
a -> m (Blocker, a)
reduceWithBlocker forall a b. (a -> b) -> a -> b
$ forall a. LensSort a => a -> Sort
getSort Type
a'
    Maybe Term
mlvl     <- forall (m :: * -> *). HasBuiltins m => VerboseKey -> m (Maybe Term)
getBuiltin' VerboseKey
builtinLevel
    forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"tc.conv.term" Int
40 forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
fsep
      [ TCMT IO Doc
"compareTerm", forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Term
m, forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Comparison
cmp, forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Term
n, TCMT IO Doc
":", forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Type
a'
      , TCMT IO Doc
"at sort", forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Sort
s]
    forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"tc.conv.level" Int
60 forall a b. (a -> b) -> a -> b
$ 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
      [ TCMT IO Doc
"a'   =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall (m :: * -> *) a. (Applicative m, Pretty a) => a -> m Doc
pretty Type
a'
      , TCMT IO Doc
"mlvl =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall (m :: * -> *) a. (Applicative m, Pretty a) => a -> m Doc
pretty Maybe Term
mlvl
      , forall (m :: * -> *). Applicative m => VerboseKey -> m Doc
text forall a b. (a -> b) -> a -> b
$ VerboseKey
"(Just (unEl a') == mlvl) = " forall a. [a] -> [a] -> [a]
++ forall a. Show a => a -> VerboseKey
show (forall a. a -> Maybe a
Just (forall t a. Type'' t a -> a
unEl Type
a') forall a. Eq a => a -> a -> Bool
== Maybe Term
mlvl)
      ]
    forall (m :: * -> *) a. MonadError TCErr m => Blocker -> m a -> m a
blockOnError Blocker
bs forall a b. (a -> b) -> a -> b
$ case Sort
s of
      Prop{} | Bool
propIrr -> forall (m :: * -> *).
MonadConversion m =>
Type -> Term -> Term -> m ()
compareIrrelevant Type
a' Term
m Term
n
      Sort
_    | Bool
isSize   -> forall (m :: * -> *).
MonadConversion m =>
Comparison -> Term -> Term -> m ()
compareSizes Comparison
cmp Term
m Term
n
      Sort
_               -> case forall t a. Type'' t a -> a
unEl Type
a' of
        Term
a | forall a. a -> Maybe a
Just Term
a forall a. Eq a => a -> a -> Bool
== Maybe Term
mlvl -> do
          Level
a <- forall (m :: * -> *). PureTCM m => Term -> m Level
levelView Term
m
          Level
b <- forall (m :: * -> *). PureTCM m => Term -> m Level
levelView Term
n
          forall (m :: * -> *). MonadConversion m => Level -> Level -> m ()
equalLevel Level
a Level
b
        a :: Term
a@Pi{}    -> MonadConversion m => Sort -> Term -> Term -> Term -> m ()
equalFun Sort
s Term
a Term
m Term
n
        Lam ArgInfo
_ Abs Term
_   -> do
          forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"tc.conv.term.sort" Int
10 forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
fsep
            [ TCMT IO Doc
"compareTerm", forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Term
m, forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Comparison
cmp, forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Term
n, TCMT IO Doc
":", forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Type
a'
            , TCMT IO Doc
"at sort", forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Sort
s
            ]
          forall a. HasCallStack => a
__IMPOSSIBLE__
        Def QName
r Elims
es  -> do
          Bool
isrec <- forall (m :: * -> *). HasConstInfo m => QName -> m Bool
isEtaRecord QName
r
          if Bool
isrec
            then do
              forall (m :: * -> *). MonadDebug m => ProfileOption -> m () -> m ()
whenProfile ProfileOption
Profile.Conversion forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *). MonadStatistics m => VerboseKey -> m ()
tick VerboseKey
"compare at eta record"
              Signature
sig <- forall (m :: * -> *). ReadTCState m => m Signature
getSignature
              let ps :: [Arg Term]
ps = forall a. a -> Maybe a -> a
fromMaybe forall a. HasCallStack => a
__IMPOSSIBLE__ forall a b. (a -> b) -> a -> b
$ forall a. [Elim' a] -> Maybe [Arg a]
allApplyElims Elims
es
              -- Andreas, 2010-10-11: allowing neutrals to be blocked things does not seem
              -- to change Agda's behavior
              --    isNeutral Blocked{}          = False
                  isNeutral :: Blocked' t Term -> m Bool
isNeutral (NotBlocked NotBlocked' t
_ Con{}) = forall (m :: * -> *) a. Monad m => a -> m a
return Bool
False
              -- Andreas, 2013-09-18 / 2015-06-29: a Def by copatterns is
              -- not neutral if it is blocked (there can be missing projections
              -- to trigger a reduction.
                  isNeutral (NotBlocked NotBlocked' t
r (Def QName
q Elims
_)) = do    -- Andreas, 2014-12-06 optimize this using r !!
                    Bool -> Bool
not forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> forall (m :: * -> *). HasConstInfo m => QName -> m Bool
usesCopatterns QName
q -- a def by copattern can reduce if projected
                  isNeutral Blocked' t Term
_                   = forall (m :: * -> *) a. Monad m => a -> m a
return Bool
True
                  isMeta :: Blocked' t Term -> Bool
isMeta Blocked' t Term
b = case forall t a. Blocked' t a -> a
ignoreBlocking Blocked' t Term
b of
                               MetaV{} -> Bool
True
                               Term
_       -> Bool
False

              forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"tc.conv.term" Int
30 forall a b. (a -> b) -> a -> b
$ forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Type
a forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> TCMT IO Doc
"is eta record type"
              Blocked Term
m <- forall a (m :: * -> *).
(Reduce a, MonadReduce m) =>
a -> m (Blocked a)
reduceB Term
m
              Bool
mNeutral <- forall {m :: * -> *} {t}.
HasConstInfo m =>
Blocked' t Term -> m Bool
isNeutral Blocked Term
m
              Blocked Term
n <- forall a (m :: * -> *).
(Reduce a, MonadReduce m) =>
a -> m (Blocked a)
reduceB Term
n
              Bool
nNeutral <- forall {m :: * -> *} {t}.
HasConstInfo m =>
Blocked' t Term -> m Bool
isNeutral Blocked Term
n
              if | forall {t}. Blocked' t Term -> Bool
isMeta Blocked Term
m Bool -> Bool -> Bool
|| forall {t}. Blocked' t Term -> Bool
isMeta Blocked Term
n -> do
                     forall (m :: * -> *). MonadDebug m => ProfileOption -> m () -> m ()
whenProfile ProfileOption
Profile.Conversion forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *). MonadStatistics m => VerboseKey -> m ()
tick VerboseKey
"compare at eta-record: meta"
                     forall (m :: * -> *).
MonadConversion m =>
Comparison -> CompareAs -> Term -> Term -> m ()
compareAtom Comparison
cmp (Type -> CompareAs
AsTermsOf Type
a') (forall t a. Blocked' t a -> a
ignoreBlocking Blocked Term
m) (forall t a. Blocked' t a -> a
ignoreBlocking Blocked Term
n)
                 | Bool
mNeutral Bool -> Bool -> Bool
&& Bool
nNeutral -> do
                     forall (m :: * -> *). MonadDebug m => ProfileOption -> m () -> m ()
whenProfile ProfileOption
Profile.Conversion forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *). MonadStatistics m => VerboseKey -> m ()
tick VerboseKey
"compare at eta-record: both neutral"
                     -- Andreas 2011-03-23: (fixing issue 396)
                     -- if we are dealing with a singleton record,
                     -- we can succeed immediately
                     let profUnitEta :: m ()
profUnitEta = forall (m :: * -> *). MonadDebug m => ProfileOption -> m () -> m ()
whenProfile ProfileOption
Profile.Conversion forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *). MonadStatistics m => VerboseKey -> m ()
tick VerboseKey
"compare at eta-record: both neutral at unit"
                     forall (m :: * -> *) a. Monad m => m Bool -> m a -> m a -> m a
ifM (forall (m :: * -> *).
(PureTCM m, MonadBlock m) =>
QName -> [Arg Term] -> m Bool
isSingletonRecordModuloRelevance QName
r [Arg Term]
ps) (m ()
profUnitEta) forall a b. (a -> b) -> a -> b
$ do
                       -- do not eta-expand if comparing two neutrals
                       forall (m :: * -> *).
MonadConversion m =>
Comparison -> CompareAs -> Term -> Term -> m ()
compareAtom Comparison
cmp (Type -> CompareAs
AsTermsOf Type
a') (forall t a. Blocked' t a -> a
ignoreBlocking Blocked Term
m) (forall t a. Blocked' t a -> a
ignoreBlocking Blocked Term
n)
                 | Bool
otherwise -> do
                     forall (m :: * -> *). MonadDebug m => ProfileOption -> m () -> m ()
whenProfile ProfileOption
Profile.Conversion forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *). MonadStatistics m => VerboseKey -> m ()
tick VerboseKey
"compare at eta-record: eta-expanding"
                     (Telescope
tel, [Arg Term]
m') <- forall (m :: * -> *).
(HasConstInfo m, MonadDebug m, ReadTCState m) =>
QName -> [Arg Term] -> Term -> m (Telescope, [Arg Term])
etaExpandRecord QName
r [Arg Term]
ps forall a b. (a -> b) -> a -> b
$ forall t a. Blocked' t a -> a
ignoreBlocking Blocked Term
m
                     (Telescope
_  , [Arg Term]
n') <- forall (m :: * -> *).
(HasConstInfo m, MonadDebug m, ReadTCState m) =>
QName -> [Arg Term] -> Term -> m (Telescope, [Arg Term])
etaExpandRecord QName
r [Arg Term]
ps forall a b. (a -> b) -> a -> b
$ forall t a. Blocked' t a -> a
ignoreBlocking Blocked Term
n
                     -- No subtyping on record terms
                     ConHead
c <- forall (m :: * -> *).
(HasConstInfo m, ReadTCState m, MonadError TCErr m) =>
QName -> m ConHead
getRecordConstructor QName
r
                     -- Record constructors are covariant (see test/succeed/CovariantConstructors).
                     forall (m :: * -> *).
MonadConversion m =>
[Polarity]
-> [IsForced] -> Type -> Term -> [Arg Term] -> [Arg Term] -> m ()
compareArgs (forall a. a -> [a]
repeat forall a b. (a -> b) -> a -> b
$ Comparison -> Polarity
polFromCmp Comparison
cmp) [] (Telescope -> Type -> Type
telePi_ Telescope
tel HasCallStack => Type
__DUMMY_TYPE__) (ConHead -> ConInfo -> Elims -> Term
Con ConHead
c ConInfo
ConOSystem []) [Arg Term]
m' [Arg Term]
n'

            else (do PathView
pathview <- forall (m :: * -> *). HasBuiltins m => Type -> m PathView
pathView Type
a'
                     MonadConversion m => PathView -> Type -> Term -> Term -> m ()
equalPath PathView
pathview Type
a' Term
m Term
n)
        Term
_ -> forall (m :: * -> *).
MonadConversion m =>
Comparison -> CompareAs -> Term -> Term -> m ()
compareAtom Comparison
cmp (Type -> CompareAs
AsTermsOf Type
a') Term
m Term
n
  where
    -- equality at function type (accounts for eta)
    equalFun :: (MonadConversion m) => Sort -> Term -> Term -> Term -> m ()
    equalFun :: MonadConversion m => Sort -> Term -> Term -> Term -> m ()
equalFun Sort
s a :: Term
a@(Pi Dom Type
dom Abs Type
b) Term
m Term
n | forall t e. Dom' t e -> Bool
domIsFinite Dom Type
dom = do
       Maybe QName
mp <- forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
fmap Term -> QName
getPrimName forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> forall (m :: * -> *). HasBuiltins m => VerboseKey -> m (Maybe Term)
getBuiltin' VerboseKey
builtinIsOne
       let asFn :: Type
asFn = forall t a. Sort' t -> a -> Type'' t a
El Sort
s (Dom Type -> Abs Type -> Term
Pi (Dom Type
dom { domIsFinite :: Bool
domIsFinite = Bool
False }) Abs Type
b)
       case 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
dom of
          Def QName
q [Apply Arg Term
phi]
              | forall a. a -> Maybe a
Just QName
q forall a. Eq a => a -> a -> Bool
== Maybe QName
mp -> forall (m :: * -> *).
MonadConversion m =>
Comparison -> Term -> Type -> Term -> Term -> m ()
compareTermOnFace Comparison
cmp (forall e. Arg e -> e
unArg Arg Term
phi) Type
asFn Term
m Term
n
          Term
_                  -> MonadConversion m => Sort -> Term -> Term -> Term -> m ()
equalFun Sort
s (forall t a. Type'' t a -> a
unEl Type
asFn) Term
m Term
n

    equalFun Sort
_ (Pi dom :: Dom Type
dom@Dom{domInfo :: forall t e. Dom' t e -> ArgInfo
domInfo = ArgInfo
info} Abs Type
b) Term
m Term
n = do
        forall (m :: * -> *). MonadDebug m => ProfileOption -> m () -> m ()
whenProfile ProfileOption
Profile.Conversion forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *). MonadStatistics m => VerboseKey -> m ()
tick VerboseKey
"compare at function type"
        let name :: VerboseKey
name = [Suggestion] -> VerboseKey
suggests [ forall a. Suggest a => a -> Suggestion
Suggestion Abs Type
b , forall a. Suggest a => a -> Suggestion
Suggestion Term
m , forall a. Suggest a => a -> Suggestion
Suggestion Term
n ]
        forall b (m :: * -> *) a.
(AddContext b, MonadAddContext m) =>
b -> m a -> m a
addContext (VerboseKey
name, Dom Type
dom) forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *).
MonadConversion m =>
Comparison -> Type -> Term -> Term -> m ()
compareTerm Comparison
cmp (forall a. Subst a => Abs a -> a
absBody Abs Type
b) Term
m' Term
n'
      where
        (Term
m',Term
n') = forall a. Subst a => Int -> a -> a
raise Int
1 (Term
m,Term
n) forall t. Apply t => t -> [Arg Term] -> t
`apply` [forall e. ArgInfo -> e -> Arg e
Arg ArgInfo
info forall a b. (a -> b) -> a -> b
$ Int -> Term
var Int
0]

    equalFun Sort
_ Term
_ Term
_ Term
_ = forall a. HasCallStack => a
__IMPOSSIBLE__

    equalPath :: (MonadConversion m) => PathView -> Type -> Term -> Term -> m ()
    equalPath :: MonadConversion m => PathView -> Type -> Term -> Term -> m ()
equalPath (PathType Sort
s QName
_ Arg Term
l Arg Term
a Arg Term
x Arg Term
y) Type
_ Term
m Term
n = do
        forall (m :: * -> *). MonadDebug m => ProfileOption -> m () -> m ()
whenProfile ProfileOption
Profile.Conversion forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *). MonadStatistics m => VerboseKey -> m ()
tick VerboseKey
"compare at path type"
        let name :: VerboseKey
name = VerboseKey
"i" :: String
        Type
interval <- forall (m :: * -> *). Functor m => m Term -> m Type
el forall (m :: * -> *).
(HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) =>
m Term
primInterval
        let (Term
m',Term
n') = forall a. Subst a => Int -> a -> a
raise Int
1 (Term
m, Term
n) forall t. Apply t => t -> Elims -> t
`applyE` [forall a. a -> a -> a -> Elim' a
IApply (forall a. Subst a => Int -> a -> a
raise Int
1 forall a b. (a -> b) -> a -> b
$ forall e. Arg e -> e
unArg Arg Term
x) (forall a. Subst a => Int -> a -> a
raise Int
1 forall a b. (a -> b) -> a -> b
$ forall e. Arg e -> e
unArg Arg Term
y) (Int -> Term
var Int
0)]
        forall b (m :: * -> *) a.
(AddContext b, MonadAddContext m) =>
b -> m a -> m a
addContext (VerboseKey
name, forall a. a -> Dom a
defaultDom Type
interval) forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *).
MonadConversion m =>
Comparison -> Type -> Term -> Term -> m ()
compareTerm Comparison
cmp (forall t a. Sort' t -> a -> Type'' t a
El (forall a. Subst a => Int -> a -> a
raise Int
1 Sort
s) forall a b. (a -> b) -> a -> b
$ forall a. Subst a => Int -> a -> a
raise Int
1 (forall e. Arg e -> e
unArg Arg Term
a) forall t. Apply t => t -> [Arg Term] -> t
`apply` [forall e. e -> Arg e
argN forall a b. (a -> b) -> a -> b
$ Int -> Term
var Int
0]) Term
m' Term
n'
    equalPath OType{} Type
a' Term
m Term
n = Type -> Term -> Term -> m ()
cmpDef Type
a' Term
m Term
n

    cmpDef :: Type -> Term -> Term -> m ()
cmpDef a' :: Type
a'@(El Sort
s Term
ty) Term
m Term
n = do
       Maybe QName
mI     <- forall (m :: * -> *).
HasBuiltins m =>
VerboseKey -> m (Maybe QName)
getBuiltinName'   VerboseKey
builtinInterval
       Maybe QName
mIsOne <- forall (m :: * -> *).
HasBuiltins m =>
VerboseKey -> m (Maybe QName)
getBuiltinName'   VerboseKey
builtinIsOne
       Maybe QName
mGlue  <- forall (m :: * -> *).
HasBuiltins m =>
VerboseKey -> m (Maybe QName)
getPrimitiveName' VerboseKey
builtinGlue
       Maybe QName
mHComp <- forall (m :: * -> *).
HasBuiltins m =>
VerboseKey -> m (Maybe QName)
getPrimitiveName' VerboseKey
builtinHComp
       Maybe QName
mSub   <- forall (m :: * -> *).
HasBuiltins m =>
VerboseKey -> m (Maybe QName)
getBuiltinName' VerboseKey
builtinSub
       Maybe Term
mUnglueU <- forall (m :: * -> *). HasBuiltins m => VerboseKey -> m (Maybe Term)
getPrimitiveTerm' VerboseKey
builtin_unglueU
       Maybe Term
mSubIn   <- forall (m :: * -> *). HasBuiltins m => VerboseKey -> m (Maybe Term)
getPrimitiveTerm' VerboseKey
builtinSubIn
       case Term
ty of
         Def QName
q Elims
es | forall a. a -> Maybe a
Just QName
q forall a. Eq a => a -> a -> Bool
== Maybe QName
mIsOne -> forall (m :: * -> *) a. Monad m => a -> m a
return ()
         Def QName
q Elims
es | forall a. a -> Maybe a
Just QName
q forall a. Eq a => a -> a -> Bool
== Maybe QName
mGlue, Just args :: [Arg Term]
args@(Arg Term
l:Arg Term
_:Arg Term
a:Arg Term
phi:[Arg Term]
_) <- forall a. [Elim' a] -> Maybe [Arg a]
allApplyElims Elims
es -> do
              Type
aty <- forall (m :: * -> *). Applicative m => m Term -> m Term -> m Type
el' (forall (f :: * -> *) a. Applicative f => a -> f a
pure forall a b. (a -> b) -> a -> b
$ forall e. Arg e -> e
unArg Arg Term
l) (forall (f :: * -> *) a. Applicative f => a -> f a
pure forall a b. (a -> b) -> a -> b
$ forall e. Arg e -> e
unArg Arg Term
a)
              Term
unglue <- forall (m :: * -> *).
(HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) =>
m Term
prim_unglue
              let mkUnglue :: Term -> Term
mkUnglue Term
m = forall t. Apply t => t -> [Arg Term] -> t
apply Term
unglue forall a b. (a -> b) -> a -> b
$ forall a b. (a -> b) -> [a] -> [b]
map (forall a. LensHiding a => Hiding -> a -> a
setHiding Hiding
Hidden) [Arg Term]
args forall a. [a] -> [a] -> [a]
++ [forall e. e -> Arg e
argN Term
m]
              forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"conv.glue" Int
20 forall a b. (a -> b) -> a -> b
$ forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM (Type
aty,Term -> Term
mkUnglue Term
m,Term -> Term
mkUnglue Term
n)

              -- When φ is an interval expression which can be
              -- decomposed into substitutions σ, then we also compare
              -- the terms m[σ] = n[σ] at the type (Glue a φ _)[σ]. This
              -- is because, under decomposing φ, the Glue type might
              -- reduce.
              [(IntMap BoolSet, [Term])]
phi' <- forall (m :: * -> *).
HasBuiltins m =>
Term -> m [(IntMap BoolSet, [Term])]
decomposeInterval' (forall e. Arg e -> e
unArg Arg Term
phi)
              -- However if φ is *not* decomposable (e.g. because it is
              -- a function application φ i, see Issue #5955), then we
              -- do not recur, otherwise we'd just end up right back
              -- here.
              forall (f :: * -> *). Applicative f => Bool -> f () -> f ()
unless (forall a. IntMap a -> Bool
IntMap.null (forall (t :: * -> *) m a.
(Foldable t, Monoid m) =>
(a -> m) -> t a -> m
foldMap forall a b. (a, b) -> a
fst [(IntMap BoolSet, [Term])]
phi')) forall a b. (a -> b) -> a -> b
$
                forall (m :: * -> *).
MonadConversion m =>
Comparison -> Term -> Type -> Term -> Term -> m ()
compareTermOnFace Comparison
cmp (forall e. Arg e -> e
unArg Arg Term
phi) Type
a' Term
m Term
n

              -- And in the general case, we compare the glued things by
              -- "eta": m and n are the same if they unglue to the same
              -- thing.
              forall (m :: * -> *).
MonadConversion m =>
Comparison -> Type -> Term -> Term -> m ()
compareTerm Comparison
cmp Type
aty (Term -> Term
mkUnglue Term
m) (Term -> Term
mkUnglue Term
n)
         Def QName
q Elims
es | forall a. a -> Maybe a
Just QName
q forall a. Eq a => a -> a -> Bool
== Maybe QName
mHComp, Just (Arg Term
sl:Arg Term
s:args :: [Arg Term]
args@[Arg Term
phi,Arg Term
u,Arg Term
u0]) <- forall a. [Elim' a] -> Maybe [Arg a]
allApplyElims Elims
es
                  , Sort (Type Level
lvl) <- forall e. Arg e -> e
unArg Arg Term
s
                  , Just Term
unglueU <- Maybe Term
mUnglueU, Just Term
subIn <- Maybe Term
mSubIn
                  -> do
              let l :: Term
l = Level -> Term
Level Level
lvl
              Type
ty <- forall (m :: * -> *). Applicative m => m Term -> m Term -> m Type
el' (forall (f :: * -> *) a. Applicative f => a -> f a
pure forall a b. (a -> b) -> a -> b
$ Term
l) (forall (f :: * -> *) a. Applicative f => a -> f a
pure forall a b. (a -> b) -> a -> b
$ forall e. Arg e -> e
unArg Arg Term
u0)
              let bA :: Term
bA = Term
subIn forall t. Apply t => t -> [Arg Term] -> t
`apply` [Arg Term
sl,Arg Term
s,Arg Term
phi,Arg Term
u0]
              let mkUnglue :: Term -> Term
mkUnglue Term
m = forall t. Apply t => t -> [Arg Term] -> t
apply Term
unglueU forall a b. (a -> b) -> a -> b
$ [forall e. e -> Arg e
argH Term
l] forall a. [a] -> [a] -> [a]
++ forall a b. (a -> b) -> [a] -> [b]
map (forall a. LensHiding a => Hiding -> a -> a
setHiding Hiding
Hidden) [Arg Term
phi,Arg Term
u]  forall a. [a] -> [a] -> [a]
++ [forall e. e -> Arg e
argH Term
bA,forall e. e -> Arg e
argN Term
m]
              forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"conv.hcompU" Int
20 forall a b. (a -> b) -> a -> b
$ forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM (Type
ty,Term -> Term
mkUnglue Term
m,Term -> Term
mkUnglue Term
n)
              forall (m :: * -> *).
MonadConversion m =>
Comparison -> Term -> Type -> Term -> Term -> m ()
compareTermOnFace Comparison
cmp (forall e. Arg e -> e
unArg Arg Term
phi) Type
ty Term
m Term
n
              forall (m :: * -> *).
MonadConversion m =>
Comparison -> Type -> Term -> Term -> m ()
compareTerm Comparison
cmp Type
ty (Term -> Term
mkUnglue Term
m) (Term -> Term
mkUnglue Term
n)
         Def QName
q Elims
es | forall a. a -> Maybe a
Just QName
q forall a. Eq a => a -> a -> Bool
== Maybe QName
mSub, Just args :: [Arg Term]
args@(Arg Term
l:Arg Term
a:[Arg Term]
_) <- forall a. [Elim' a] -> Maybe [Arg a]
allApplyElims Elims
es -> do
              Type
ty <- forall (m :: * -> *). Applicative m => m Term -> m Term -> m Type
el' (forall (f :: * -> *) a. Applicative f => a -> f a
pure forall a b. (a -> b) -> a -> b
$ forall e. Arg e -> e
unArg Arg Term
l) (forall (f :: * -> *) a. Applicative f => a -> f a
pure forall a b. (a -> b) -> a -> b
$ forall e. Arg e -> e
unArg Arg Term
a)
              Term
out <- forall (m :: * -> *).
(HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) =>
m Term
primSubOut
              let mkOut :: Term -> Term
mkOut Term
m = forall t. Apply t => t -> [Arg Term] -> t
apply Term
out forall a b. (a -> b) -> a -> b
$ forall a b. (a -> b) -> [a] -> [b]
map (forall a. LensHiding a => Hiding -> a -> a
setHiding Hiding
Hidden) [Arg Term]
args forall a. [a] -> [a] -> [a]
++ [forall e. e -> Arg e
argN Term
m]
              forall (m :: * -> *).
MonadConversion m =>
Comparison -> Type -> Term -> Term -> m ()
compareTerm Comparison
cmp Type
ty (Term -> Term
mkOut Term
m) (Term -> Term
mkOut Term
n)
         Def QName
q [] | forall a. a -> Maybe a
Just QName
q forall a. Eq a => a -> a -> Bool
== Maybe QName
mI -> forall (m :: * -> *).
MonadConversion m =>
Comparison -> Type -> Term -> Term -> m ()
compareInterval Comparison
cmp Type
a' Term
m Term
n
         Term
_ -> forall (m :: * -> *).
MonadConversion m =>
Comparison -> CompareAs -> Term -> Term -> m ()
compareAtom Comparison
cmp (Type -> CompareAs
AsTermsOf Type
a') Term
m Term
n

compareAtomDir :: MonadConversion m => CompareDirection -> CompareAs -> Term -> Term -> m ()
compareAtomDir :: forall (m :: * -> *).
MonadConversion m =>
CompareDirection -> CompareAs -> Term -> Term -> m ()
compareAtomDir CompareDirection
dir CompareAs
a = forall a c.
(Comparison -> a -> a -> c) -> CompareDirection -> a -> a -> c
dirToCmp (forall (m :: * -> *).
MonadConversion m =>
Comparison -> CompareAs -> Term -> Term -> m ()
`compareAtom` CompareAs
a) CompareDirection
dir

-- | Compute the head type of an elimination. For projection-like functions
--   this requires inferring the type of the principal argument.
computeElimHeadType :: MonadConversion m => QName -> Elims -> Elims -> m Type
computeElimHeadType :: forall (m :: * -> *).
MonadConversion m =>
QName -> Elims -> Elims -> m Type
computeElimHeadType QName
f Elims
es Elims
es' = do
  Definition
def <- forall (m :: * -> *). HasConstInfo m => QName -> m Definition
getConstInfo QName
f
  -- To compute the type @a@ of a projection-like @f@,
  -- we have to infer the type of its first argument.
  if Definition -> Int
projectionArgs Definition
def forall a. Ord a => a -> a -> Bool
<= Int
0 then forall (m :: * -> *) a. Monad m => a -> m a
return forall a b. (a -> b) -> a -> b
$ Definition -> Type
defType Definition
def else do
    -- Find a first argument to @f@.
    let arg :: Arg Term
arg = case (Elims
es, Elims
es') of
              (Apply Arg Term
arg : Elims
_, Elims
_) -> Arg Term
arg
              (Elims
_, Apply Arg Term
arg : Elims
_) -> Arg Term
arg
              (Elims, Elims)
_ -> forall a. HasCallStack => a
__IMPOSSIBLE__
    -- Infer its type.
    forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"tc.conv.infer" Int
30 forall a b. (a -> b) -> a -> b
$
      TCMT IO Doc
"inferring type of internal arg: " forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Arg Term
arg
    Type
targ <- forall (m :: * -> *). MonadCheckInternal m => Term -> m Type
infer forall a b. (a -> b) -> a -> b
$ forall e. Arg e -> e
unArg Arg Term
arg
    forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"tc.conv.infer" Int
30 forall a b. (a -> b) -> a -> b
$
      TCMT IO Doc
"inferred type: " forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Type
targ
    -- getDefType wants the argument type reduced.
    -- Andreas, 2016-02-09, Issue 1825: The type of arg might be
    -- a meta-variable, e.g. in interactive development.
    -- In this case, we postpone.
    Type
targ <- forall (m :: * -> *) t.
(MonadReduce m, MonadBlock m, IsMeta t, Reduce t) =>
t -> m t
abortIfBlocked Type
targ
    forall (m :: * -> *) a. Monad m => m a -> m (Maybe a) -> m a
fromMaybeM forall a. HasCallStack => a
__IMPOSSIBLE__ forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *). PureTCM m => QName -> Type -> m (Maybe Type)
getDefType QName
f Type
targ


-- | Syntax directed equality on atomic values
--
compareAtom :: forall m. MonadConversion m => Comparison -> CompareAs -> Term -> Term -> m ()
compareAtom :: forall (m :: * -> *).
MonadConversion m =>
Comparison -> CompareAs -> Term -> Term -> m ()
compareAtom Comparison
cmp CompareAs
t Term
m Term
n =
  forall (m :: * -> *) a.
MonadDebug m =>
VerboseKey -> Int -> VerboseKey -> m a -> m a
verboseBracket VerboseKey
"tc.conv.atom" Int
20 VerboseKey
"compareAtom" forall a b. (a -> b) -> a -> b
$
  -- if a PatternErr is thrown, rebuild constraint!
  (forall (m :: * -> *).
MonadConstraint m =>
Constraint -> m () -> m ()
catchConstraint (Comparison -> CompareAs -> Term -> Term -> Constraint
ValueCmp Comparison
cmp CompareAs
t Term
m Term
n) :: m () -> m ()) forall a b. (a -> b) -> a -> b
$ do
    forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> VerboseKey -> m ()
reportSLn VerboseKey
"tc.conv.atom.size" Int
50 forall a b. (a -> b) -> a -> b
$ VerboseKey
"compareAtom term size:  " forall a. [a] -> [a] -> [a]
++ forall a. Show a => a -> VerboseKey
show (forall a. TermSize a => a -> Int
termSize Term
m, forall a. TermSize a => a -> Int
termSize Term
n)
    forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"tc.conv.atom" Int
50 forall a b. (a -> b) -> a -> b
$
      TCMT IO Doc
"compareAtom" 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 (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Term
m forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Comparison
cmp
                             , forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Term
n
                             , forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM CompareAs
t
                             ]
    forall (m :: * -> *). MonadDebug m => ProfileOption -> m () -> m ()
whenProfile ProfileOption
Profile.Conversion forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *). MonadStatistics m => VerboseKey -> m ()
tick VerboseKey
"compare by reduction"
    -- Are we currently defining mutual functions? Which?
    Set QName
currentMutuals <- forall b a. b -> (a -> b) -> Maybe a -> b
maybe (forall (f :: * -> *) a. Applicative f => a -> f a
pure forall a. Set a
Set.empty) (MutualBlock -> Set QName
mutualNames forall (m :: * -> *) b c a.
Functor m =>
(b -> c) -> (a -> m b) -> a -> m c
<.> forall (tcm :: * -> *).
ReadTCState tcm =>
MutualId -> tcm MutualBlock
lookupMutualBlock) forall (m :: * -> *) a b. Monad m => (a -> m b) -> m a -> m b
=<< forall (m :: * -> *) a. MonadTCEnv m => (TCEnv -> a) -> m a
asksTC TCEnv -> Maybe MutualId
envMutualBlock

    -- Andreas: what happens if I cut out the eta expansion here?
    -- Answer: Triggers issue 245, does not resolve 348
    (Blocked Term
mb',Blocked Term
nb') <- do
      Blocked Term
mb' <- forall (m :: * -> *) t.
(MonadReduce m, MonadMetaSolver m, IsMeta t, Reduce t) =>
Blocked t -> m (Blocked t)
etaExpandBlocked forall (m :: * -> *) a b. Monad m => (a -> m b) -> m a -> m b
=<< forall a (m :: * -> *).
(Reduce a, MonadReduce m) =>
a -> m (Blocked a)
reduceB Term
m
      Blocked Term
nb' <- forall (m :: * -> *) t.
(MonadReduce m, MonadMetaSolver m, IsMeta t, Reduce t) =>
Blocked t -> m (Blocked t)
etaExpandBlocked forall (m :: * -> *) a b. Monad m => (a -> m b) -> m a -> m b
=<< forall a (m :: * -> *).
(Reduce a, MonadReduce m) =>
a -> m (Blocked a)
reduceB Term
n
      forall (m :: * -> *) a. Monad m => a -> m a
return (Blocked Term
mb', Blocked Term
nb')
    let blocker :: Blocker
blocker = Blocker -> Blocker -> Blocker
unblockOnEither (forall t a. Blocked' t a -> Blocker
getBlocker Blocked Term
mb') (forall t a. Blocked' t a -> Blocker
getBlocker Blocked Term
nb')
    forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> VerboseKey -> m ()
reportSLn VerboseKey
"tc.conv.atom.size" Int
50 forall a b. (a -> b) -> a -> b
$ VerboseKey
"term size after reduce: " forall a. [a] -> [a] -> [a]
++ forall a. Show a => a -> VerboseKey
show (forall a. TermSize a => a -> Int
termSize forall a b. (a -> b) -> a -> b
$ forall t a. Blocked' t a -> a
ignoreBlocking Blocked Term
mb', forall a. TermSize a => a -> Int
termSize forall a b. (a -> b) -> a -> b
$ forall t a. Blocked' t a -> a
ignoreBlocking Blocked Term
nb')

    -- constructorForm changes literal to constructors
    -- only needed if the other side is not a literal
    (Blocked Term
mb'', Blocked Term
nb'') <- case (forall t a. Blocked' t a -> a
ignoreBlocking Blocked Term
mb', forall t a. Blocked' t a -> a
ignoreBlocking Blocked Term
nb') of
      (Lit Literal
_, Lit Literal
_) -> forall (m :: * -> *) a. Monad m => a -> m a
return (Blocked Term
mb', Blocked Term
nb')
      (Term, Term)
_ -> (,) forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> forall (t :: * -> *) (f :: * -> *) a b.
(Traversable t, Applicative f) =>
(a -> f b) -> t a -> f (t b)
traverse forall (m :: * -> *). HasBuiltins m => Term -> m Term
constructorForm Blocked Term
mb'
               forall (f :: * -> *) a b. Applicative f => f (a -> b) -> f a -> f b
<*> forall (t :: * -> *) (f :: * -> *) a b.
(Traversable t, Applicative f) =>
(a -> f b) -> t a -> f (t b)
traverse forall (m :: * -> *). HasBuiltins m => Term -> m Term
constructorForm Blocked Term
nb'

    Blocked Term
mb <- forall (t :: * -> *) (f :: * -> *) a b.
(Traversable t, Applicative f) =>
(a -> f b) -> t a -> f (t b)
traverse forall (m :: * -> *). HasBuiltins m => Term -> m Term
unLevel Blocked Term
mb''
    Blocked Term
nb <- forall (t :: * -> *) (f :: * -> *) a b.
(Traversable t, Applicative f) =>
(a -> f b) -> t a -> f (t b)
traverse forall (m :: * -> *). HasBuiltins m => Term -> m Term
unLevel Blocked Term
nb''

    Bool
cmpBlocked <- forall (m :: * -> *) a. MonadTCEnv m => Lens' a TCEnv -> m a
viewTC Lens' Bool TCEnv
eCompareBlocked

    let m :: Term
m = forall t a. Blocked' t a -> a
ignoreBlocking Blocked Term
mb
        n :: Term
n = forall t a. Blocked' t a -> a
ignoreBlocking Blocked Term
nb

        checkDefinitionalEquality :: m ()
checkDefinitionalEquality = forall (m :: * -> *). Monad m => m Bool -> m () -> m ()
unlessM (forall (m :: * -> *).
(PureTCM m, MonadBlock m) =>
Comparison -> CompareAs -> Term -> Term -> m Bool
pureCompareAs Comparison
CmpEq CompareAs
t Term
m Term
n) m ()
notEqual

        notEqual :: m ()
notEqual = forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
TypeError -> m a
typeError forall a b. (a -> b) -> a -> b
$ Comparison -> Term -> Term -> CompareAs -> TypeError
UnequalTerms Comparison
cmp Term
m Term
n CompareAs
t

        dir :: CompareDirection
dir = Comparison -> CompareDirection
fromCmp Comparison
cmp
        rid :: CompareDirection
rid = CompareDirection -> CompareDirection
flipCmp CompareDirection
dir     -- The reverse direction.  Bad name, I know.

        assign :: CompareDirection -> MetaId -> Elims -> Term -> m ()
assign CompareDirection
dir MetaId
x Elims
es Term
v = forall (m :: * -> *).
MonadConversion m =>
CompareDirection
-> MetaId
-> Elims
-> Term
-> CompareAs
-> (Term -> Term -> m ())
-> m ()
assignE CompareDirection
dir MetaId
x Elims
es Term
v CompareAs
t forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *).
MonadConversion m =>
CompareDirection -> CompareAs -> Term -> Term -> m ()
compareAtomDir CompareDirection
dir CompareAs
t

    forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"tc.conv.atom" Int
30 forall a b. (a -> b) -> a -> b
$
      TCMT IO Doc
"compareAtom" 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 (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Blocked Term
mb forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Comparison
cmp
                             , forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Blocked Term
nb
                             , forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM CompareAs
t
                             ]
    forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"tc.conv.atom" Int
80 forall a b. (a -> b) -> a -> b
$
      TCMT IO Doc
"compareAtom" 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 :: * -> *) a. (Applicative m, Pretty a) => a -> m Doc
pretty Blocked Term
mb forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Comparison
cmp
                                  , forall (m :: * -> *) a. (Applicative m, Pretty a) => a -> m Doc
pretty Blocked Term
nb
                                  , TCMT IO Doc
":" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall (m :: * -> *) a. (Applicative m, Pretty a) => a -> m Doc
pretty CompareAs
t ]
    case (Blocked Term
mb, Blocked Term
nb) of
      -- equate two metas x and y.  if y is the younger meta,
      -- try first y := x and then x := y
      (Blocked Term, Blocked Term)
_ | MetaV MetaId
x Elims
xArgs <- forall t a. Blocked' t a -> a
ignoreBlocking Blocked Term
mb,   -- Can be either Blocked or NotBlocked depending on
          MetaV MetaId
y Elims
yArgs <- forall t a. Blocked' t a -> a
ignoreBlocking Blocked Term
nb -> -- envCompareBlocked check above.
        forall (m :: * -> *).
MonadConversion m =>
Comparison
-> CompareAs -> MetaId -> Elims -> MetaId -> Elims -> m ()
compareMetas Comparison
cmp CompareAs
t MetaId
x Elims
xArgs MetaId
y Elims
yArgs

      -- one side a meta
      (Blocked Term, Blocked Term)
_ | MetaV MetaId
x Elims
es <- forall t a. Blocked' t a -> a
ignoreBlocking Blocked Term
mb -> CompareDirection -> MetaId -> Elims -> Term -> m ()
assign CompareDirection
dir MetaId
x Elims
es Term
n
      (Blocked Term, Blocked Term)
_ | MetaV MetaId
x Elims
es <- forall t a. Blocked' t a -> a
ignoreBlocking Blocked Term
nb -> CompareDirection -> MetaId -> Elims -> Term -> m ()
assign CompareDirection
rid MetaId
x Elims
es Term
m
      (Blocked{}, Blocked{}) | Bool -> Bool
not Bool
cmpBlocked  -> m ()
checkDefinitionalEquality
      (Blocked Blocker
b Term
_, Blocked Term
_) | Bool -> Bool
not Bool
cmpBlocked -> forall (m :: * -> *).
MonadConversion m =>
CompareDirection -> Blocker -> CompareAs -> Term -> Term -> m ()
useInjectivity (Comparison -> CompareDirection
fromCmp Comparison
cmp) Blocker
b CompareAs
t Term
m Term
n   -- The blocked term  goes first
      (Blocked Term
_, Blocked Blocker
b Term
_) | Bool -> Bool
not Bool
cmpBlocked -> forall (m :: * -> *).
MonadConversion m =>
CompareDirection -> Blocker -> CompareAs -> Term -> Term -> m ()
useInjectivity (CompareDirection -> CompareDirection
flipCmp forall a b. (a -> b) -> a -> b
$ Comparison -> CompareDirection
fromCmp Comparison
cmp) Blocker
b CompareAs
t Term
n Term
m
      (Blocked Term, Blocked Term)
bs -> do
        forall (m :: * -> *) a. MonadError TCErr m => Blocker -> m a -> m a
blockOnError Blocker
blocker forall a b. (a -> b) -> a -> b
$ do
        -- -- Andreas, 2013-10-20 put projection-like function
        -- -- into the spine, to make compareElims work.
        -- -- 'False' means: leave (Def f []) unchanged even for
        -- -- proj-like funs.
        -- m <- elimView False m
        -- n <- elimView False n
        -- Andreas, 2015-07-01, actually, don't put them into the spine.
        -- Polarity cannot be communicated properly if projection-like
        -- functions are post-fix.
        case (Term
m, Term
n) of
          (Pi{}, Pi{}) -> Term -> Term -> m ()
equalFun Term
m Term
n

          (Sort Sort
s1, Sort Sort
s2) ->
            forall (m :: * -> *) a. Monad m => m Bool -> m a -> m a -> m a
ifM (PragmaOptions -> Bool
optCumulativity forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> forall (m :: * -> *). HasOptions m => m PragmaOptions
pragmaOptions)
              (forall (m :: * -> *).
MonadConversion m =>
Comparison -> Sort -> Sort -> m ()
compareSort Comparison
cmp Sort
s1 Sort
s2)
              (forall (m :: * -> *). MonadConversion m => Sort -> Sort -> m ()
equalSort Sort
s1 Sort
s2)

          (Lit Literal
l1, Lit Literal
l2) | Literal
l1 forall a. Eq a => a -> a -> Bool
== Literal
l2 -> forall (m :: * -> *) a. Monad m => a -> m a
return ()
          (Var Int
i Elims
es, Var Int
i' Elims
es') | Int
i forall a. Eq a => a -> a -> Bool
== Int
i' -> do
              Type
a <- forall (m :: * -> *).
(Applicative m, MonadFail m, MonadTCEnv m) =>
Int -> m Type
typeOfBV Int
i
              -- Variables are invariant in their arguments
              forall (m :: * -> *).
MonadConversion m =>
[Polarity] -> [IsForced] -> Type -> Term -> Elims -> Elims -> m ()
compareElims [] [] Type
a (Int -> Term
var Int
i) Elims
es Elims
es'

          -- The case of definition application:
          (Def QName
f Elims
es, Def QName
f' Elims
es') -> do

              -- 1. All absurd lambdas are equal.
              forall (m :: * -> *). Monad m => m Bool -> m () -> m ()
unlessM (forall (m :: * -> *). MonadConversion m => QName -> QName -> m Bool
bothAbsurd QName
f QName
f') forall a b. (a -> b) -> a -> b
$ do

              -- 2. If the heads are unequal, the only chance is subtyping between SIZE and SIZELT.
              if QName
f forall a. Eq a => a -> a -> Bool
/= QName
f' then forall (m :: * -> *).
MonadConversion m =>
Comparison
-> CompareAs
-> Term
-> Term
-> QName
-> Elims
-> QName
-> Elims
-> m ()
trySizeUniv Comparison
cmp CompareAs
t Term
m Term
n QName
f Elims
es QName
f' Elims
es' else do

              -- 3. If the heads are equal:
              -- 3a. If there are no arguments, we are done.
              forall (f :: * -> *). Applicative f => Bool -> f () -> f ()
unless (forall (t :: * -> *) a. Foldable t => t a -> Bool
null Elims
es Bool -> Bool -> Bool
&& forall (t :: * -> *) a. Foldable t => t a -> Bool
null Elims
es') forall a b. (a -> b) -> a -> b
$ do

              -- 3b. If some cubical magic kicks in, we are done.
              forall (m :: * -> *). Monad m => m Bool -> m () -> m ()
unlessM (MonadConversion m => QName -> Elims -> Elims -> m Bool
compareEtaPrims QName
f Elims
es Elims
es') forall a b. (a -> b) -> a -> b
$ do

              -- 3c. Oh no, we actually have to work and compare the eliminations!
               Type
a <- forall (m :: * -> *).
MonadConversion m =>
QName -> Elims -> Elims -> m Type
computeElimHeadType QName
f Elims
es Elims
es'
               -- The polarity vector of projection-like functions
               -- does not include the parameters.
               [Polarity]
pol <- forall (m :: * -> *).
HasConstInfo m =>
Comparison -> QName -> m [Polarity]
getPolarity' Comparison
cmp QName
f
               forall (m :: * -> *).
MonadConversion m =>
[Polarity] -> [IsForced] -> Type -> Term -> Elims -> Elims -> m ()
compareElims [Polarity]
pol [] Type
a (QName -> Elims -> Term
Def QName
f []) Elims
es Elims
es'

          -- Due to eta-expansion, these constructors are fully applied.
          (Con ConHead
x ConInfo
ci Elims
xArgs, Con ConHead
y ConInfo
_ Elims
yArgs)
              | ConHead
x forall a. Eq a => a -> a -> Bool
== ConHead
y -> do
                  -- Get the type of the constructor instantiated to the datatype parameters.
                  Type
a' <- case CompareAs
t of
                    AsTermsOf Type
a -> forall {m :: * -> *}.
(MonadBlock m, PureTCM m) =>
ConHead -> Type -> m Type
conType ConHead
x Type
a
                    CompareAs
AsSizes   -> forall a. HasCallStack => a
__IMPOSSIBLE__
                    CompareAs
AsTypes   -> forall a. HasCallStack => a
__IMPOSSIBLE__
                  [IsForced]
forcedArgs <- forall (m :: * -> *). HasConstInfo m => QName -> m [IsForced]
getForcedArgs forall a b. (a -> b) -> a -> b
$ ConHead -> QName
conName ConHead
x
                  -- Constructors are covariant in their arguments
                  -- (see test/succeed/CovariantConstructors).
                  forall (m :: * -> *).
MonadConversion m =>
[Polarity] -> [IsForced] -> Type -> Term -> Elims -> Elims -> m ()
compareElims (forall a. a -> [a]
repeat forall a b. (a -> b) -> a -> b
$ Comparison -> Polarity
polFromCmp Comparison
cmp) [IsForced]
forcedArgs Type
a' (ConHead -> ConInfo -> Elims -> Term
Con ConHead
x ConInfo
ci []) Elims
xArgs Elims
yArgs
          (Term, Term)
_ -> m ()
notEqual
    where
        -- returns True in case we handled the comparison already.
        compareEtaPrims :: MonadConversion m => QName -> Elims -> Elims -> m Bool
        compareEtaPrims :: MonadConversion m => QName -> Elims -> Elims -> m Bool
compareEtaPrims QName
q Elims
es Elims
es' = do
          Maybe QName
munglue <- forall (m :: * -> *).
HasBuiltins m =>
VerboseKey -> m (Maybe QName)
getPrimitiveName' VerboseKey
builtin_unglue
          Maybe QName
munglueU <- forall (m :: * -> *).
HasBuiltins m =>
VerboseKey -> m (Maybe QName)
getPrimitiveName' VerboseKey
builtin_unglueU
          Maybe QName
msubout <- forall (m :: * -> *).
HasBuiltins m =>
VerboseKey -> m (Maybe QName)
getPrimitiveName' VerboseKey
builtinSubOut
          case () of
            ()
_ | forall a. a -> Maybe a
Just QName
q forall a. Eq a => a -> a -> Bool
== Maybe QName
munglue -> QName -> Elims -> Elims -> m Bool
compareUnglueApp QName
q Elims
es Elims
es'
            ()
_ | forall a. a -> Maybe a
Just QName
q forall a. Eq a => a -> a -> Bool
== Maybe QName
munglueU -> MonadConversion m => QName -> Elims -> Elims -> m Bool
compareUnglueUApp QName
q Elims
es Elims
es'
            ()
_ | forall a. a -> Maybe a
Just QName
q forall a. Eq a => a -> a -> Bool
== Maybe QName
msubout -> QName -> Elims -> Elims -> m Bool
compareSubApp QName
q Elims
es Elims
es'
            ()
_                     -> forall (m :: * -> *) a. Monad m => a -> m a
return Bool
False
        compareSubApp :: QName -> Elims -> Elims -> m Bool
compareSubApp QName
q Elims
es Elims
es' = do
          let (Elims
as,Elims
bs) = forall a. Int -> [a] -> ([a], [a])
splitAt Int
5 Elims
es; (Elims
as',Elims
bs') = forall a. Int -> [a] -> ([a], [a])
splitAt Int
5 Elims
es'
          case (forall a. [Elim' a] -> Maybe [Arg a]
allApplyElims Elims
as, forall a. [Elim' a] -> Maybe [Arg a]
allApplyElims Elims
as') of
            (Just [Arg Term
a,Arg Term
bA,Arg Term
phi,Arg Term
u,Arg Term
x], Just [Arg Term
a',Arg Term
bA',Arg Term
phi',Arg Term
u',Arg Term
x']) -> do
              Term
tSub <- forall (m :: * -> *).
(HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) =>
m Term
primSub
              -- Andrea, 28-07-16:
              -- comparing the types is most probably wasteful,
              -- since b and b' should be neutral terms, but it's a
              -- precondition for the compareAtom call to make
              -- sense.
              forall (m :: * -> *). MonadConversion m => Type -> Type -> m ()
equalType (forall t a. Sort' t -> a -> Type'' t a
El (Term -> Sort
tmSSort forall a b. (a -> b) -> a -> b
$ forall e. Arg e -> e
unArg Arg Term
a) forall a b. (a -> b) -> a -> b
$ forall t. Apply t => t -> [Arg Term] -> t
apply Term
tSub forall a b. (a -> b) -> a -> b
$ Arg Term
a forall a. a -> [a] -> [a]
: forall a b. (a -> b) -> [a] -> [b]
map (forall a. LensHiding a => Hiding -> a -> a
setHiding Hiding
NotHidden) [Arg Term
bA,Arg Term
phi,Arg Term
u])
                        (forall t a. Sort' t -> a -> Type'' t a
El (Term -> Sort
tmSSort forall a b. (a -> b) -> a -> b
$ forall e. Arg e -> e
unArg Arg Term
a) forall a b. (a -> b) -> a -> b
$ forall t. Apply t => t -> [Arg Term] -> t
apply Term
tSub forall a b. (a -> b) -> a -> b
$ Arg Term
a forall a. a -> [a] -> [a]
: forall a b. (a -> b) -> [a] -> [b]
map (forall a. LensHiding a => Hiding -> a -> a
setHiding Hiding
NotHidden) [Arg Term
bA',Arg Term
phi',Arg Term
u'])
              forall (m :: * -> *).
MonadConversion m =>
Comparison -> CompareAs -> Term -> Term -> m ()
compareAtom Comparison
cmp (Type -> CompareAs
AsTermsOf forall a b. (a -> b) -> a -> b
$ forall t a. Sort' t -> a -> Type'' t a
El (Term -> Sort
tmSSort forall a b. (a -> b) -> a -> b
$ forall e. Arg e -> e
unArg Arg Term
a) forall a b. (a -> b) -> a -> b
$ forall t. Apply t => t -> [Arg Term] -> t
apply Term
tSub forall a b. (a -> b) -> a -> b
$ Arg Term
a forall a. a -> [a] -> [a]
: forall a b. (a -> b) -> [a] -> [b]
map (forall a. LensHiding a => Hiding -> a -> a
setHiding Hiding
NotHidden) [Arg Term
bA,Arg Term
phi,Arg Term
u])
                              (forall e. Arg e -> e
unArg Arg Term
x) (forall e. Arg e -> e
unArg Arg Term
x')
              forall (m :: * -> *).
MonadConversion m =>
[Polarity] -> [IsForced] -> Type -> Term -> Elims -> Elims -> m ()
compareElims [] [] (forall t a. Sort' t -> a -> Type'' t a
El (Term -> Sort
tmSort (forall e. Arg e -> e
unArg Arg Term
a)) (forall e. Arg e -> e
unArg Arg Term
bA)) (QName -> Elims -> Term
Def QName
q Elims
as) Elims
bs Elims
bs'
              forall (m :: * -> *) a. Monad m => a -> m a
return Bool
True
            (Maybe [Arg Term], Maybe [Arg Term])
_  -> forall (m :: * -> *) a. Monad m => a -> m a
return Bool
False
        compareUnglueApp :: QName -> Elims -> Elims -> m Bool
compareUnglueApp QName
q Elims
es Elims
es' = do
          let (Elims
as,Elims
bs) = forall a. Int -> [a] -> ([a], [a])
splitAt Int
7 Elims
es; (Elims
as',Elims
bs') = forall a. Int -> [a] -> ([a], [a])
splitAt Int
7 Elims
es'
          case (forall a. [Elim' a] -> Maybe [Arg a]
allApplyElims Elims
as, forall a. [Elim' a] -> Maybe [Arg a]
allApplyElims Elims
as') of
            (Just [Arg Term
la,Arg Term
lb,Arg Term
bA,Arg Term
phi,Arg Term
bT,Arg Term
e,Arg Term
b], Just [Arg Term
la',Arg Term
lb',Arg Term
bA',Arg Term
phi',Arg Term
bT',Arg Term
e',Arg Term
b']) -> do
              Term
tGlue <- forall (m :: * -> *).
(HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) =>
VerboseKey -> m Term
getPrimitiveTerm VerboseKey
builtinGlue
              -- Andrea, 28-07-16:
              -- comparing the types is most probably wasteful,
              -- since b and b' should be neutral terms, but it's a
              -- precondition for the compareAtom call to make
              -- sense.
              -- equalType (El (tmSort (unArg lb)) $ apply tGlue $ [la,lb] ++ map (setHiding NotHidden) [bA,phi,bT,e])
              --           (El (tmSort (unArg lb')) $ apply tGlue $ [la',lb'] ++ map (setHiding NotHidden) [bA',phi',bT',e'])
              forall (m :: * -> *).
MonadConversion m =>
Comparison -> CompareAs -> Term -> Term -> m ()
compareAtom Comparison
cmp (Type -> CompareAs
AsTermsOf forall a b. (a -> b) -> a -> b
$ forall t a. Sort' t -> a -> Type'' t a
El (Term -> Sort
tmSort (forall e. Arg e -> e
unArg Arg Term
lb)) forall a b. (a -> b) -> a -> b
$ forall t. Apply t => t -> [Arg Term] -> t
apply Term
tGlue forall a b. (a -> b) -> a -> b
$ [Arg Term
la,Arg Term
lb] forall a. [a] -> [a] -> [a]
++ forall a b. (a -> b) -> [a] -> [b]
map (forall a. LensHiding a => Hiding -> a -> a
setHiding Hiding
NotHidden) [Arg Term
bA,Arg Term
phi,Arg Term
bT,Arg Term
e])
                              (forall e. Arg e -> e
unArg Arg Term
b) (forall e. Arg e -> e
unArg Arg Term
b')
              forall (m :: * -> *).
MonadConversion m =>
[Polarity] -> [IsForced] -> Type -> Term -> Elims -> Elims -> m ()
compareElims [] [] (forall t a. Sort' t -> a -> Type'' t a
El (Term -> Sort
tmSort (forall e. Arg e -> e
unArg Arg Term
la)) (forall e. Arg e -> e
unArg Arg Term
bA)) (QName -> Elims -> Term
Def QName
q Elims
as) Elims
bs Elims
bs'
              forall (m :: * -> *) a. Monad m => a -> m a
return Bool
True
            (Maybe [Arg Term], Maybe [Arg Term])
_  -> forall (m :: * -> *) a. Monad m => a -> m a
return Bool
False
        compareUnglueUApp :: MonadConversion m => QName -> Elims -> Elims -> m Bool
        compareUnglueUApp :: MonadConversion m => QName -> Elims -> Elims -> m Bool
compareUnglueUApp QName
q Elims
es Elims
es' = do
          let (Elims
as,Elims
bs) = forall a. Int -> [a] -> ([a], [a])
splitAt Int
5 Elims
es; (Elims
as',Elims
bs') = forall a. Int -> [a] -> ([a], [a])
splitAt Int
5 Elims
es'
          case (forall a. [Elim' a] -> Maybe [Arg a]
allApplyElims Elims
as, forall a. [Elim' a] -> Maybe [Arg a]
allApplyElims Elims
as') of
            (Just [Arg Term
la,Arg Term
phi,Arg Term
bT,Arg Term
bAS,Arg Term
b], Just [Arg Term
la',Arg Term
phi',Arg Term
bT',Arg Term
bA',Arg Term
b']) -> do
              Term
tHComp <- forall (m :: * -> *).
(HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) =>
m Term
primHComp
              Term
tLSuc <- forall (m :: * -> *).
(HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) =>
m Term
primLevelSuc
              Term
tSubOut <- forall (m :: * -> *).
(HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) =>
m Term
primSubOut
              Term
iz <- forall (m :: * -> *).
(HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) =>
m Term
primIZero
              let lsuc :: Term -> Term
lsuc Term
t = Term
tLSuc forall t. Apply t => t -> [Arg Term] -> t
`apply` [forall e. e -> Arg e
argN Term
t]
                  s :: Sort
s = Term -> Sort
tmSort forall a b. (a -> b) -> a -> b
$ forall e. Arg e -> e
unArg Arg Term
la
                  sucla :: Arg Term
sucla = Term -> Term
lsuc forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> Arg Term
la
              Term
bA <- forall (m :: * -> *) a. Names -> NamesT m a -> m a
runNamesT [] forall a b. (a -> b) -> a -> b
$ do
                [NamesT m Term
la,NamesT m Term
phi,NamesT m Term
bT,NamesT m Term
bAS] <- forall (t :: * -> *) (m :: * -> *) a b.
(Traversable t, Monad m) =>
(a -> m b) -> t a -> m (t b)
mapM (forall (m :: * -> *) a.
(MonadFail m, Subst a) =>
a -> NamesT m (NamesT m a)
open forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall e. Arg e -> e
unArg) [Arg Term
la,Arg Term
phi,Arg Term
bT,Arg Term
bAS]
                (forall (f :: * -> *) a. Applicative f => a -> f a
pure Term
tSubOut forall (m :: * -> *). Applicative m => m Term -> m Term -> m Term
<#> (forall (f :: * -> *) a. Applicative f => a -> f a
pure Term
tLSuc forall (m :: * -> *). Applicative m => m Term -> m Term -> m Term
<@> NamesT m Term
la) forall (m :: * -> *). Applicative m => m Term -> m Term -> m Term
<#> (Sort -> Term
Sort forall b c a. (b -> c) -> (a -> b) -> a -> c
. Term -> Sort
tmSort forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> NamesT m Term
la) forall (m :: * -> *). Applicative m => m Term -> m Term -> m Term
<#> NamesT m Term
phi forall (m :: * -> *). Applicative m => m Term -> m Term -> m Term
<#> (NamesT m Term
bT forall (m :: * -> *). Applicative m => m Term -> m Term -> m Term
<@> forall (m :: * -> *).
(HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) =>
m Term
primIZero) forall (m :: * -> *). Applicative m => m Term -> m Term -> m Term
<@> NamesT m Term
bAS)
              forall (m :: * -> *).
MonadConversion m =>
Comparison -> CompareAs -> Term -> Term -> m ()
compareAtom Comparison
cmp (Type -> CompareAs
AsTermsOf forall a b. (a -> b) -> a -> b
$ forall t a. Sort' t -> a -> Type'' t a
El (Term -> Sort
tmSort forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall e. Arg e -> e
unArg forall a b. (a -> b) -> a -> b
$ Arg Term
sucla) forall a b. (a -> b) -> a -> b
$ forall t. Apply t => t -> [Arg Term] -> t
apply Term
tHComp forall a b. (a -> b) -> a -> b
$ [Arg Term
sucla, forall e. e -> Arg e
argH (Sort -> Term
Sort Sort
s), Arg Term
phi] forall a. [a] -> [a] -> [a]
++ [forall e. e -> Arg e
argH (forall e. Arg e -> e
unArg Arg Term
bT), forall e. e -> Arg e
argH Term
bA])
                              (forall e. Arg e -> e
unArg Arg Term
b) (forall e. Arg e -> e
unArg Arg Term
b')
              forall (m :: * -> *).
MonadConversion m =>
[Polarity] -> [IsForced] -> Type -> Term -> Elims -> Elims -> m ()
compareElims [] [] (forall t a. Sort' t -> a -> Type'' t a
El Sort
s Term
bA) (QName -> Elims -> Term
Def QName
q Elims
as) Elims
bs Elims
bs'
              forall (m :: * -> *) a. Monad m => a -> m a
return Bool
True
            (Maybe [Arg Term], Maybe [Arg Term])
_  -> forall (m :: * -> *) a. Monad m => a -> m a
return Bool
False
        -- Andreas, 2013-05-15 due to new postponement strategy, type can now be blocked
        conType :: ConHead -> Type -> m Type
conType ConHead
c Type
t = do
          Type
t <- forall (m :: * -> *) t.
(MonadReduce m, MonadBlock m, IsMeta t, Reduce t) =>
t -> m t
abortIfBlocked Type
t
          let impossible :: m Type
impossible = do
                forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"impossible" Int
10 forall a b. (a -> b) -> a -> b
$
                  TCMT IO Doc
"expected data/record type, found " 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 =>
VerboseKey -> Int -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"impossible" Int
70 forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *). Functor m => Int -> m Doc -> m Doc
nest Int
2 forall a b. (a -> b) -> a -> b
$ TCMT IO Doc
"raw =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall (m :: * -> *) a. (Applicative m, Pretty a) => a -> m Doc
pretty Type
t
                -- __IMPOSSIBLE__
                -- Andreas, 2013-10-20:  in case termination checking fails
                -- we might get some unreduced types here.
                -- In issue 921, this happens during the final attempt
                -- to solve left-over constraints.
                -- Thus, instead of crashing, just give up gracefully.
                forall (m :: * -> *) a. MonadBlock m => Blocker -> m a
patternViolation Blocker
neverUnblock
          forall b a. b -> (a -> b) -> Maybe a -> b
maybe m Type
impossible (forall (m :: * -> *) a. Monad m => a -> m a
return forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall a b. (a, b) -> b
snd) forall (m :: * -> *) a b. Monad m => (a -> m b) -> m a -> m b
=<< forall (m :: * -> *).
PureTCM m =>
ConHead -> Type -> m (Maybe ((QName, Type, [Arg Term]), Type))
getFullyAppliedConType ConHead
c Type
t
        equalFun :: Term -> Term -> m ()
equalFun Term
t1 Term
t2 = case (Term
t1, Term
t2) of
          (Pi Dom Type
dom1 Abs Type
b1, Pi Dom Type
dom2 Abs Type
b2) -> do
            forall (m :: * -> *) a.
MonadDebug m =>
VerboseKey -> Int -> VerboseKey -> m a -> m a
verboseBracket VerboseKey
"tc.conv.fun" Int
15 VerboseKey
"compare function types" forall a b. (a -> b) -> a -> b
$ do
              forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"tc.conv.fun" Int
20 forall a b. (a -> b) -> a -> b
$ 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
"t1 =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Term
t1
                , TCMT IO Doc
"t2 =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Term
t2
                ]
              forall (m :: * -> *) c b.
(MonadConversion m, Free c) =>
Comparison
-> Dom Type
-> Dom Type
-> Abs b
-> Abs c
-> m ()
-> m ()
-> m ()
-> m ()
-> m ()
-> m ()
-> m ()
compareDom Comparison
cmp Dom Type
dom2 Dom Type
dom1 Abs Type
b1 Abs Type
b2 m ()
errH m ()
errR m ()
errQ m ()
errC m ()
errF forall a b. (a -> b) -> a -> b
$
                forall (m :: * -> *).
MonadConversion m =>
Comparison -> Type -> Type -> m ()
compareType Comparison
cmp (forall a. Subst a => Abs a -> a
absBody Abs Type
b1) (forall a. Subst a => Abs a -> a
absBody Abs Type
b2)
            where
            errH :: m ()
errH = forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
TypeError -> m a
typeError forall a b. (a -> b) -> a -> b
$ Term -> Term -> TypeError
UnequalHiding Term
t1 Term
t2
            errR :: m ()
errR = forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
TypeError -> m a
typeError forall a b. (a -> b) -> a -> b
$ Comparison -> Term -> Term -> TypeError
UnequalRelevance Comparison
cmp Term
t1 Term
t2
            errQ :: m ()
errQ = forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
TypeError -> m a
typeError forall a b. (a -> b) -> a -> b
$ Comparison -> Term -> Term -> TypeError
UnequalQuantity  Comparison
cmp Term
t1 Term
t2
            errC :: m ()
errC = forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
TypeError -> m a
typeError forall a b. (a -> b) -> a -> b
$ Comparison -> Term -> Term -> TypeError
UnequalCohesion Comparison
cmp Term
t1 Term
t2
            errF :: m ()
errF = forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
TypeError -> m a
typeError forall a b. (a -> b) -> a -> b
$ Comparison -> Term -> Term -> TypeError
UnequalFiniteness Comparison
cmp Term
t1 Term
t2
          (Term, Term)
_ -> forall a. HasCallStack => a
__IMPOSSIBLE__

-- | Check whether @x xArgs `cmp` y yArgs@
compareMetas :: MonadConversion m => Comparison -> CompareAs -> MetaId -> Elims -> MetaId -> Elims -> m ()
compareMetas :: forall (m :: * -> *).
MonadConversion m =>
Comparison
-> CompareAs -> MetaId -> Elims -> MetaId -> Elims -> m ()
compareMetas Comparison
cmp CompareAs
t MetaId
x Elims
xArgs MetaId
y Elims
yArgs | MetaId
x forall a. Eq a => a -> a -> Bool
== MetaId
y = forall (m :: * -> *) a. MonadError TCErr m => Blocker -> m a -> m a
blockOnError (MetaId -> Blocker
unblockOnMeta MetaId
x) forall a b. (a -> b) -> a -> b
$ do
  Bool
cmpBlocked <- forall (m :: * -> *) a. MonadTCEnv m => Lens' a TCEnv -> m a
viewTC Lens' Bool TCEnv
eCompareBlocked
  let ok :: m ()
ok    = forall (m :: * -> *) a. Monad m => a -> m a
return ()
      notOk :: m a
notOk = forall (m :: * -> *) a. MonadBlock m => Blocker -> m a
patternViolation Blocker
neverUnblock
      fallback :: m ()
fallback = do
        -- Fallback: check definitional equality
        Type
a <- forall (m :: * -> *). ReadTCState m => MetaId -> m Type
metaType MetaId
x
        forall (m :: * -> *) a.
(MonadBlock m, PureTCM m, Show a) =>
PureConversionT m a -> m (Maybe a)
runPureConversion (forall (m :: * -> *).
MonadConversion m =>
[Polarity] -> [IsForced] -> Type -> Term -> Elims -> Elims -> m ()
compareElims [] [] Type
a (MetaId -> Elims -> Term
MetaV MetaId
x []) Elims
xArgs Elims
yArgs) forall (m :: * -> *) a b. Monad m => m a -> (a -> m b) -> m b
>>= \case
          Just{}  -> m ()
ok
          Maybe ()
Nothing -> forall {a}. m a
notOk
  if | Bool
cmpBlocked -> do
         Type
a <- forall (m :: * -> *). ReadTCState m => MetaId -> m Type
metaType MetaId
x
         forall (m :: * -> *).
MonadConversion m =>
[Polarity] -> [IsForced] -> Type -> Term -> Elims -> Elims -> m ()
compareElims [] [] Type
a (MetaId -> Elims -> Term
MetaV MetaId
x []) Elims
xArgs Elims
yArgs
     | Bool
otherwise -> case Elims -> Elims -> Maybe [Bool]
intersectVars Elims
xArgs Elims
yArgs of
         -- all relevant arguments are variables
         Just [Bool]
kills -> do
           -- kills is a list with 'True' for each different var
           PruneResult
killResult <- forall (m :: * -> *).
MonadMetaSolver m =>
[Bool] -> MetaId -> m PruneResult
killArgs [Bool]
kills MetaId
x
           case PruneResult
killResult of
             PruneResult
NothingToPrune   -> m ()
ok
             PruneResult
PrunedEverything -> m ()
ok
             PruneResult
PrunedNothing    -> m ()
fallback
             PruneResult
PrunedSomething  -> m ()
fallback
         -- not all relevant arguments are variables
         Maybe [Bool]
Nothing -> m ()
fallback
compareMetas Comparison
cmp CompareAs
t MetaId
x Elims
xArgs MetaId
y Elims
yArgs = do
  [MetaPriority
p1, MetaPriority
p2] <- forall (t :: * -> *) (m :: * -> *) a b.
(Traversable t, Monad m) =>
(a -> m b) -> t a -> m (t b)
mapM forall (m :: * -> *).
(HasCallStack, MonadDebug m, ReadTCState m) =>
MetaId -> m MetaPriority
getMetaPriority [MetaId
x,MetaId
y]
  let dir :: CompareDirection
dir = Comparison -> CompareDirection
fromCmp Comparison
cmp
      rid :: CompareDirection
rid = CompareDirection -> CompareDirection
flipCmp CompareDirection
dir     -- The reverse direction.  Bad name, I know.
      retry :: m a
retry = forall (m :: * -> *) a. MonadBlock m => Blocker -> m a
patternViolation Blocker
alwaysUnblock
  -- First try the one with the highest priority. If that doesn't
  -- work, try the low priority one.
  let (m ()
solve1, m ()
solve2)
        | (MetaPriority
p1, MetaId
x) forall a. Ord a => a -> a -> Bool
> (MetaPriority
p2, MetaId
y) = (m ()
l1, m ()
r2)
        | Bool
otherwise         = (m ()
r1, m ()
l2)
        where l1 :: m ()
l1 = forall (m :: * -> *).
MonadConversion m =>
CompareDirection
-> MetaId
-> Elims
-> Term
-> CompareAs
-> (Term -> Term -> m ())
-> m ()
assignE CompareDirection
dir MetaId
x Elims
xArgs (MetaId -> Elims -> Term
MetaV MetaId
y Elims
yArgs) CompareAs
t forall a b. (a -> b) -> a -> b
$ \ Term
_ Term
_ -> forall {a}. m a
retry
              r1 :: m ()
r1 = forall (m :: * -> *).
MonadConversion m =>
CompareDirection
-> MetaId
-> Elims
-> Term
-> CompareAs
-> (Term -> Term -> m ())
-> m ()
assignE CompareDirection
rid MetaId
y Elims
yArgs (MetaId -> Elims -> Term
MetaV MetaId
x Elims
xArgs) CompareAs
t forall a b. (a -> b) -> a -> b
$ \ Term
_ Term
_ -> forall {a}. m a
retry
              -- Careful: the first attempt might prune the low
              -- priority meta! (Issue #2978)
              l2 :: m ()
l2 = forall (m :: * -> *) a. Monad m => m Bool -> m a -> m a -> m a
ifM (forall a (m :: * -> *).
(IsInstantiatedMeta a, MonadFail m, ReadTCState m) =>
a -> m Bool
isInstantiatedMeta MetaId
x) forall {a}. m a
retry m ()
l1
              r2 :: m ()
r2 = forall (m :: * -> *) a. Monad m => m Bool -> m a -> m a -> m a
ifM (forall a (m :: * -> *).
(IsInstantiatedMeta a, MonadFail m, ReadTCState m) =>
a -> m Bool
isInstantiatedMeta MetaId
y) forall {a}. m a
retry m ()
r1

  -- Unblock on both unblockers of solve1 and solve2
  forall (m :: * -> *) a.
MonadBlock m =>
(Blocker -> m a) -> m a -> m a
catchPatternErr (forall (m :: * -> *) a. MonadBlock m => Blocker -> m a -> m a
`addOrUnblocker` m ()
solve2) m ()
solve1

-- | Check whether @a1 `cmp` a2@ and continue in context extended by @a1@.
compareDom :: (MonadConversion m , Free c)
  => Comparison -- ^ @cmp@ The comparison direction
  -> Dom Type   -- ^ @a1@  The smaller domain.
  -> Dom Type   -- ^ @a2@  The other domain.
  -> Abs b      -- ^ @b1@  The smaller codomain.
  -> Abs c      -- ^ @b2@  The bigger codomain.
  -> m ()     -- ^ Continuation if mismatch in 'Hiding'.
  -> m ()     -- ^ Continuation if mismatch in 'Relevance'.
  -> m ()     -- ^ Continuation if mismatch in 'Quantity'.
  -> m ()     -- ^ Continuation if mismatch in 'Cohesion'.
  -> m ()     -- ^ Continuation if mismatch in 'annFinite'.
  -> m ()     -- ^ Continuation if comparison is successful.
  -> m ()
compareDom :: forall (m :: * -> *) c b.
(MonadConversion m, Free c) =>
Comparison
-> Dom Type
-> Dom Type
-> Abs b
-> Abs c
-> m ()
-> m ()
-> m ()
-> m ()
-> m ()
-> m ()
-> m ()
compareDom Comparison
cmp0
  dom1 :: Dom Type
dom1@(Dom{domInfo :: forall t e. Dom' t e -> ArgInfo
domInfo = ArgInfo
i1, unDom :: forall t e. Dom' t e -> e
unDom = Type
a1})
  dom2 :: Dom Type
dom2@(Dom{domInfo :: forall t e. Dom' t e -> ArgInfo
domInfo = ArgInfo
i2, unDom :: forall t e. Dom' t e -> e
unDom = Type
a2})
  Abs b
b1 Abs c
b2 m ()
errH m ()
errR m ()
errQ m ()
errC m ()
errF m ()
cont = do
  if | Bool -> Bool
not forall a b. (a -> b) -> a -> b
$ forall a b. (LensHiding a, LensHiding b) => a -> b -> Bool
sameHiding Dom Type
dom1 Dom Type
dom2 -> m ()
errH
     | Bool -> Bool
not forall a b. (a -> b) -> a -> b
$ forall a. Eq a => a -> a -> Bool
(==)         (forall a. LensRelevance a => a -> Relevance
getRelevance Dom Type
dom1) (forall a. LensRelevance a => a -> Relevance
getRelevance Dom Type
dom2) -> m ()
errR
     | Bool -> Bool
not forall a b. (a -> b) -> a -> b
$ Quantity -> Quantity -> Bool
sameQuantity (forall a. LensQuantity a => a -> Quantity
getQuantity  Dom Type
dom1) (forall a. LensQuantity a => a -> Quantity
getQuantity  Dom Type
dom2) -> m ()
errQ
     | Bool -> Bool
not forall a b. (a -> b) -> a -> b
$ Cohesion -> Cohesion -> Bool
sameCohesion (forall a. LensCohesion a => a -> Cohesion
getCohesion  Dom Type
dom1) (forall a. LensCohesion a => a -> Cohesion
getCohesion  Dom Type
dom2) -> m ()
errC
     | Bool -> Bool
not forall a b. (a -> b) -> a -> b
$ forall t e. Dom' t e -> Bool
domIsFinite Dom Type
dom1 forall a. Eq a => a -> a -> Bool
== forall t e. Dom' t e -> Bool
domIsFinite Dom Type
dom2 -> m ()
errF
     | Bool
otherwise -> do
      let r :: Relevance
r = forall a. Ord a => a -> a -> a
max (forall a. LensRelevance a => a -> Relevance
getRelevance Dom Type
dom1) (forall a. LensRelevance a => a -> Relevance
getRelevance Dom Type
dom2)
              -- take "most irrelevant"
          dependent :: Bool
dependent = (Relevance
r forall a. Eq a => a -> a -> Bool
/= Relevance
Irrelevant) Bool -> Bool -> Bool
&& forall a. Free a => Abs a -> Bool
isBinderUsed Abs c
b2
      ProblemId
pid <- forall (m :: * -> *) a.
(MonadFresh ProblemId m, MonadConstraint m) =>
m a -> m ProblemId
newProblem_ forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *).
MonadConversion m =>
Comparison -> Type -> Type -> m ()
compareType Comparison
cmp0 Type
a1 Type
a2
      Dom Type
dom <- if Bool
dependent
             then (\ Type
a -> Dom Type
dom1 {unDom :: Type
unDom = Type
a}) forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> forall (m :: * -> *).
(MonadMetaSolver m, MonadFresh Int m) =>
Type -> ProblemId -> m Type
blockTypeOnProblem Type
a1 ProblemId
pid
             else forall (m :: * -> *) a. Monad m => a -> m a
return Dom Type
dom1
        -- We only need to require a1 == a2 if b2 is dependent
        -- If it's non-dependent it doesn't matter what we add to the context.
      let name :: VerboseKey
name = [Suggestion] -> VerboseKey
suggests [ forall a. Suggest a => a -> Suggestion
Suggestion Abs b
b1 , forall a. Suggest a => a -> Suggestion
Suggestion Abs c
b2 ]
      forall b (m :: * -> *) a.
(AddContext b, MonadAddContext m) =>
b -> m a -> m a
addContext (VerboseKey
name, Dom Type
dom) forall a b. (a -> b) -> a -> b
$ m ()
cont
      forall (m :: * -> *). MonadConstraint m => ProblemId -> m ()
stealConstraints ProblemId
pid
        -- Andreas, 2013-05-15 Now, comparison of codomains is not
        -- blocked any more by getting stuck on domains.
        -- Only the domain type in context will be blocked.
        -- But see issue #1258.

-- | When comparing argument spines (in compareElims) where the first arguments
--   don't match, we keep going, substituting the anti-unification of the two
--   terms in the telescope. More precisely:
--
--  @@
--    (u = v : A)[pid]   w = antiUnify pid A u v   us = vs : Δ[w/x]
--    -------------------------------------------------------------
--                    u us = v vs : (x : A) Δ
--  @@
--
--   The simplest case of anti-unification is to return a fresh metavariable
--   (created by blockTermOnProblem), but if there's shared structure between
--   the two terms we can expose that.
--
--   This is really a crutch that lets us get away with things that otherwise
--   would require heterogenous conversion checking. See for instance issue
--   #2384.
antiUnify :: MonadConversion m => ProblemId -> Type -> Term -> Term -> m Term
antiUnify :: forall (m :: * -> *).
MonadConversion m =>
ProblemId -> Type -> Term -> Term -> m Term
antiUnify ProblemId
pid Type
a Term
u Term
v = do
  forall a (m :: * -> *) b.
(Instantiate a, SynEq a, MonadReduce m) =>
a -> a -> (a -> a -> m b) -> (a -> a -> m b) -> m b
SynEq.checkSyntacticEquality Term
u Term
v (\Term
u Term
_ -> forall (m :: * -> *) a. Monad m => a -> m a
return Term
u) forall a b. (a -> b) -> a -> b
$ \Term
u Term
v -> do
  (Term
u, Term
v) <- forall a (m :: * -> *). (Reduce a, MonadReduce m) => a -> m a
reduce (Term
u, Term
v)
  forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"tc.conv.antiUnify" 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
"antiUnify"
    , TCMT IO Doc
"a =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Type
a
    , TCMT IO Doc
"u =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Term
u
    , TCMT IO Doc
"v =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Term
v
    ]
  case (Term
u, Term
v) of
    (Pi Dom Type
ua Abs Type
ub, Pi Dom Type
va Abs Type
vb) -> do
      Type
wa0 <- forall (m :: * -> *).
MonadConversion m =>
ProblemId -> Type -> Type -> m Type
antiUnifyType ProblemId
pid (forall t e. Dom' t e -> e
unDom Dom Type
ua) (forall t e. Dom' t e -> e
unDom Dom Type
va)
      let wa :: Dom Type
wa = Type
wa0 forall (f :: * -> *) a b. Functor f => a -> f b -> f a
<$ Dom Type
ua
      Type
wb <- forall b (m :: * -> *) a.
(AddContext b, MonadAddContext m) =>
b -> m a -> m a
addContext Dom Type
wa forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *).
MonadConversion m =>
ProblemId -> Type -> Type -> m Type
antiUnifyType ProblemId
pid (forall a. Subst a => Abs a -> a
absBody Abs Type
ub) (forall a. Subst a => Abs a -> a
absBody Abs Type
vb)
      forall (m :: * -> *) a. Monad m => a -> m a
return forall a b. (a -> b) -> a -> b
$ Dom Type -> Abs Type -> Term
Pi Dom Type
wa (forall a. (Subst a, Free a) => VerboseKey -> a -> Abs a
mkAbs (forall a. Abs a -> VerboseKey
absName Abs Type
ub) Type
wb)
    (Lam ArgInfo
i Abs Term
u, Lam ArgInfo
_ Abs Term
v) ->
      forall a (m :: * -> *). (Reduce a, MonadReduce m) => a -> m a
reduce (forall t a. Type'' t a -> a
unEl Type
a) forall (m :: * -> *) a b. Monad m => m a -> (a -> m b) -> m b
>>= \case
        Pi Dom Type
a Abs Type
b -> ArgInfo -> Abs Term -> Term
Lam ArgInfo
i forall b c a. (b -> c) -> (a -> b) -> a -> c
. (forall a. (Subst a, Free a) => VerboseKey -> a -> Abs a
mkAbs (forall a. Abs a -> VerboseKey
absName Abs Term
u)) forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> forall b (m :: * -> *) a.
(AddContext b, MonadAddContext m) =>
b -> m a -> m a
addContext Dom Type
a (forall (m :: * -> *).
MonadConversion m =>
ProblemId -> Type -> Term -> Term -> m Term
antiUnify ProblemId
pid (forall a. Subst a => Abs a -> a
absBody Abs Type
b) (forall a. Subst a => Abs a -> a
absBody Abs Term
u) (forall a. Subst a => Abs a -> a
absBody Abs Term
v))
        Term
_      -> m Term
fallback
    (Var Int
i Elims
us, Var Int
j Elims
vs) | Int
i forall a. Eq a => a -> a -> Bool
== Int
j -> m Term -> m Term
maybeGiveUp forall a b. (a -> b) -> a -> b
$ do
      Type
a <- forall (m :: * -> *).
(Applicative m, MonadFail m, MonadTCEnv m) =>
Int -> m Type
typeOfBV Int
i
      forall (m :: * -> *).
MonadConversion m =>
ProblemId -> Type -> Term -> Elims -> Elims -> m Term
antiUnifyElims ProblemId
pid Type
a (Int -> Term
var Int
i) Elims
us Elims
vs
    -- Andreas, 2017-07-27:
    -- It seems that nothing guarantees here that the constructors are fully
    -- applied!?  Thus, @a@ could be a function type and we need the robust
    -- @getConType@ here.
    -- (Note that @patternViolation@ swallows exceptions coming from @getConType@
    -- thus, we would not see clearly if we used @getFullyAppliedConType@ instead.)
    (Con ConHead
x ConInfo
ci Elims
us, Con ConHead
y ConInfo
_ Elims
vs) | ConHead
x forall a. Eq a => a -> a -> Bool
== ConHead
y -> m Term -> m Term
maybeGiveUp forall a b. (a -> b) -> a -> b
$ do
      Type
a <- forall b a. b -> (a -> b) -> Maybe a -> b
maybe forall {a}. m a
abort (forall (m :: * -> *) a. Monad m => a -> m a
return forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall a b. (a, b) -> b
snd) forall (m :: * -> *) a b. Monad m => (a -> m b) -> m a -> m b
=<< forall (m :: * -> *).
PureTCM m =>
ConHead -> Type -> m (Maybe ((QName, Type, [Arg Term]), Type))
getConType ConHead
x Type
a
      forall (m :: * -> *).
MonadConversion m =>
ProblemId -> Type -> Term -> Elims -> Elims -> m Term
antiUnifyElims ProblemId
pid Type
a (ConHead -> ConInfo -> Elims -> Term
Con ConHead
x ConInfo
ci []) Elims
us Elims
vs
    (Def QName
f [], Def QName
g []) | QName
f forall a. Eq a => a -> a -> Bool
== QName
g -> forall (m :: * -> *) a. Monad m => a -> m a
return (QName -> Elims -> Term
Def QName
f [])
    (Def QName
f Elims
us, Def QName
g Elims
vs) | QName
f forall a. Eq a => a -> a -> Bool
== QName
g, forall (t :: * -> *) a. Foldable t => t a -> Int
length Elims
us forall a. Eq a => a -> a -> Bool
== forall (t :: * -> *) a. Foldable t => t a -> Int
length Elims
vs -> m Term -> m Term
maybeGiveUp forall a b. (a -> b) -> a -> b
$ do
      Type
a <- forall (m :: * -> *).
MonadConversion m =>
QName -> Elims -> Elims -> m Type
computeElimHeadType QName
f Elims
us Elims
vs
      forall (m :: * -> *).
MonadConversion m =>
ProblemId -> Type -> Term -> Elims -> Elims -> m Term
antiUnifyElims ProblemId
pid Type
a (QName -> Elims -> Term
Def QName
f []) Elims
us Elims
vs
    (Term, Term)
_ -> m Term
fallback
  where
    maybeGiveUp :: m Term -> m Term
maybeGiveUp = forall (m :: * -> *) a.
MonadBlock m =>
(Blocker -> m a) -> m a -> m a
catchPatternErr forall a b. (a -> b) -> a -> b
$ \ Blocker
_ -> m Term
fallback
    abort :: m a
abort = forall (m :: * -> *) a. MonadBlock m => Blocker -> m a
patternViolation Blocker
neverUnblock -- caught by maybeGiveUp
    fallback :: m Term
fallback = forall (m :: * -> *).
(MonadMetaSolver m, MonadFresh Int m) =>
Type -> Term -> ProblemId -> m Term
blockTermOnProblem Type
a Term
u ProblemId
pid

antiUnifyArgs :: MonadConversion m => ProblemId -> Dom Type -> Arg Term -> Arg Term -> m (Arg Term)
antiUnifyArgs :: forall (m :: * -> *).
MonadConversion m =>
ProblemId -> Dom Type -> Arg Term -> Arg Term -> m (Arg Term)
antiUnifyArgs ProblemId
pid Dom Type
dom Arg Term
u Arg Term
v
  | Bool -> Bool
not (forall a b. (LensModality a, LensModality b) => a -> b -> Bool
sameModality (forall a. LensModality a => a -> Modality
getModality Arg Term
u) (forall a. LensModality a => a -> Modality
getModality Arg Term
v))
              = forall (m :: * -> *) a. MonadBlock m => Blocker -> m a
patternViolation Blocker
neverUnblock
  | Bool
otherwise = forall (tcm :: * -> *) m a.
(MonadTCEnv tcm, LensModality m) =>
m -> tcm a -> tcm a
applyModalityToContext Arg Term
u forall a b. (a -> b) -> a -> b
$
    forall (m :: * -> *) a. Monad m => m Bool -> m a -> m a -> m a
ifM (forall a (m :: * -> *).
(LensRelevance a, LensSort a, PrettyTCM a, PureTCM m,
 MonadBlock m) =>
a -> m Bool
isIrrelevantOrPropM Dom Type
dom)
    {-then-} (forall (m :: * -> *) a. Monad m => a -> m a
return Arg Term
u)
    {-else-} ((forall (f :: * -> *) a b. Functor f => a -> f b -> f a
<$ Arg Term
u) forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> forall (m :: * -> *).
MonadConversion m =>
ProblemId -> Type -> Term -> Term -> m Term
antiUnify ProblemId
pid (forall t e. Dom' t e -> e
unDom Dom Type
dom) (forall e. Arg e -> e
unArg Arg Term
u) (forall e. Arg e -> e
unArg Arg Term
v))

antiUnifyType :: MonadConversion m => ProblemId -> Type -> Type -> m Type
antiUnifyType :: forall (m :: * -> *).
MonadConversion m =>
ProblemId -> Type -> Type -> m Type
antiUnifyType ProblemId
pid (El Sort
s Term
a) (El Sort
_ Term
b) = forall (m :: * -> *) a.
(MonadTCEnv m, HasOptions m, MonadDebug m) =>
m a -> m a
workOnTypes forall a b. (a -> b) -> a -> b
$ forall t a. Sort' t -> a -> Type'' t a
El Sort
s forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> forall (m :: * -> *).
MonadConversion m =>
ProblemId -> Type -> Term -> Term -> m Term
antiUnify ProblemId
pid (Sort -> Type
sort Sort
s) Term
a Term
b

antiUnifyElims :: MonadConversion m => ProblemId -> Type -> Term -> Elims -> Elims -> m Term
antiUnifyElims :: forall (m :: * -> *).
MonadConversion m =>
ProblemId -> Type -> Term -> Elims -> Elims -> m Term
antiUnifyElims ProblemId
pid Type
a Term
self [] [] = forall (m :: * -> *) a. Monad m => a -> m a
return Term
self
antiUnifyElims ProblemId
pid Type
a Term
self (Proj ProjOrigin
o QName
f : Elims
es1) (Proj ProjOrigin
_ QName
g : Elims
es2) | QName
f forall a. Eq a => a -> a -> Bool
== QName
g = do
  Maybe (Dom Type, Term, Type)
res <- forall (m :: * -> *).
PureTCM m =>
Term
-> Type -> ProjOrigin -> QName -> m (Maybe (Dom Type, Term, Type))
projectTyped Term
self Type
a ProjOrigin
o QName
f
  case Maybe (Dom Type, Term, Type)
res of
    Just (Dom Type
_, Term
self, Type
a) -> forall (m :: * -> *).
MonadConversion m =>
ProblemId -> Type -> Term -> Elims -> Elims -> m Term
antiUnifyElims ProblemId
pid Type
a Term
self Elims
es1 Elims
es2
    Maybe (Dom Type, Term, Type)
Nothing -> forall (m :: * -> *) a. MonadBlock m => Blocker -> m a
patternViolation Blocker
neverUnblock -- can fail for projection like
antiUnifyElims ProblemId
pid Type
a Term
self (Apply Arg Term
u : Elims
es1) (Apply Arg Term
v : Elims
es2) = do
  forall a (m :: * -> *). (Reduce a, MonadReduce m) => a -> m a
reduce (forall t a. Type'' t a -> a
unEl Type
a) forall (m :: * -> *) a b. Monad m => m a -> (a -> m b) -> m b
>>= \case
    Pi Dom Type
a Abs Type
b -> do
      Arg Term
w <- forall (m :: * -> *).
MonadConversion m =>
ProblemId -> Dom Type -> Arg Term -> Arg Term -> m (Arg Term)
antiUnifyArgs ProblemId
pid Dom Type
a Arg Term
u Arg Term
v
      forall (m :: * -> *).
MonadConversion m =>
ProblemId -> Type -> Term -> Elims -> Elims -> m Term
antiUnifyElims ProblemId
pid (Abs Type
b forall a. Subst a => Abs a -> SubstArg a -> a
`lazyAbsApp` forall e. Arg e -> e
unArg Arg Term
w) (forall t. Apply t => t -> [Arg Term] -> t
apply Term
self [Arg Term
w]) Elims
es1 Elims
es2
    Term
_ -> forall (m :: * -> *) a. MonadBlock m => Blocker -> m a
patternViolation Blocker
neverUnblock
antiUnifyElims ProblemId
_ Type
_ Term
_ Elims
_ Elims
_ = forall (m :: * -> *) a. MonadBlock m => Blocker -> m a
patternViolation Blocker
neverUnblock -- trigger maybeGiveUp in antiUnify

-- | @compareElims pols a v els1 els2@ performs type-directed equality on eliminator spines.
--   @t@ is the type of the head @v@.
compareElims :: forall m. MonadConversion m => [Polarity] -> [IsForced] -> Type -> Term -> [Elim] -> [Elim] -> m ()
compareElims :: forall (m :: * -> *).
MonadConversion m =>
[Polarity] -> [IsForced] -> Type -> Term -> Elims -> Elims -> m ()
compareElims [Polarity]
pols0 [IsForced]
fors0 Type
a Term
v Elims
els01 Elims
els02 =
  forall (m :: * -> *) a.
MonadDebug m =>
VerboseKey -> Int -> VerboseKey -> m a -> m a
verboseBracket VerboseKey
"tc.conv.elim" Int
20 VerboseKey
"compareElims" forall a b. (a -> b) -> a -> b
$
  (forall (m :: * -> *).
MonadConstraint m =>
Constraint -> m () -> m ()
catchConstraint ([Polarity]
-> [IsForced] -> Type -> Term -> Elims -> Elims -> Constraint
ElimCmp [Polarity]
pols0 [IsForced]
fors0 Type
a Term
v Elims
els01 Elims
els02) :: m () -> m ()) forall a b. (a -> b) -> a -> b
$ do
  let v1 :: Term
v1 = forall t. Apply t => t -> Elims -> t
applyE Term
v Elims
els01
      v2 :: Term
v2 = forall t. Apply t => t -> Elims -> t
applyE Term
v Elims
els02
      failure :: m ()
failure = forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
TypeError -> m a
typeError forall a b. (a -> b) -> a -> b
$ Comparison -> Term -> Term -> CompareAs -> TypeError
UnequalTerms Comparison
CmpEq Term
v1 Term
v2 (Type -> CompareAs
AsTermsOf Type
a)
        -- Andreas, 2013-03-15 since one of the spines is empty, @a@
        -- is the correct type here.
  forall (f :: * -> *). Applicative f => Bool -> f () -> f ()
unless (forall (t :: * -> *) a. Foldable t => t a -> Bool
null Elims
els01) forall a b. (a -> b) -> a -> b
$ do
    forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"tc.conv.elim" Int
25 forall a b. (a -> b) -> a -> b
$ TCMT IO Doc
"compareElims" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
$$ do
     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
"a     =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Type
a
      , TCMT IO Doc
"pols0 (truncated to 10) =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
hsep (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
$ forall a. Int -> [a] -> [a]
take Int
10 [Polarity]
pols0)
      , TCMT IO Doc
"fors0 (truncated to 10) =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
hsep (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
$ forall a. Int -> [a] -> [a]
take Int
10 [IsForced]
fors0)
      , TCMT IO Doc
"v     =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Term
v
      , TCMT IO Doc
"els01 =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Elims
els01
      , TCMT IO Doc
"els02 =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Elims
els02
      ]
  case (Elims
els01, Elims
els02) of
    ([]         , []         ) -> forall (m :: * -> *) a. Monad m => a -> m a
return ()
    ([]         , Proj{}:Elims
_   ) -> m ()
failure -- not impossible, see issue 821
    (Proj{}  : Elims
_, []         ) -> m ()
failure -- could be x.p =?= x for projection p
    ([]         , Apply{} : Elims
_) -> m ()
failure -- not impossible, see issue 878
    (Apply{} : Elims
_, []         ) -> m ()
failure
    ([]         , IApply{} : Elims
_) -> m ()
failure
    (IApply{} : Elims
_, []         ) -> m ()
failure
    (Apply{} : Elims
_, Proj{}  : Elims
_) -> forall a. HasCallStack => a
__IMPOSSIBLE__ forall (f :: * -> *) a b. Functor f => a -> f b -> f a
<$ forall (m :: * -> *). MonadConstraint m => Bool -> m ()
solveAwakeConstraints' Bool
True -- NB: popped up in issue 889
    (Proj{}  : Elims
_, Apply{} : Elims
_) -> forall a. HasCallStack => a
__IMPOSSIBLE__ forall (f :: * -> *) a b. Functor f => a -> f b -> f a
<$ forall (m :: * -> *). MonadConstraint m => Bool -> m ()
solveAwakeConstraints' Bool
True -- but should be impossible (but again in issue 1467)
    (IApply{} : Elims
_, Proj{}  : Elims
_) -> forall a. HasCallStack => a
__IMPOSSIBLE__ forall (f :: * -> *) a b. Functor f => a -> f b -> f a
<$ forall (m :: * -> *). MonadConstraint m => Bool -> m ()
solveAwakeConstraints' Bool
True
    (Proj{}  : Elims
_, IApply{} : Elims
_) -> forall a. HasCallStack => a
__IMPOSSIBLE__ forall (f :: * -> *) a b. Functor f => a -> f b -> f a
<$ forall (m :: * -> *). MonadConstraint m => Bool -> m ()
solveAwakeConstraints' Bool
True
    (IApply{} : Elims
_, Apply{}  : Elims
_) -> forall a. HasCallStack => a
__IMPOSSIBLE__ forall (f :: * -> *) a b. Functor f => a -> f b -> f a
<$ forall (m :: * -> *). MonadConstraint m => Bool -> m ()
solveAwakeConstraints' Bool
True
    (Apply{}  : Elims
_, IApply{} : Elims
_) -> forall a. HasCallStack => a
__IMPOSSIBLE__ forall (f :: * -> *) a b. Functor f => a -> f b -> f a
<$ forall (m :: * -> *). MonadConstraint m => Bool -> m ()
solveAwakeConstraints' Bool
True
    (e :: Elim' Term
e@(IApply Term
x1 Term
y1 Term
r1) : Elims
els1, IApply Term
x2 Term
y2 Term
r2 : Elims
els2) -> do
      forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"tc.conv.elim" Int
25 forall a b. (a -> b) -> a -> b
$ TCMT IO Doc
"compareElims IApply"
       -- Andrea: copying stuff from the Apply case..
      let (Polarity
pol, [Polarity]
pols) = [Polarity] -> (Polarity, [Polarity])
nextPolarity [Polarity]
pols0
      Type
a  <- forall (m :: * -> *) t.
(MonadReduce m, MonadBlock m, IsMeta t, Reduce t) =>
t -> m t
abortIfBlocked Type
a
      PathView
va <- forall (m :: * -> *). HasBuiltins m => Type -> m PathView
pathView Type
a
      forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"tc.conv.elim.iapply" Int
60 forall a b. (a -> b) -> a -> b
$ TCMT IO Doc
"compareElims IApply" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
$$ do
        forall (m :: * -> *). Functor m => Int -> m Doc -> m Doc
nest Int
2 forall a b. (a -> b) -> a -> b
$ TCMT IO Doc
"va =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall (m :: * -> *). Applicative m => VerboseKey -> m Doc
text (forall a. Show a => a -> VerboseKey
show (PathView -> Bool
isPathType PathView
va))
      case PathView
va of
        PathType Sort
s QName
path Arg Term
l Arg Term
bA Arg Term
x Arg Term
y -> do
          Type
b <- forall (m :: * -> *).
(HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) =>
m Type
primIntervalType
          forall (m :: * -> *) a.
MonadConversion m =>
Polarity -> (Comparison -> a -> a -> m ()) -> a -> a -> m ()
compareWithPol Polarity
pol (forall a b c. (a -> b -> c) -> b -> a -> c
flip forall (m :: * -> *).
MonadConversion m =>
Comparison -> Type -> Term -> Term -> m ()
compareTerm Type
b)
                              Term
r1 Term
r2
          -- TODO: compare (x1,x2) and (y1,y2) ?
          let r :: Term
r = Term
r1 -- TODO Andrea:  do blocking
          Type
codom <- forall (m :: * -> *). Applicative m => m Term -> m Term -> m Type
el' (forall (f :: * -> *) a. Applicative f => a -> f a
pure forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall e. Arg e -> e
unArg forall a b. (a -> b) -> a -> b
$ Arg Term
l) ((forall (f :: * -> *) a. Applicative f => a -> f a
pure forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall e. Arg e -> e
unArg forall a b. (a -> b) -> a -> b
$ Arg Term
bA) forall (m :: * -> *). Applicative m => m Term -> m Term -> m Term
<@> forall (f :: * -> *) a. Applicative f => a -> f a
pure Term
r)
          forall (m :: * -> *).
MonadConversion m =>
[Polarity] -> [IsForced] -> Type -> Term -> Elims -> Elims -> m ()
compareElims [Polarity]
pols [] Type
codom -- Path non-dependent (codom `lazyAbsApp` unArg arg)
                            (forall t. Apply t => t -> Elims -> t
applyE Term
v [Elim' Term
e]) Elims
els1 Elims
els2
        -- We allow for functions (i : I) -> ... to also be heads of a IApply,
        -- because @etaContract@ can produce such terms
        OType t :: Type
t@(El Sort
_ Pi{}) -> forall (m :: * -> *).
MonadConversion m =>
[Polarity] -> [IsForced] -> Type -> Term -> Elims -> Elims -> m ()
compareElims [Polarity]
pols0 [IsForced]
fors0 Type
t Term
v (forall a. Arg a -> Elim' a
Apply (forall e. e -> Arg e
defaultArg Term
r1) forall a. a -> [a] -> [a]
: Elims
els1) (forall a. Arg a -> Elim' a
Apply (forall e. e -> Arg e
defaultArg Term
r2) forall a. a -> [a] -> [a]
: Elims
els2)

        OType Type
t -> forall (m :: * -> *) a. MonadBlock m => Blocker -> m a
patternViolation (forall t. AllMetas t => t -> Blocker
unblockOnAnyMetaIn Type
t) -- Can we get here? We know a is not blocked.

    (Apply Arg Term
arg1 : Elims
els1, Apply Arg Term
arg2 : Elims
els2) ->
      (forall (m :: * -> *) a.
MonadDebug m =>
VerboseKey -> Int -> VerboseKey -> m a -> m a
verboseBracket VerboseKey
"tc.conv.elim" Int
20 VerboseKey
"compare Apply" :: m () -> m ()) forall a b. (a -> b) -> a -> b
$ do
      forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"tc.conv.elim" Int
10 forall a b. (a -> b) -> a -> b
$ 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
"a    =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Type
a
        , TCMT IO Doc
"v    =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Term
v
        , TCMT IO Doc
"arg1 =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Arg Term
arg1
        , TCMT IO Doc
"arg2 =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Arg Term
arg2
        ]
      forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"tc.conv.elim" Int
50 forall a b. (a -> b) -> a -> b
$ 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
"raw:"
        , TCMT IO Doc
"a    =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall (m :: * -> *) a. (Applicative m, Pretty a) => a -> m Doc
pretty Type
a
        , TCMT IO Doc
"v    =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall (m :: * -> *) a. (Applicative m, Pretty a) => a -> m Doc
pretty Term
v
        , TCMT IO Doc
"arg1 =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall (m :: * -> *) a. (Applicative m, Pretty a) => a -> m Doc
pretty Arg Term
arg1
        , TCMT IO Doc
"arg2 =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall (m :: * -> *) a. (Applicative m, Pretty a) => a -> m Doc
pretty Arg Term
arg2
        ]
      let (Polarity
pol, [Polarity]
pols) = [Polarity] -> (Polarity, [Polarity])
nextPolarity [Polarity]
pols0
          (IsForced
for, [IsForced]
fors) = [IsForced] -> (IsForced, [IsForced])
nextIsForced [IsForced]
fors0
      Type
a <- forall (m :: * -> *) t.
(MonadReduce m, MonadBlock m, IsMeta t, Reduce t) =>
t -> m t
abortIfBlocked Type
a
      forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> VerboseKey -> m ()
reportSLn VerboseKey
"tc.conv.elim" Int
40 forall a b. (a -> b) -> a -> b
$ VerboseKey
"type is not blocked"
      case forall t a. Type'' t a -> a
unEl Type
a of
        (Pi (Dom{domInfo :: forall t e. Dom' t e -> ArgInfo
domInfo = ArgInfo
info, unDom :: forall t e. Dom' t e -> e
unDom = Type
b}) Abs Type
codom) -> do
          forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> VerboseKey -> m ()
reportSLn VerboseKey
"tc.conv.elim" Int
40 forall a b. (a -> b) -> a -> b
$ VerboseKey
"type is a function type"
          Maybe Term
mlvl <- forall e (m :: * -> *) a.
(MonadError e m, Functor m) =>
m a -> m (Maybe a)
tryMaybe forall (m :: * -> *).
(HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) =>
m Term
primLevel
          let freeInCoDom :: Abs a -> Bool
freeInCoDom (Abs VerboseKey
_ a
c) = Int
0 forall a. Free a => Int -> a -> Bool
`freeInIgnoringSorts` a
c
              freeInCoDom Abs a
_         = Bool
False
              dependent :: Bool
dependent = (forall a. a -> Maybe a
Just (forall t a. Type'' t a -> a
unEl Type
b) forall a. Eq a => a -> a -> Bool
/= Maybe Term
mlvl) Bool -> Bool -> Bool
&& forall a. Free a => Abs a -> Bool
freeInCoDom Abs Type
codom
                -- Level-polymorphism (x : Level) -> ... does not count as dependency here
                   -- NB: we could drop the free variable test and still be sound.
                   -- It is a trade-off between the administrative effort of
                   -- creating a blocking and traversing a term for free variables.
                   -- Apparently, it is believed that checking free vars is cheaper.
                   -- Andreas, 2013-05-15

-- NEW, Andreas, 2013-05-15

          -- compare arg1 and arg2
          ProblemId
pid <- forall (m :: * -> *) a.
(MonadFresh ProblemId m, MonadConstraint m) =>
m a -> m ProblemId
newProblem_ forall a b. (a -> b) -> a -> b
$ forall (tcm :: * -> *) m a.
(MonadTCEnv tcm, LensModality m) =>
m -> tcm a -> tcm a
applyModalityToContext ArgInfo
info forall a b. (a -> b) -> a -> b
$
              if IsForced -> Bool
isForced IsForced
for then
                forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> VerboseKey -> m ()
reportSLn VerboseKey
"tc.conv.elim" Int
40 forall a b. (a -> b) -> a -> b
$ VerboseKey
"argument is forced"
              else if forall a. LensRelevance a => a -> Bool
isIrrelevant ArgInfo
info then do
                forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> VerboseKey -> m ()
reportSLn VerboseKey
"tc.conv.elim" Int
40 forall a b. (a -> b) -> a -> b
$ VerboseKey
"argument is irrelevant"
                forall (m :: * -> *).
MonadConversion m =>
Type -> Term -> Term -> m ()
compareIrrelevant Type
b (forall e. Arg e -> e
unArg Arg Term
arg1) (forall e. Arg e -> e
unArg Arg Term
arg2)
              else do
                forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> VerboseKey -> m ()
reportSLn VerboseKey
"tc.conv.elim" Int
40 forall a b. (a -> b) -> a -> b
$ VerboseKey
"argument has polarity " forall a. [a] -> [a] -> [a]
++ forall a. Show a => a -> VerboseKey
show Polarity
pol
                forall (m :: * -> *) a.
MonadConversion m =>
Polarity -> (Comparison -> a -> a -> m ()) -> a -> a -> m ()
compareWithPol Polarity
pol (forall a b c. (a -> b -> c) -> b -> a -> c
flip forall (m :: * -> *).
MonadConversion m =>
Comparison -> Type -> Term -> Term -> m ()
compareTerm Type
b)
                  (forall e. Arg e -> e
unArg Arg Term
arg1) (forall e. Arg e -> e
unArg Arg Term
arg2)
          -- if comparison got stuck and function type is dependent, block arg
          Bool
solved <- forall (m :: * -> *).
(MonadTCEnv m, ReadTCState m) =>
ProblemId -> m Bool
isProblemSolved ProblemId
pid
          forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> VerboseKey -> m ()
reportSLn VerboseKey
"tc.conv.elim" Int
40 forall a b. (a -> b) -> a -> b
$ VerboseKey
"solved = " forall a. [a] -> [a] -> [a]
++ forall a. Show a => a -> VerboseKey
show Bool
solved
          Arg Term
arg <- if Bool
dependent Bool -> Bool -> Bool
&& Bool -> Bool
not Bool
solved
                 then forall (tcm :: * -> *) m a.
(MonadTCEnv tcm, LensModality m) =>
m -> tcm a -> tcm a
applyModalityToContext ArgInfo
info forall a b. (a -> b) -> a -> b
$ do
                  forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"tc.conv.elims" Int
50 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
"Trying antiUnify:"
                    , forall (m :: * -> *). Functor m => Int -> m Doc -> m Doc
nest Int
2 forall a b. (a -> b) -> a -> b
$ TCMT IO Doc
"b    =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Type
b
                    , forall (m :: * -> *). Functor m => Int -> m Doc -> m Doc
nest Int
2 forall a b. (a -> b) -> a -> b
$ TCMT IO Doc
"arg1 =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Arg Term
arg1
                    , forall (m :: * -> *). Functor m => Int -> m Doc -> m Doc
nest Int
2 forall a b. (a -> b) -> a -> b
$ TCMT IO Doc
"arg2 =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Arg Term
arg2
                    ]
                  Arg Term
arg <- (Arg Term
arg1 forall (f :: * -> *) a b. Functor f => f a -> b -> f b
$>) forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> forall (m :: * -> *).
MonadConversion m =>
ProblemId -> Type -> Term -> Term -> m Term
antiUnify ProblemId
pid Type
b (forall e. Arg e -> e
unArg Arg Term
arg1) (forall e. Arg e -> e
unArg Arg Term
arg2)
                  forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"tc.conv.elims" Int
50 forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *).
Applicative m =>
m Doc -> Int -> m Doc -> m Doc
hang TCMT IO Doc
"Anti-unification:" Int
2 (forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Arg Term
arg)
                  forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"tc.conv.elims" Int
70 forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *). Functor m => Int -> m Doc -> m Doc
nest Int
2 forall a b. (a -> b) -> a -> b
$ TCMT IO Doc
"raw:" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall (m :: * -> *) a. (Applicative m, Pretty a) => a -> m Doc
pretty Arg Term
arg
                  forall (m :: * -> *) a. Monad m => a -> m a
return Arg Term
arg
                 else forall (m :: * -> *) a. Monad m => a -> m a
return Arg Term
arg1
          -- continue, possibly with blocked instantiation
          forall (m :: * -> *).
MonadConversion m =>
[Polarity] -> [IsForced] -> Type -> Term -> Elims -> Elims -> m ()
compareElims [Polarity]
pols [IsForced]
fors (Abs Type
codom forall a. Subst a => Abs a -> SubstArg a -> a
`lazyAbsApp` forall e. Arg e -> e
unArg Arg Term
arg) (forall t. Apply t => t -> [Arg Term] -> t
apply Term
v [Arg Term
arg]) Elims
els1 Elims
els2
          -- any left over constraints of arg are associated to the comparison
          forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> VerboseKey -> m ()
reportSLn VerboseKey
"tc.conv.elim" Int
40 forall a b. (a -> b) -> a -> b
$ VerboseKey
"stealing constraints from problem " forall a. [a] -> [a] -> [a]
++ forall a. Show a => a -> VerboseKey
show ProblemId
pid
          forall (m :: * -> *). MonadConstraint m => ProblemId -> m ()
stealConstraints ProblemId
pid
          {- Stealing solves this issue:

             Does not create enough blocked tc-problems,
             see test/fail/DontPrune.
             (There are remaining problems which do not show up as yellow.)
             Need to find a way to associate pid also to result of compareElims.
          -}
        Term
a -> do
          forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"impossible" Int
10 forall a b. (a -> b) -> a -> b
$
            TCMT IO Doc
"unexpected type when comparing apply eliminations " forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Term
a
          forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"impossible" Int
50 forall a b. (a -> b) -> a -> b
$ TCMT IO Doc
"raw type:" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall (m :: * -> *) a. (Applicative m, Pretty a) => a -> m Doc
pretty Term
a
          forall (m :: * -> *) a. MonadBlock m => Blocker -> m a
patternViolation (forall t. AllMetas t => t -> Blocker
unblockOnAnyMetaIn Term
a)
          -- Andreas, 2013-10-22
          -- in case of disabled reductions (due to failing termination check)
          -- we might get stuck, so do not crash, but fail gently.
          -- __IMPOSSIBLE__

    -- case: f == f' are projections
    (Proj ProjOrigin
o QName
f : Elims
els1, Proj ProjOrigin
_ QName
f' : Elims
els2)
      | QName
f forall a. Eq a => a -> a -> Bool
/= QName
f'   -> forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
TypeError -> m a
typeError forall b c a. (b -> c) -> (a -> b) -> a -> c
. Doc -> TypeError
GenericDocError forall (m :: * -> *) a b. Monad m => (a -> m b) -> m a -> m b
=<< forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM QName
f forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> m Doc
"/=" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM QName
f'
      | Bool
otherwise -> do
        Type
a   <- forall (m :: * -> *) t.
(MonadReduce m, MonadBlock m, IsMeta t, Reduce t) =>
t -> m t
abortIfBlocked Type
a
        Maybe (Dom Type, Term, Type)
res <- forall (m :: * -> *).
PureTCM m =>
Term
-> Type -> ProjOrigin -> QName -> m (Maybe (Dom Type, Term, Type))
projectTyped Term
v Type
a ProjOrigin
o QName
f -- fails only if f is proj.like but parameters cannot be retrieved
        case Maybe (Dom Type, Term, Type)
res of
          Just (Dom Type
_, Term
u, Type
t) -> do
            -- Andreas, 2015-07-01:
            -- The arguments following the principal argument of a projection
            -- are invariant.  (At least as long as we have no explicit polarity
            -- annotations.)
            forall (m :: * -> *).
MonadConversion m =>
[Polarity] -> [IsForced] -> Type -> Term -> Elims -> Elims -> m ()
compareElims [] [] Type
t Term
u Elims
els1 Elims
els2
          Maybe (Dom Type, Term, Type)
Nothing -> do
            forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"tc.conv.elims" Int
30 forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
sep
              [ forall (m :: * -> *). Applicative m => VerboseKey -> m Doc
text forall a b. (a -> b) -> a -> b
$ VerboseKey
"projection " forall a. [a] -> [a] -> [a]
++ forall a. Pretty a => a -> VerboseKey
prettyShow QName
f
              , forall (m :: * -> *). Applicative m => VerboseKey -> m Doc
text   VerboseKey
"applied to value " 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 => VerboseKey -> m Doc
text   VerboseKey
"of unexpected type " forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Type
a
              ]
            forall (m :: * -> *) a. MonadBlock m => Blocker -> m a
patternViolation (forall t. AllMetas t => t -> Blocker
unblockOnAnyMetaIn Type
a)


-- | "Compare" two terms in irrelevant position.  This always succeeds.
--   However, we can dig for solutions of irrelevant metas in the
--   terms we compare.
--   (Certainly not the systematic solution, that'd be proof search...)
compareIrrelevant :: MonadConversion m => Type -> Term -> Term -> m ()
{- 2012-04-02 DontCare no longer present
compareIrrelevant t (DontCare v) w = compareIrrelevant t v w
compareIrrelevant t v (DontCare w) = compareIrrelevant t v w
-}
compareIrrelevant :: forall (m :: * -> *).
MonadConversion m =>
Type -> Term -> Term -> m ()
compareIrrelevant Type
t Term
v0 Term
w0 = do
  let v :: Term
v = Term -> Term
stripDontCare Term
v0
      w :: Term
w = Term -> Term
stripDontCare Term
w0
  forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"tc.conv.irr" Int
20 forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
vcat
    [ TCMT IO Doc
"compareIrrelevant"
    , forall (m :: * -> *). Functor m => Int -> m Doc -> m Doc
nest Int
2 forall a b. (a -> b) -> a -> b
$ TCMT IO Doc
"v =" 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 :: * -> *). Functor m => Int -> m Doc -> m Doc
nest Int
2 forall a b. (a -> b) -> a -> b
$ TCMT IO Doc
"w =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Term
w
    ]
  forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"tc.conv.irr" Int
50 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
$ TCMT IO Doc
"v =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall (m :: * -> *) a. (Applicative m, Pretty a) => a -> m Doc
pretty Term
v
    , forall (m :: * -> *). Functor m => Int -> m Doc -> m Doc
nest Int
2 forall a b. (a -> b) -> a -> b
$ TCMT IO Doc
"w =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall (m :: * -> *) a. (Applicative m, Pretty a) => a -> m Doc
pretty Term
w
    ]
  forall (m :: * -> *). MonadDebug m => ProfileOption -> m () -> m ()
whenProfile ProfileOption
Profile.Conversion forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *). MonadStatistics m => VerboseKey -> m ()
tick VerboseKey
"compare irrelevant"
  Term -> Term -> m () -> m ()
try Term
v Term
w forall a b. (a -> b) -> a -> b
$ Term -> Term -> m () -> m ()
try Term
w Term
v forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *) a. Monad m => a -> m a
return ()
  where
    try :: Term -> Term -> m () -> m ()
try (MetaV MetaId
x Elims
es) Term
w m ()
fallback = do
      MetaInstantiation
mi <- forall (m :: * -> *).
ReadTCState m =>
MetaId -> m MetaInstantiation
lookupMetaInstantiation MetaId
x
      Modality
mm <- forall (m :: * -> *). ReadTCState m => MetaId -> m Modality
lookupMetaModality MetaId
x
      let rel :: Relevance
rel  = forall a. LensRelevance a => a -> Relevance
getRelevance Modality
mm
          inst :: Bool
inst = case MetaInstantiation
mi of
                   InstV{} -> Bool
True
                   MetaInstantiation
_       -> Bool
False
      forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"tc.conv.irr" Int
20 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 => VerboseKey -> m Doc
text forall a b. (a -> b) -> a -> b
$ VerboseKey
"rel  = " forall a. [a] -> [a] -> [a]
++ forall a. Show a => a -> VerboseKey
show Relevance
rel
        , forall (m :: * -> *). Functor m => Int -> m Doc -> m Doc
nest Int
2 forall a b. (a -> b) -> a -> b
$ TCMT IO Doc
"inst =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall (m :: * -> *) a. (Applicative m, Pretty a) => a -> m Doc
pretty Bool
inst
        ]
      if Bool -> Bool
not (forall a. LensRelevance a => a -> Bool
isIrrelevant Relevance
rel) Bool -> Bool -> Bool
|| Bool
inst
        then m ()
fallback
        -- Andreas, 2016-08-08, issue #2131:
        -- Mining for solutions for irrelevant metas is not definite.
        -- Thus, in case of error, leave meta unsolved.
        else forall (m :: * -> *).
MonadConversion m =>
CompareDirection
-> MetaId
-> Elims
-> Term
-> CompareAs
-> (Term -> Term -> m ())
-> m ()
assignE CompareDirection
DirEq MetaId
x Elims
es Term
w (Type -> CompareAs
AsTermsOf Type
t) (forall (m :: * -> *).
MonadConversion m =>
Type -> Term -> Term -> m ()
compareIrrelevant Type
t) forall e (m :: * -> *) a.
MonadError e m =>
m a -> (e -> m a) -> m a
`catchError` \ TCErr
_ -> m ()
fallback
        -- the value of irrelevant or unused meta does not matter
    try Term
v Term
w m ()
fallback = m ()
fallback

compareWithPol :: MonadConversion m => Polarity -> (Comparison -> a -> a -> m ()) -> a -> a -> m ()
compareWithPol :: forall (m :: * -> *) a.
MonadConversion m =>
Polarity -> (Comparison -> a -> a -> m ()) -> a -> a -> m ()
compareWithPol Polarity
Invariant     Comparison -> a -> a -> m ()
cmp a
x a
y = Comparison -> a -> a -> m ()
cmp Comparison
CmpEq a
x a
y
compareWithPol Polarity
Covariant     Comparison -> a -> a -> m ()
cmp a
x a
y = Comparison -> a -> a -> m ()
cmp Comparison
CmpLeq a
x a
y
compareWithPol Polarity
Contravariant Comparison -> a -> a -> m ()
cmp a
x a
y = Comparison -> a -> a -> m ()
cmp Comparison
CmpLeq a
y a
x
compareWithPol Polarity
Nonvariant    Comparison -> a -> a -> m ()
cmp a
x a
y = forall (m :: * -> *) a. Monad m => a -> m a
return ()

polFromCmp :: Comparison -> Polarity
polFromCmp :: Comparison -> Polarity
polFromCmp Comparison
CmpLeq = Polarity
Covariant
polFromCmp Comparison
CmpEq  = Polarity
Invariant

-- | Type-directed equality on argument lists
--
compareArgs :: MonadConversion m => [Polarity] -> [IsForced] -> Type -> Term -> Args -> Args -> m ()
compareArgs :: forall (m :: * -> *).
MonadConversion m =>
[Polarity]
-> [IsForced] -> Type -> Term -> [Arg Term] -> [Arg Term] -> m ()
compareArgs [Polarity]
pol [IsForced]
for Type
a Term
v [Arg Term]
args1 [Arg Term]
args2 =
  forall (m :: * -> *).
MonadConversion m =>
[Polarity] -> [IsForced] -> Type -> Term -> Elims -> Elims -> m ()
compareElims [Polarity]
pol [IsForced]
for Type
a Term
v (forall a b. (a -> b) -> [a] -> [b]
map forall a. Arg a -> Elim' a
Apply [Arg Term]
args1) (forall a b. (a -> b) -> [a] -> [b]
map forall a. Arg a -> Elim' a
Apply [Arg Term]
args2)

---------------------------------------------------------------------------
-- * Types
---------------------------------------------------------------------------

-- | Equality on Types
compareType :: MonadConversion m => Comparison -> Type -> Type -> m ()
compareType :: forall (m :: * -> *).
MonadConversion m =>
Comparison -> Type -> Type -> m ()
compareType Comparison
cmp ty1 :: Type
ty1@(El Sort
s1 Term
a1) ty2 :: Type
ty2@(El Sort
s2 Term
a2) =
    forall (m :: * -> *) a.
(MonadTCEnv m, HasOptions m, MonadDebug m) =>
m a -> m a
workOnTypes forall a b. (a -> b) -> a -> b
$
    forall (m :: * -> *) a.
MonadDebug m =>
VerboseKey -> Int -> VerboseKey -> m a -> m a
verboseBracket VerboseKey
"tc.conv.type" Int
20 VerboseKey
"compareType" forall a b. (a -> b) -> a -> b
$ do
        forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"tc.conv.type" Int
50 forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
vcat
          [ TCMT IO Doc
"compareType" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
sep [ forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Type
ty1 forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Comparison
cmp
                                       , forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Type
ty2 ]
          , forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
hsep [ TCMT IO Doc
"   sorts:", forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Sort
s1, TCMT IO Doc
" and ", forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Sort
s2 ]
          ]
        forall (m :: * -> *).
MonadConversion m =>
Comparison -> CompareAs -> Term -> Term -> m ()
compareAs Comparison
cmp CompareAs
AsTypes Term
a1 Term
a2

leqType :: MonadConversion m => Type -> Type -> m ()
leqType :: forall (m :: * -> *). MonadConversion m => Type -> Type -> m ()
leqType = forall (m :: * -> *).
MonadConversion m =>
Comparison -> Type -> Type -> m ()
compareType Comparison
CmpLeq

-- | @coerce v a b@ coerces @v : a@ to type @b@, returning a @v' : b@
--   with maybe extra hidden applications or hidden abstractions.
--
--   In principle, this function can host coercive subtyping, but
--   currently it only tries to fix problems with hidden function types.
--
coerce :: (MonadConversion m, MonadTCM m) => Comparison -> Term -> Type -> Type -> m Term
coerce :: forall (m :: * -> *).
(MonadConversion m, MonadTCM m) =>
Comparison -> Term -> Type -> Type -> m Term
coerce Comparison
cmp Term
v Type
t1 Type
t2 = forall (m :: * -> *).
(MonadMetaSolver m, MonadConstraint m, MonadFresh Int m,
 MonadFresh ProblemId m) =>
Type -> m Term -> m Term
blockTerm Type
t2 forall a b. (a -> b) -> a -> b
$ do
  forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> m () -> m ()
verboseS VerboseKey
"tc.conv.coerce" Int
10 forall a b. (a -> b) -> a -> b
$ do
    (Type
a1,Type
a2) <- forall i (m :: * -> *).
(Reify i, MonadReify m) =>
i -> m (ReifiesTo i)
reify (Type
t1,Type
t2)
    let dbglvl :: Int
dbglvl = Int
30
    forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"tc.conv.coerce" Int
dbglvl forall a b. (a -> b) -> a -> b
$
      TCMT IO Doc
"coerce" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
vcat
        [ TCMT IO Doc
"term      v  =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Term
v
        , TCMT IO Doc
"from type t1 =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Type
a1
        , TCMT IO Doc
"to type   t2 =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Type
a2
        , TCMT IO Doc
"comparison   =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Comparison
cmp
        ]
    forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"tc.conv.coerce" Int
70 forall a b. (a -> b) -> a -> b
$
      TCMT IO Doc
"coerce" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
vcat
        [ TCMT IO Doc
"term      v  =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall (m :: * -> *) a. (Applicative m, Pretty a) => a -> m Doc
pretty Term
v
        , TCMT IO Doc
"from type t1 =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall (m :: * -> *) a. (Applicative m, Pretty a) => a -> m Doc
pretty Type
t1
        , TCMT IO Doc
"to type   t2 =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall (m :: * -> *) a. (Applicative m, Pretty a) => a -> m Doc
pretty Type
t2
        , TCMT IO Doc
"comparison   =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall (m :: * -> *) a. (Applicative m, Pretty a) => a -> m Doc
pretty Comparison
cmp
        ]
  -- v <$ do workOnTypes $ leqType t1 t2
  -- take off hidden/instance domains from t1 and t2
  TelV Telescope
tel1 Type
b1 <- forall (m :: * -> *).
(MonadReduce m, MonadAddContext m) =>
Int -> (Dom Type -> Bool) -> Type -> m (TelV Type)
telViewUpTo' (-Int
1) forall a. LensHiding a => a -> Bool
notVisible Type
t1
  TelV Telescope
tel2 Type
b2 <- forall (m :: * -> *).
(MonadReduce m, MonadAddContext m) =>
Int -> (Dom Type -> Bool) -> Type -> m (TelV Type)
telViewUpTo' (-Int
1) forall a. LensHiding a => a -> Bool
notVisible Type
t2
  let n :: Int
n = forall a. Sized a => a -> Int
size Telescope
tel1 forall a. Num a => a -> a -> a
- forall a. Sized a => a -> Int
size Telescope
tel2
  -- the crude solution would be
  --   v' = λ {tel2} → v {tel1}
  -- however, that may introduce unneccessary many function types
  -- If n  > 0 and b2 is not blocked, it is safe to
  -- insert n many hidden args
  if Int
n forall a. Ord a => a -> a -> Bool
<= Int
0 then m Term
fallback else do
    forall t (m :: * -> *) a.
(Reduce t, IsMeta t, MonadReduce m) =>
t -> (Blocker -> t -> m a) -> (NotBlocked -> t -> m a) -> m a
ifBlocked Type
b2 (\ Blocker
_ Type
_ -> m Term
fallback) forall a b. (a -> b) -> a -> b
$ \ NotBlocked
_ Type
_ -> do
      ([Arg Term]
args, Type
t1') <- forall (m :: * -> *).
(PureTCM m, MonadMetaSolver m, MonadTCM m) =>
Int -> (Hiding -> Bool) -> Type -> m ([Arg Term], Type)
implicitArgs Int
n forall a. LensHiding a => a -> Bool
notVisible Type
t1
      let v' :: Term
v' = Term
v forall t. Apply t => t -> [Arg Term] -> t
`apply` [Arg Term]
args
      Term
v' forall (f :: * -> *) a b. Functor f => a -> f b -> f a
<$ forall (m :: * -> *).
MonadConversion m =>
(Type -> Type -> m ()) -> Term -> Type -> Type -> m ()
coerceSize (forall (m :: * -> *).
MonadConversion m =>
Comparison -> Type -> Type -> m ()
compareType Comparison
cmp) Term
v' Type
t1' Type
t2
  where
    fallback :: m Term
fallback = Term
v forall (f :: * -> *) a b. Functor f => a -> f b -> f a
<$ forall (m :: * -> *).
MonadConversion m =>
(Type -> Type -> m ()) -> Term -> Type -> Type -> m ()
coerceSize (forall (m :: * -> *).
MonadConversion m =>
Comparison -> Type -> Type -> m ()
compareType Comparison
cmp) Term
v Type
t1 Type
t2

-- | Account for situations like @k : (Size< j) <= (Size< k + 1)@
--
--   Actually, the semantics is
--   @(Size<= k) ∩ (Size< j) ⊆ rhs@
--   which gives a disjunctive constraint.  Mmmh, looks like stuff
--   TODO.
--
--   For now, we do a cheap heuristics.
--
coerceSize :: MonadConversion m => (Type -> Type -> m ()) -> Term -> Type -> Type -> m ()
coerceSize :: forall (m :: * -> *).
MonadConversion m =>
(Type -> Type -> m ()) -> Term -> Type -> Type -> m ()
coerceSize Type -> Type -> m ()
leqType Term
v Type
t1 Type
t2 = forall (m :: * -> *) a.
MonadDebug m =>
VerboseKey -> Int -> VerboseKey -> m a -> m a
verboseBracket VerboseKey
"tc.conv.size.coerce" Int
45 VerboseKey
"coerceSize" forall a b. (a -> b) -> a -> b
$
  forall (m :: * -> *) a.
(MonadTCEnv m, HasOptions m, MonadDebug m) =>
m a -> m a
workOnTypes forall a b. (a -> b) -> a -> b
$ do
    forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"tc.conv.size.coerce" Int
70 forall a b. (a -> b) -> a -> b
$
      TCMT IO Doc
"coerceSize" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
vcat
        [ TCMT IO Doc
"term      v  =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall (m :: * -> *) a. (Applicative m, Pretty a) => a -> m Doc
pretty Term
v
        , TCMT IO Doc
"from type t1 =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall (m :: * -> *) a. (Applicative m, Pretty a) => a -> m Doc
pretty Type
t1
        , TCMT IO Doc
"to type   t2 =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall (m :: * -> *) a. (Applicative m, Pretty a) => a -> m Doc
pretty Type
t2
        ]
    let fallback :: m ()
fallback = Type -> Type -> m ()
leqType Type
t1 Type
t2
        done :: m ()
done = forall (m :: * -> *) a b.
Monad m =>
m (Maybe a) -> m b -> (a -> m b) -> m b
caseMaybeM (forall a (m :: * -> *).
(IsSizeType a, HasOptions m, HasBuiltins m) =>
a -> m (Maybe BoundedSize)
isSizeType forall (m :: * -> *) a b. Monad m => (a -> m b) -> m a -> m b
=<< forall a (m :: * -> *). (Reduce a, MonadReduce m) => a -> m a
reduce Type
t1) m ()
fallback forall a b. (a -> b) -> a -> b
$ \ BoundedSize
_ -> forall (m :: * -> *) a. Monad m => a -> m a
return ()
    -- Andreas, 2015-07-22, Issue 1615:
    -- If t1 is a meta and t2 a type like Size< v2, we need to make sure we do not miss
    -- the constraint v < v2!
    forall (m :: * -> *) a b.
Monad m =>
m (Maybe a) -> m b -> (a -> m b) -> m b
caseMaybeM (forall a (m :: * -> *).
(IsSizeType a, HasOptions m, HasBuiltins m) =>
a -> m (Maybe BoundedSize)
isSizeType forall (m :: * -> *) a b. Monad m => (a -> m b) -> m a -> m b
=<< forall a (m :: * -> *). (Reduce a, MonadReduce m) => a -> m a
reduce Type
t2) m ()
fallback forall a b. (a -> b) -> a -> b
$ \ BoundedSize
b2 -> do
      -- Andreas, 2017-01-20, issue #2329:
      -- If v is not a size suitable for the solver, like a neutral term,
      -- we can only rely on the type.
      SizeMaxView
mv <- forall (m :: * -> *). PureTCM m => Term -> m SizeMaxView
sizeMaxView Term
v
      if forall (t :: * -> *) a. Foldable t => (a -> Bool) -> t a -> Bool
any (\case{ DOtherSize{} -> Bool
True; DeepSizeView
_ -> Bool
False }) SizeMaxView
mv then m ()
fallback else do
      -- Andreas, 2015-02-11 do not instantiate metas here (triggers issue 1203).
      forall (m :: * -> *). Monad m => m Bool -> m () -> m ()
unlessM (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
$ Type -> Type -> m ()
leqType Type
t1 Type
t2) forall a b. (a -> b) -> a -> b
$ do
        -- A (most probably weaker) alternative is to just check syn.eq.
        -- ifM (snd <$> checkSyntacticEquality t1 t2) (return v) $ {- else -} do
        forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"tc.conv.size.coerce" Int
20 forall a b. (a -> b) -> a -> b
$ TCMT IO Doc
"coercing to a size type"
        case BoundedSize
b2 of
          -- @t2 = Size@.  We are done!
          BoundedSize
BoundedNo -> m ()
done
          -- @t2 = Size< v2@
          BoundedLt Term
v2 -> do
            SizeView
sv2 <- forall (m :: * -> *).
(HasBuiltins m, MonadTCEnv m, ReadTCState m) =>
Term -> m SizeView
sizeView Term
v2
            case SizeView
sv2 of
              SizeView
SizeInf     -> m ()
done
              OtherSize{} -> do
                -- Andreas, 2014-06-16:
                -- Issue 1203: For now, just treat v < v2 as suc v <= v2
                -- TODO: Need proper < comparison
                Term
vinc <- forall (m :: * -> *). HasBuiltins m => Int -> Term -> m Term
sizeSuc Int
1 Term
v
                forall (m :: * -> *).
MonadConversion m =>
Comparison -> Term -> Term -> m ()
compareSizes Comparison
CmpLeq Term
vinc Term
v2
                m ()
done
              -- @v2 = a2 + 1@: In this case, we can try @v <= a2@
              SizeSuc Term
a2 -> do
                forall (m :: * -> *).
MonadConversion m =>
Comparison -> Term -> Term -> m ()
compareSizes Comparison
CmpLeq Term
v Term
a2
                m ()
done  -- to pass Issue 1136

---------------------------------------------------------------------------
-- * Sorts and levels
---------------------------------------------------------------------------

compareLevel :: MonadConversion m => Comparison -> Level -> Level -> m ()
compareLevel :: forall (m :: * -> *).
MonadConversion m =>
Comparison -> Level -> Level -> m ()
compareLevel Comparison
CmpLeq Level
u Level
v = forall (m :: * -> *). MonadConversion m => Level -> Level -> m ()
leqLevel Level
u Level
v
compareLevel Comparison
CmpEq  Level
u Level
v = forall (m :: * -> *). MonadConversion m => Level -> Level -> m ()
equalLevel Level
u Level
v

compareSort :: MonadConversion m => Comparison -> Sort -> Sort -> m ()
compareSort :: forall (m :: * -> *).
MonadConversion m =>
Comparison -> Sort -> Sort -> m ()
compareSort Comparison
CmpEq  = forall (m :: * -> *). MonadConversion m => Sort -> Sort -> m ()
equalSort
compareSort Comparison
CmpLeq = forall (m :: * -> *). MonadConversion m => Sort -> Sort -> m ()
leqSort

-- | Check that the first sort is less or equal to the second.
--
--   We can put @SizeUniv@ below @Inf@, but otherwise, it is
--   unrelated to the other universes.
--
leqSort :: forall m. MonadConversion m => Sort -> Sort -> m ()
leqSort :: forall (m :: * -> *). MonadConversion m => Sort -> Sort -> m ()
leqSort Sort
s1 Sort
s2 = do
  forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"tc.conv.sort" Int
30 forall a b. (a -> b) -> a -> b
$
    forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
sep [ TCMT IO Doc
"leqSort"
        , 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
fsep [ forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Sort
s1 forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> TCMT IO Doc
"=<"
                        , forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Sort
s2 ]
        ]
  forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"tc.conv.sort" Int
60 forall a b. (a -> b) -> a -> b
$
    forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
sep [ TCMT IO Doc
"leqSort"
        , 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
fsep [ forall (m :: * -> *) a. (Applicative m, Pretty a) => a -> m Doc
pretty Sort
s1 forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> TCMT IO Doc
"=<"
                        , forall (m :: * -> *) a. (Applicative m, Pretty a) => a -> m Doc
pretty Sort
s2 ]
        ]
  forall (m :: * -> *). MonadDebug m => ProfileOption -> m () -> m ()
whenProfile ProfileOption
Profile.Conversion forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *). MonadStatistics m => VerboseKey -> m ()
tick VerboseKey
"compare sorts"

  forall a (m :: * -> *) b.
(Instantiate a, SynEq a, MonadReduce m) =>
a -> a -> (a -> a -> m b) -> (a -> a -> m b) -> m b
SynEq.checkSyntacticEquality Sort
s1 Sort
s2 (\Sort
_ Sort
_ -> forall (m :: * -> *) a. Monad m => a -> m a
return ()) forall a b. (a -> b) -> a -> b
$ \Sort
s1 Sort
s2 -> do

    Blocked Sort
s1b <- forall a (m :: * -> *).
(Reduce a, MonadReduce m) =>
a -> m (Blocked a)
reduceB Sort
s1
    Blocked Sort
s2b <- forall a (m :: * -> *).
(Reduce a, MonadReduce m) =>
a -> m (Blocked a)
reduceB Sort
s2

    let (Sort
s1,Sort
s2) = (forall t a. Blocked' t a -> a
ignoreBlocking Blocked Sort
s1b , forall t a. Blocked' t a -> a
ignoreBlocking Blocked Sort
s2b)
        blocker :: Blocker
blocker = Blocker -> Blocker -> Blocker
unblockOnEither (forall t a. Blocked' t a -> Blocker
getBlocker Blocked Sort
s1b) (forall t a. Blocked' t a -> Blocker
getBlocker Blocked Sort
s2b)
        postpone :: m ()
postpone = forall (m :: * -> *) a. MonadBlock m => Blocker -> m a
patternViolation Blocker
blocker

    let postponeIfBlocked :: m () -> m ()
postponeIfBlocked = forall (m :: * -> *) a.
MonadBlock m =>
(Blocker -> m a) -> m a -> m a
catchPatternErr forall a b. (a -> b) -> a -> b
$ \Blocker
blocker -> do
          if | Blocker
blocker forall a. Eq a => a -> a -> Bool
== Blocker
neverUnblock -> forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
TypeError -> m a
typeError forall a b. (a -> b) -> a -> b
$ Sort -> Sort -> TypeError
NotLeqSort Sort
s1 Sort
s2
             | Bool
otherwise -> do
                 forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"tc.conv.sort" 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
"Postponing constraint"
                   , 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
fsep [ forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Sort
s1 forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> TCMT IO Doc
"=<"
                                   , forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Sort
s2 ]
                   ]
                 forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"tc.conv.sort" Int
60 forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
vcat
                   [ TCMT IO Doc
"Postponing constraint"
                   , 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
fsep [ forall (m :: * -> *) a. (Applicative m, Pretty a) => a -> m Doc
pretty Sort
s1 forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> TCMT IO Doc
"=<"
                                   , forall (m :: * -> *) a. (Applicative m, Pretty a) => a -> m Doc
pretty Sort
s2 ]
                   ]
                 forall (m :: * -> *).
MonadConstraint m =>
Blocker -> Constraint -> m ()
addConstraint Blocker
blocker forall a b. (a -> b) -> a -> b
$ Comparison -> Sort -> Sort -> Constraint
SortCmp Comparison
CmpLeq Sort
s1 Sort
s2

    Bool
propEnabled <- forall (m :: * -> *). HasOptions m => m Bool
isPropEnabled
    Bool
typeInTypeEnabled <- forall (m :: * -> *). HasOptions m => m Bool
typeInType
    Bool
omegaInOmegaEnabled <- PragmaOptions -> Bool
optOmegaInOmega forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> forall (m :: * -> *). HasOptions m => m PragmaOptions
pragmaOptions

    let fvsRHS :: Int -> Bool
fvsRHS = (Int -> IntSet -> Bool
`IntSet.member` forall t. Free t => t -> IntSet
allFreeVars Sort
s2)
    Bool
badRigid <- Sort
s1 forall (m :: * -> *) a.
(PureTCM m, AnyRigid a) =>
a -> (Int -> Bool) -> m Bool
`rigidVarsNotContainedIn` Int -> Bool
fvsRHS

    m () -> m ()
postponeIfBlocked forall a b. (a -> b) -> a -> b
$ case (Sort
s1, Sort
s2) of
      -- Andreas, 2018-09-03: crash on dummy sort
      (DummyS VerboseKey
s, Sort
_) -> forall {m :: * -> *} {a} {b}.
(ReportS [a], MonadDebug m, IsString a) =>
a -> m b
impossibleSort VerboseKey
s
      (Sort
_, DummyS VerboseKey
s) -> forall {m :: * -> *} {a} {b}.
(ReportS [a], MonadDebug m, IsString a) =>
a -> m b
impossibleSort VerboseKey
s

      -- The most basic rule: @Set l =< Set l'@ iff @l =< l'@
      (Type Level
a  , Type Level
b  ) -> forall (m :: * -> *). MonadConversion m => Level -> Level -> m ()
leqLevel Level
a Level
b

      -- Likewise for @Prop@
      (Prop Level
a  , Prop Level
b  ) -> forall (m :: * -> *). MonadConversion m => Level -> Level -> m ()
leqLevel Level
a Level
b

      -- Likewise for @SSet@
      (SSet Level
a  , SSet Level
b  ) -> forall (m :: * -> *). MonadConversion m => Level -> Level -> m ()
leqLevel Level
a Level
b

      -- @Prop l@ is below @Set l@
      (Prop Level
a  , Type Level
b  ) -> forall (m :: * -> *). MonadConversion m => Level -> Level -> m ()
leqLevel Level
a Level
b
      (Type Level
a  , Prop Level
b  ) -> forall {a}. m a
no

      -- @Setωᵢ@ is above all small sorts (spelling out all cases
      -- for the exhaustiveness checker)
      (Inf IsFibrant
f Integer
m , Inf IsFibrant
f' Integer
n) ->
        if IsFibrant -> IsFibrant -> Bool
leqFib IsFibrant
f IsFibrant
f' Bool -> Bool -> Bool
&& (Integer
m forall a. Ord a => a -> a -> Bool
<= Integer
n Bool -> Bool -> Bool
|| Bool
typeInTypeEnabled Bool -> Bool -> Bool
|| Bool
omegaInOmegaEnabled) then m ()
yes else forall {a}. m a
no
      (Type{}  , Inf IsFibrant
f Integer
_) -> m ()
yes
      (Prop{}  , Inf IsFibrant
f Integer
_) -> m ()
yes
      (Inf IsFibrant
f Integer
_, Type{}  ) -> if IsFibrant
f forall a. Eq a => a -> a -> Bool
== IsFibrant
IsFibrant Bool -> Bool -> Bool
&& Bool
typeInTypeEnabled then m ()
yes else forall {a}. m a
no
      (Inf IsFibrant
f Integer
_, SSet{}  ) -> if IsFibrant
f forall a. Eq a => a -> a -> Bool
== IsFibrant
IsStrict  Bool -> Bool -> Bool
&& Bool
typeInTypeEnabled then m ()
yes else forall {a}. m a
no
      (Inf IsFibrant
_ Integer
_, Prop{}  ) -> forall {a}. m a
no

      -- @Set l@ is below @SSet l@
      (Type Level
a  , SSet Level
b  ) -> forall (m :: * -> *). MonadConversion m => Level -> Level -> m ()
leqLevel Level
a Level
b
      (SSet Level
a  , Type Level
b  ) -> forall {a}. m a
no

      -- @Prop l@ is below @SSet l@
      (Prop Level
a  , SSet Level
b  ) -> forall (m :: * -> *). MonadConversion m => Level -> Level -> m ()
leqLevel Level
a Level
b
      (SSet Level
a  , Prop Level
b  ) -> forall {a}. m a
no

      -- @SSet@ is below @SSetω@
      (SSet{}  , Inf IsFibrant
IsStrict Integer
_) -> m ()
yes
      (SSet{}  , Inf IsFibrant
IsFibrant Integer
_) -> forall {a}. m a
no

      -- @LockUniv@, @IntervalUniv@, @SizeUniv@, and @Prop0@ are bottom sorts.
      -- So is @Set0@ if @Prop@ is not enabled.
      (Sort
_       , Sort
LockUniv) -> forall (m :: * -> *). MonadConversion m => Sort -> Sort -> m ()
equalSort Sort
s1 Sort
s2
      (Sort
_       , Sort
IntervalUniv) -> forall (m :: * -> *). MonadConversion m => Sort -> Sort -> m ()
equalSort Sort
s1 Sort
s2
      (Sort
_       , Sort
SizeUniv) -> forall (m :: * -> *). MonadConversion m => Sort -> Sort -> m ()
equalSort Sort
s1 Sort
s2
      (Sort
_       , Prop (Max Integer
0 [])) -> forall (m :: * -> *). MonadConversion m => Sort -> Sort -> m ()
equalSort Sort
s1 Sort
s2
      (Sort
_       , Type (Max Integer
0 []))
        | Bool -> Bool
not Bool
propEnabled  -> forall (m :: * -> *). MonadConversion m => Sort -> Sort -> m ()
equalSort Sort
s1 Sort
s2

      -- @SizeUniv@ and @LockUniv@ are unrelated to any @Set l@ or @Prop l@
      (Sort
SizeUniv, Type{}  ) -> forall {a}. m a
no
      (Sort
SizeUniv, Prop{}  ) -> forall {a}. m a
no
      (Sort
SizeUniv , Inf{}  ) -> forall {a}. m a
no
      (Sort
SizeUniv, SSet{}  ) -> forall {a}. m a
no
      (Sort
LockUniv, Type{}  ) -> forall {a}. m a
no
      (Sort
LockUniv, Prop{}  ) -> forall {a}. m a
no
      (Sort
LockUniv , Inf{}  ) -> forall {a}. m a
no
      (Sort
LockUniv, SSet{}  ) -> forall {a}. m a
no

      -- @IntervalUniv@ is below @SSet l@, but not @Set l@ or @Prop l@
      (Sort
IntervalUniv, Type{}) -> forall {a}. m a
no
      (Sort
IntervalUniv, Prop{}) -> forall {a}. m a
no
      (Sort
IntervalUniv , Inf IsFibrant
IsStrict Integer
_) -> m ()
yes
      (Sort
IntervalUniv , Inf IsFibrant
IsFibrant Integer
_) -> forall {a}. m a
no
      (Sort
IntervalUniv , SSet Level
b) -> forall (m :: * -> *). MonadConversion m => Level -> Level -> m ()
leqLevel (Integer -> Level
ClosedLevel Integer
0) Level
b

      -- If the first sort is a small sort that rigidly depends on a
      -- variable and the second sort does not mention this variable,
      -- the second sort must be at least @Setω@.
      (Sort
_       , Sort
_       ) | Right (SmallSort IsFibrant
f) <- Sort -> Either Blocker SizeOfSort
sizeOfSort Sort
s1 , Bool
badRigid -> forall (m :: * -> *). MonadConversion m => Sort -> Sort -> m ()
leqSort (forall t. IsFibrant -> Integer -> Sort' t
Inf IsFibrant
f Integer
0) Sort
s2

      -- PiSort, FunSort, UnivSort and MetaS might reduce once we instantiate
      -- more metas, so we postpone.
      (PiSort{}, Sort
_       ) -> m ()
postpone
      (Sort
_       , PiSort{}) -> m ()
postpone
      (FunSort{}, Sort
_      ) -> m ()
postpone
      (Sort
_      , FunSort{}) -> m ()
postpone
      (UnivSort{}, Sort
_     ) -> m ()
postpone
      (Sort
_     , UnivSort{}) -> m ()
postpone
      (MetaS{} , Sort
_       ) -> m ()
postpone
      (Sort
_       , MetaS{} ) -> m ()
postpone

      -- DefS are postulated sorts, so they do not reduce.
      (DefS{} , Sort
_     ) -> forall {a}. m a
no
      (Sort
_      , DefS{}) -> forall {a}. m a
no

  where
  no :: m a
no  = forall (m :: * -> *) a. MonadBlock m => Blocker -> m a
patternViolation Blocker
neverUnblock
  yes :: m ()
yes = forall (m :: * -> *) a. Monad m => a -> m a
return ()

  leqFib :: IsFibrant -> IsFibrant -> Bool
leqFib IsFibrant
IsFibrant IsFibrant
_ = Bool
True
  leqFib IsFibrant
IsStrict IsFibrant
IsStrict = Bool
True
  leqFib IsFibrant
_ IsFibrant
_ = Bool
False
  impossibleSort :: a -> m b
impossibleSort a
s = do
    forall a (m :: * -> *).
(ReportS a, MonadDebug m) =>
VerboseKey -> Int -> a -> m ()
reportS VerboseKey
"impossible" Int
10
      [ a
"leqSort: found dummy sort with description:"
      , a
s
      ]
    forall a. HasCallStack => a
__IMPOSSIBLE__

leqLevel :: MonadConversion m => Level -> Level -> m ()
leqLevel :: forall (m :: * -> *). MonadConversion m => Level -> Level -> m ()
leqLevel Level
a Level
b = forall (m :: * -> *).
MonadConstraint m =>
Constraint -> m () -> m ()
catchConstraint (Comparison -> Level -> Level -> Constraint
LevelCmp Comparison
CmpLeq Level
a Level
b) forall a b. (a -> b) -> a -> b
$ do
      forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"tc.conv.level" Int
30 forall a b. (a -> b) -> a -> b
$
        TCMT IO Doc
"compareLevel" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+>
          forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
sep [ forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Level
a forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> TCMT IO Doc
"=<"
              , forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Level
b ]
      forall (m :: * -> *). MonadDebug m => ProfileOption -> m () -> m ()
whenProfile ProfileOption
Profile.Conversion forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *). MonadStatistics m => VerboseKey -> m ()
tick VerboseKey
"compare levels"

      (Level
a, Level
b) <- forall a (m :: * -> *). (Normalise a, MonadReduce m) => a -> m a
normalise (Level
a, Level
b)
      forall a (m :: * -> *) b.
(Instantiate a, SynEq a, MonadReduce m) =>
a -> a -> (a -> a -> m b) -> (a -> a -> m b) -> m b
SynEq.checkSyntacticEquality' Level
a Level
b
        (\Level
_ Level
_ ->
          forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"tc.conv.level" Int
60
            TCMT IO Doc
"checkSyntacticEquality returns True") forall a b. (a -> b) -> a -> b
$ \Level
a Level
b -> do
      forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"tc.conv.level" Int
60
        TCMT IO Doc
"checkSyntacticEquality returns False"

      let notok :: m ()
notok    = forall (m :: * -> *). Monad m => m Bool -> m () -> m ()
unlessM forall (m :: * -> *). HasOptions m => m Bool
typeInType forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
TypeError -> m a
typeError forall a b. (a -> b) -> a -> b
$ Sort -> Sort -> TypeError
NotLeqSort (forall t. Level' t -> Sort' t
Type Level
a) (forall t. Level' t -> Sort' t
Type Level
b)
          postpone :: m ()
postpone = forall (m :: * -> *) a. MonadBlock m => Blocker -> m a
patternViolation (forall t. AllMetas t => t -> Blocker
unblockOnAnyMetaIn (Level
a, Level
b))

          wrap :: m () -> m ()
wrap m ()
m = m ()
m forall e (m :: * -> *) a.
MonadError e m =>
m a -> (e -> m a) -> m a
`catchError` \case
            TypeError{} -> m ()
notok
            TCErr
err         -> forall e (m :: * -> *) a. MonadError e m => e -> m a
throwError TCErr
err

      Bool
cumulativity <- PragmaOptions -> Bool
optCumulativity forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> forall (m :: * -> *). HasOptions m => m PragmaOptions
pragmaOptions
      Bool
areWeComputingOverlap <- forall (m :: * -> *) a. MonadTCEnv m => Lens' a TCEnv -> m a
viewTC Lens' Bool TCEnv
eConflComputingOverlap
      forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"tc.conv.level" Int
40 forall a b. (a -> b) -> a -> b
$
        TCMT IO Doc
"compareLevelView" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+>
          forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
sep [ forall (m :: * -> *) (t :: * -> *).
(Applicative m, Semigroup (m Doc), Foldable t) =>
t (m Doc) -> m Doc
prettyList_ forall a b. (a -> b) -> a -> b
$ forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
fmap (forall (m :: * -> *) a. (Applicative m, Pretty a) => a -> m Doc
pretty forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall t. SingleLevel' t -> Level' t
unSingleLevel) forall a b. (a -> b) -> a -> b
$ forall t. Level' t -> List1 (SingleLevel' t)
levelMaxView Level
a
              , TCMT IO Doc
"=<"
              , forall (m :: * -> *) (t :: * -> *).
(Applicative m, Semigroup (m Doc), Foldable t) =>
t (m Doc) -> m Doc
prettyList_ forall a b. (a -> b) -> a -> b
$ forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
fmap (forall (m :: * -> *) a. (Applicative m, Pretty a) => a -> m Doc
pretty forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall t. SingleLevel' t -> Level' t
unSingleLevel) forall a b. (a -> b) -> a -> b
$ forall t. Level' t -> List1 (SingleLevel' t)
levelMaxView Level
b
              ]

      -- Extra reduce on level atoms, but should be cheap since they are already reduced.
      Level' (Blocked Term)
aB <- forall (t :: * -> *) (m :: * -> *) a b.
(Traversable t, Monad m) =>
(a -> m b) -> t a -> m (t b)
mapM forall a (m :: * -> *).
(Reduce a, MonadReduce m) =>
a -> m (Blocked a)
reduceB Level
a
      Level' (Blocked Term)
bB <- forall (t :: * -> *) (m :: * -> *) a b.
(Traversable t, Monad m) =>
(a -> m b) -> t a -> m (t b)
mapM forall a (m :: * -> *).
(Reduce a, MonadReduce m) =>
a -> m (Blocked a)
reduceB Level
b

      m () -> m ()
wrap forall a b. (a -> b) -> a -> b
$ case (forall t. Level' t -> List1 (SingleLevel' t)
levelMaxView Level' (Blocked Term)
aB, forall t. Level' t -> List1 (SingleLevel' t)
levelMaxView Level' (Blocked Term)
bB) of

        -- 0 ≤ any
        (SingleClosed Integer
0 :| [] , List1 (SingleLevel' (Blocked Term))
_) -> forall (m :: * -> *) a. Monad m => a -> m a
return ()

        -- any ≤ 0
        (List1 (SingleLevel' (Blocked Term))
as , SingleClosed Integer
0 :| []) ->
          forall (t :: * -> *) (m :: * -> *) a b.
(Foldable t, Monad m) =>
t a -> (a -> m b) -> m ()
forM_ List1 (SingleLevel' (Blocked Term))
as forall a b. (a -> b) -> a -> b
$ \ SingleLevel' (Blocked Term)
a' -> forall (m :: * -> *). MonadConversion m => Level -> Level -> m ()
equalLevel (forall t. SingleLevel' t -> Level' t
unSingleLevel forall a b. (a -> b) -> a -> b
$ forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
fmap forall t a. Blocked' t a -> a
ignoreBlocking SingleLevel' (Blocked Term)
a') (Integer -> Level
ClosedLevel Integer
0)

        -- closed ≤ closed
        (SingleClosed Integer
m :| [], SingleClosed Integer
n :| []) -> forall (f :: * -> *). Applicative f => Bool -> f () -> f ()
unless (Integer
m forall a. Ord a => a -> a -> Bool
<= Integer
n) m ()
notok

        -- closed ≤ b
        (SingleClosed Integer
m :| [] , List1 (SingleLevel' (Blocked Term))
_)
          | Integer
m forall a. Ord a => a -> a -> Bool
<= Level -> Integer
levelLowerBound Level
b -> forall (m :: * -> *) a. Monad m => a -> m a
return ()

        -- as ≤ neutral/closed
        (List1 (SingleLevel' (Blocked Term))
as, List1 (SingleLevel' (Blocked Term))
bs)
          | forall (t :: * -> *) a. Foldable t => (a -> Bool) -> t a -> Bool
all forall {t} {a}. SingleLevel' (Blocked' t a) -> Bool
neutralOrClosed List1 (SingleLevel' (Blocked Term))
bs , Level -> Integer
levelLowerBound Level
a forall a. Ord a => a -> a -> Bool
> Level -> Integer
levelLowerBound Level
b -> m ()
notok

        -- ⊔ as ≤ single
        (as :: List1 (SingleLevel' (Blocked Term))
as@(SingleLevel' (Blocked Term)
_:|SingleLevel' (Blocked Term)
_:[SingleLevel' (Blocked Term)]
_), SingleLevel' (Blocked Term)
b :| []) ->
          forall (t :: * -> *) (m :: * -> *) a b.
(Foldable t, Monad m) =>
t a -> (a -> m b) -> m ()
forM_ List1 (SingleLevel' (Blocked Term))
as forall a b. (a -> b) -> a -> b
$ \ SingleLevel' (Blocked Term)
a' -> forall (m :: * -> *). MonadConversion m => Level -> Level -> m ()
leqLevel (forall t. SingleLevel' t -> Level' t
unSingleLevel forall a b. (a -> b) -> a -> b
$ forall t a. Blocked' t a -> a
ignoreBlocking forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> SingleLevel' (Blocked Term)
a')
                                      (forall t. SingleLevel' t -> Level' t
unSingleLevel forall a b. (a -> b) -> a -> b
$ forall t a. Blocked' t a -> a
ignoreBlocking forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> SingleLevel' (Blocked Term)
b)

        -- reduce constants
        (List1 (SingleLevel' (Blocked Term))
as, List1 (SingleLevel' (Blocked Term))
bs)
          | let minN :: Integer
minN = forall a. Ord a => a -> a -> a
min (forall a b. (a, b) -> a
fst forall a b. (a -> b) -> a -> b
$ Level -> (Integer, Level)
levelPlusView Level
a) (forall a b. (a, b) -> a
fst forall a b. (a -> b) -> a -> b
$ Level -> (Integer, Level)
levelPlusView Level
b)
                a' :: Level
a'   = forall a. a -> Maybe a -> a
fromMaybe forall a. HasCallStack => a
__IMPOSSIBLE__ forall a b. (a -> b) -> a -> b
$ Integer -> Level -> Maybe Level
subLevel Integer
minN Level
a
                b' :: Level
b'   = forall a. a -> Maybe a -> a
fromMaybe forall a. HasCallStack => a
__IMPOSSIBLE__ forall a b. (a -> b) -> a -> b
$ Integer -> Level -> Maybe Level
subLevel Integer
minN Level
b
          , Integer
minN forall a. Ord a => a -> a -> Bool
> Integer
0 -> forall (m :: * -> *). MonadConversion m => Level -> Level -> m ()
leqLevel Level
a' Level
b'

        -- remove subsumed
        -- Andreas, 2014-04-07: This is ok if we do not go back to equalLevel
        (List1 (SingleLevel' (Blocked Term))
as, List1 (SingleLevel' (Blocked Term))
bs)
          | (subsumed :: [SingleLevel' (Blocked Term)]
subsumed@(SingleLevel' (Blocked Term)
_:[SingleLevel' (Blocked Term)]
_) , [SingleLevel' (Blocked Term)]
as') <- forall a. (a -> Bool) -> NonEmpty a -> ([a], [a])
List1.partition (SingleLevel' Term -> Bool
isSubsumed forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
fmap forall t a. Blocked' t a -> a
ignoreBlocking) List1 (SingleLevel' (Blocked Term))
as
          -> forall (m :: * -> *). MonadConversion m => Level -> Level -> m ()
leqLevel ([SingleLevel' Term] -> Level
unSingleLevels forall a b. (a -> b) -> a -> b
$ (forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
fmap forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
fmap) forall t a. Blocked' t a -> a
ignoreBlocking [SingleLevel' (Blocked Term)]
as') Level
b
          where
            isSubsumed :: SingleLevel' Term -> Bool
isSubsumed SingleLevel' Term
a = forall (t :: * -> *) a. Foldable t => (a -> Bool) -> t a -> Bool
any (SingleLevel' Term -> SingleLevel' Term -> Bool
`subsumes` SingleLevel' Term
a) forall a b. (a -> b) -> a -> b
$ (forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
fmap forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
fmap) forall t a. Blocked' t a -> a
ignoreBlocking List1 (SingleLevel' (Blocked Term))
bs

            subsumes :: SingleLevel -> SingleLevel -> Bool
            subsumes :: SingleLevel' Term -> SingleLevel' Term -> Bool
subsumes (SingleClosed Integer
m)        (SingleClosed Integer
n)        = Integer
m forall a. Ord a => a -> a -> Bool
>= Integer
n
            subsumes (SinglePlus (Plus Integer
m Term
_)) (SingleClosed Integer
n)        = Integer
m forall a. Ord a => a -> a -> Bool
>= Integer
n
            subsumes (SinglePlus (Plus Integer
m Term
a)) (SinglePlus (Plus Integer
n Term
b)) = Term
a forall a. Eq a => a -> a -> Bool
== Term
b Bool -> Bool -> Bool
&& Integer
m forall a. Ord a => a -> a -> Bool
>= Integer
n
            subsumes SingleLevel' Term
_ SingleLevel' Term
_ = Bool
False

        -- as ≤ _l x₁ .. xₙ ⊔ bs
        -- We can solve _l := λ x₁ .. xₙ -> as ⊔ (_l' x₁ .. xₙ)
        -- (where _l' is a new metavariable)
        (List1 (SingleLevel' (Blocked Term))
as , List1 (SingleLevel' (Blocked Term))
bs)
          | Bool
cumulativity
          , Bool -> Bool
not Bool
areWeComputingOverlap
          , Just (mb :: Term
mb@(MetaV MetaId
x Elims
es) , [SingleLevel' Term]
bs') <- [SingleLevel' Term] -> Maybe (Term, [SingleLevel' Term])
singleMetaView forall a b. (a -> b) -> a -> b
$ (forall a b. (a -> b) -> [a] -> [b]
map forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
fmap) forall t a. Blocked' t a -> a
ignoreBlocking (forall l. IsList l => l -> [Item l]
List1.toList List1 (SingleLevel' (Blocked Term))
bs)
          , forall (t :: * -> *) a. Foldable t => t a -> Bool
null [SingleLevel' Term]
bs' Bool -> Bool -> Bool
|| forall a. AllMetas a => a -> Bool
noMetas (Level -> Term
Level Level
a , [SingleLevel' Term] -> Level
unSingleLevels [SingleLevel' Term]
bs') -> do
            MetaVariable
mv <- forall (m :: * -> *).
(HasCallStack, MonadDebug m, ReadTCState m) =>
MetaId -> m MetaVariable
lookupLocalMeta MetaId
x
            -- Jesper, 2019-10-13: abort if this is an interaction
            -- meta or a generalizable meta
            Bool
abort <- (forall a. Maybe a -> Bool
isJust forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> forall (m :: * -> *).
ReadTCState m =>
MetaId -> m (Maybe InteractionId)
isInteractionMeta MetaId
x) forall (m :: * -> *). Monad m => m Bool -> m Bool -> m Bool
`or2M`
                     ((forall a. Eq a => a -> a -> Bool
== DoGeneralize
YesGeneralizeVar) forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> forall (m :: * -> *).
(HasCallStack, MonadDebug m, ReadTCState m) =>
MetaId -> m DoGeneralize
isGeneralizableMeta MetaId
x)
            if | Bool
abort -> m ()
postpone
               | Bool
otherwise -> do
                  MetaId
x' <- case MetaVariable -> Judgement MetaId
mvJudgement MetaVariable
mv of
                    IsSort{} -> forall a. HasCallStack => a
__IMPOSSIBLE__
                    HasType MetaId
_ Comparison
cmp Type
t -> do
                      TelV Telescope
tel Type
t' <- forall (m :: * -> *).
(MonadReduce m, MonadAddContext m) =>
Type -> m (TelV Type)
telView Type
t
                      forall (m :: * -> *) a.
MonadMetaSolver m =>
Frozen
-> MetaInfo
-> MetaPriority
-> Permutation
-> Judgement a
-> m MetaId
newMeta Frozen
Instantiable (MetaVariable -> MetaInfo
mvInfo MetaVariable
mv) MetaPriority
normalMetaPriority (Int -> Permutation
idP forall a b. (a -> b) -> a -> b
$ forall a. Sized a => a -> Int
size Telescope
tel) forall a b. (a -> b) -> a -> b
$ forall a. a -> Comparison -> Type -> Judgement a
HasType () Comparison
cmp Type
t
                  forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"tc.conv.level" Int
20 forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
fsep
                    [ TCMT IO Doc
"attempting to solve" , forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM (MetaId -> Elims -> Term
MetaV MetaId
x Elims
es) , TCMT IO Doc
"to the maximum of"
                    , forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM (Level -> Term
Level Level
a) , TCMT IO Doc
"and the fresh meta" , forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM (MetaId -> Elims -> Term
MetaV MetaId
x' Elims
es)
                    ]
                  forall (m :: * -> *). MonadConversion m => Level -> Level -> m ()
equalLevel (forall t. t -> Level' t
atomicLevel Term
mb) forall a b. (a -> b) -> a -> b
$ Level -> Level -> Level
levelLub Level
a (forall t. t -> Level' t
atomicLevel forall a b. (a -> b) -> a -> b
$ MetaId -> Elims -> Term
MetaV MetaId
x' Elims
es)


        -- Andreas, 2016-09-28: This simplification loses the solution lzero.
        -- Thus, it is invalid.
        -- See test/Succeed/LevelMetaLeqNeutralLevel.agda.
        -- -- [a] ≤ [neutral]
        -- ([a@(Plus n _)], [b@(Plus m NeutralLevel{})])
        --   | m == n -> equalLevel' (Max [a]) (Max [b])
        --   -- Andreas, 2014-04-07: This call to equalLevel is ok even if we removed
        --   -- subsumed terms from the lhs.

        -- anything else
        (List1 (SingleLevel' (Blocked Term)),
 List1 (SingleLevel' (Blocked Term)))
_ | forall a. AllMetas a => a -> Bool
noMetas (Level
a, Level
b) -> m ()
notok
          | Bool
otherwise      -> m ()
postpone
      where
        neutralOrClosed :: SingleLevel' (Blocked' t a) -> Bool
neutralOrClosed (SingleClosed Integer
_)                   = Bool
True
        neutralOrClosed (SinglePlus (Plus Integer
_ NotBlocked{})) = Bool
True
        neutralOrClosed SingleLevel' (Blocked' t a)
_                                  = Bool
False

        -- Is there exactly one @MetaV@ in the list of single levels?
        singleMetaView :: [SingleLevel] -> Maybe (Term, [SingleLevel])
        singleMetaView :: [SingleLevel' Term] -> Maybe (Term, [SingleLevel' Term])
singleMetaView (SinglePlus (Plus Integer
0 l :: Term
l@(MetaV MetaId
m Elims
es)) : [SingleLevel' Term]
ls)
          | forall (t :: * -> *) a. Foldable t => (a -> Bool) -> t a -> Bool
all (Bool -> Bool
not forall b c a. (b -> c) -> (a -> b) -> a -> c
. SingleLevel' Term -> Bool
isMetaLevel) [SingleLevel' Term]
ls = forall a. a -> Maybe a
Just (Term
l,[SingleLevel' Term]
ls)
        singleMetaView (SingleLevel' Term
l : [SingleLevel' Term]
ls)
          | Bool -> Bool
not forall a b. (a -> b) -> a -> b
$ SingleLevel' Term -> Bool
isMetaLevel SingleLevel' Term
l = forall (a :: * -> * -> *) b c d.
Arrow a =>
a b c -> a (d, b) (d, c)
second (SingleLevel' Term
lforall a. a -> [a] -> [a]
:) forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> [SingleLevel' Term] -> Maybe (Term, [SingleLevel' Term])
singleMetaView [SingleLevel' Term]
ls
        singleMetaView [SingleLevel' Term]
_ = forall a. Maybe a
Nothing

        isMetaLevel :: SingleLevel -> Bool
        isMetaLevel :: SingleLevel' Term -> Bool
isMetaLevel (SinglePlus (Plus Integer
_ MetaV{})) = Bool
True
        isMetaLevel SingleLevel' Term
_                             = Bool
False

equalLevel :: forall m. MonadConversion m => Level -> Level -> m ()
equalLevel :: forall (m :: * -> *). MonadConversion m => Level -> Level -> m ()
equalLevel Level
a Level
b = do
  forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"tc.conv.level" 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
"equalLevel", forall (m :: * -> *). Functor m => Int -> m Doc -> m Doc
nest Int
2 forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *). Functor m => m Doc -> m Doc
parens forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *) a. (Applicative m, Pretty a) => a -> m Doc
pretty Level
a, forall (m :: * -> *). Functor m => Int -> m Doc -> m Doc
nest Int
2 forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *). Functor m => m Doc -> m Doc
parens forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *) a. (Applicative m, Pretty a) => a -> m Doc
pretty Level
b ]
  forall (m :: * -> *). MonadDebug m => ProfileOption -> m () -> m ()
whenProfile ProfileOption
Profile.Conversion forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *). MonadStatistics m => VerboseKey -> m ()
tick VerboseKey
"compare levels"
  -- Andreas, 2013-10-31 remove common terms (that don't contain metas!)
  -- THAT's actually UNSOUND when metas are instantiated, because
  --     max a b == max a c  does not imply  b == c
  -- as <- return $ Set.fromList $ closed0 as
  -- bs <- return $ Set.fromList $ closed0 bs
  -- let cs = Set.filter (not . hasMeta) $ Set.intersection as bs
  -- as <- return $ Set.toList $ as Set.\\ cs
  -- bs <- return $ Set.toList $ bs Set.\\ cs

  forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"tc.conv.level" Int
40 forall a b. (a -> b) -> a -> b
$
    forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
sep [ TCMT IO Doc
"equalLevel"
        , 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 :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
sep [ forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Level
a forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> TCMT IO Doc
"=="
                              , forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Level
b
                              ]
               ]
        ]
  forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"tc.conv.level" Int
80 forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
sep [ TCMT IO Doc
"equalLevel", forall (m :: * -> *). Functor m => Int -> m Doc -> m Doc
nest Int
2 forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *). Functor m => m Doc -> m Doc
parens forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *) a. (Applicative m, Pretty a) => a -> m Doc
pretty Level
a, forall (m :: * -> *). Functor m => Int -> m Doc -> m Doc
nest Int
2 forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *). Functor m => m Doc -> m Doc
parens forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *) a. (Applicative m, Pretty a) => a -> m Doc
pretty Level
b ]

  (Level
a, Level
b) <- forall a (m :: * -> *). (Normalise a, MonadReduce m) => a -> m a
normalise (Level
a, Level
b)

  -- Jesper, 2014-02-02 remove terms that certainly do not contribute
  -- to the maximum
  let (Level
a', Level
b') = Level -> Level -> (Level, Level)
removeSubsumed Level
a Level
b

  forall a (m :: * -> *) b.
(Instantiate a, SynEq a, MonadReduce m) =>
a -> a -> (a -> a -> m b) -> (a -> a -> m b) -> m b
SynEq.checkSyntacticEquality' Level
a' Level
b'
    (\Level
_ Level
_ ->
      forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"tc.conv.level" Int
60
        TCMT IO Doc
"checkSyntacticEquality returns True") forall a b. (a -> b) -> a -> b
$ \Level
a Level
b -> do

  forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"tc.conv.level" Int
60 TCMT IO Doc
"checkSyntacticEquality returns False"

  let notok :: m ()
notok    = forall (m :: * -> *). Monad m => m Bool -> m () -> m ()
unlessM forall (m :: * -> *). HasOptions m => m Bool
typeInType m ()
notOk
      notOk :: m ()
notOk    = forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
TypeError -> m a
typeError forall a b. (a -> b) -> a -> b
$ Comparison -> Level -> Level -> TypeError
UnequalLevel Comparison
CmpEq Level
a' Level
b'
      postpone :: m ()
postpone = do
        forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"tc.conv.level" Int
30 forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *).
Applicative m =>
m Doc -> Int -> m Doc -> m Doc
hang TCMT IO Doc
"postponing:" Int
2 forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *).
Applicative m =>
m Doc -> Int -> m Doc -> m Doc
hang (forall (m :: * -> *) a. (Applicative m, Pretty a) => a -> m Doc
pretty Level
a' forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> TCMT IO Doc
"==") Int
0 (forall (m :: * -> *) a. (Applicative m, Pretty a) => a -> m Doc
pretty Level
b')
        Blocker
blocker <- forall t. AllMetas t => t -> Blocker
unblockOnAnyMetaIn forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> forall a (m :: * -> *).
(InstantiateFull a, MonadReduce m) =>
a -> m a
instantiateFull (Level
a', Level
b')
        forall (m :: * -> *) a. MonadBlock m => Blocker -> m a
patternViolation Blocker
blocker

  forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"tc.conv.level" 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
"equalLevel (w/o subsumed)"
        , 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 :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
sep [ forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Level
a' forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> TCMT IO Doc
"=="
                              , forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Level
b'
                              ]
               ]
        ]

  let as :: List1 (SingleLevel' Term)
as  = forall t. Level' t -> List1 (SingleLevel' t)
levelMaxView Level
a'
      bs :: List1 (SingleLevel' Term)
bs  = forall t. Level' t -> List1 (SingleLevel' t)
levelMaxView Level
b'
  forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"tc.conv.level" Int
50 forall a b. (a -> b) -> a -> b
$
    forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
sep [ forall (m :: * -> *). Applicative m => VerboseKey -> m Doc
text VerboseKey
"equalLevel"
        , 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 :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
sep [ forall (m :: * -> *) (t :: * -> *).
(Applicative m, Semigroup (m Doc), Foldable t) =>
t (m Doc) -> m Doc
prettyList_ forall a b. (a -> b) -> a -> b
$ forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
fmap (forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall t. SingleLevel' t -> Level' t
unSingleLevel) List1 (SingleLevel' Term)
as
                              , TCMT IO Doc
"=="
                              , forall (m :: * -> *) (t :: * -> *).
(Applicative m, Semigroup (m Doc), Foldable t) =>
t (m Doc) -> m Doc
prettyList_ forall a b. (a -> b) -> a -> b
$ forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
fmap (forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall t. SingleLevel' t -> Level' t
unSingleLevel) List1 (SingleLevel' Term)
bs
                              ]
               ]
        ]

  forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"tc.conv.level" Int
80 forall a b. (a -> b) -> a -> b
$
    forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
sep [ forall (m :: * -> *). Applicative m => VerboseKey -> m Doc
text VerboseKey
"equalLevel"
        , 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 :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
sep [ forall (m :: * -> *) (t :: * -> *).
(Applicative m, Semigroup (m Doc), Foldable t) =>
t (m Doc) -> m Doc
prettyList_ forall a b. (a -> b) -> a -> b
$ forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
fmap (forall (m :: * -> *) a. (Applicative m, Pretty a) => a -> m Doc
pretty forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall t. SingleLevel' t -> Level' t
unSingleLevel) List1 (SingleLevel' Term)
as
                              , TCMT IO Doc
"=="
                              , forall (m :: * -> *) (t :: * -> *).
(Applicative m, Semigroup (m Doc), Foldable t) =>
t (m Doc) -> m Doc
prettyList_ forall a b. (a -> b) -> a -> b
$ forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
fmap (forall (m :: * -> *) a. (Applicative m, Pretty a) => a -> m Doc
pretty forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall t. SingleLevel' t -> Level' t
unSingleLevel) List1 (SingleLevel' Term)
bs
                              ]
               ]
        ]

  -- Extra reduce on level atoms, but should be cheap since they are already reduced.
  List1 (SingleLevel' (Blocked Term))
as <- (forall (t :: * -> *) (m :: * -> *) a b.
(Traversable t, Monad m) =>
(a -> m b) -> t a -> m (t b)
mapM forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall (t :: * -> *) (m :: * -> *) a b.
(Traversable t, Monad m) =>
(a -> m b) -> t a -> m (t b)
mapM) forall a (m :: * -> *).
(Reduce a, MonadReduce m) =>
a -> m (Blocked a)
reduceB List1 (SingleLevel' Term)
as
  List1 (SingleLevel' (Blocked Term))
bs <- (forall (t :: * -> *) (m :: * -> *) a b.
(Traversable t, Monad m) =>
(a -> m b) -> t a -> m (t b)
mapM forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall (t :: * -> *) (m :: * -> *) a b.
(Traversable t, Monad m) =>
(a -> m b) -> t a -> m (t b)
mapM) forall a (m :: * -> *).
(Reduce a, MonadReduce m) =>
a -> m (Blocked a)
reduceB List1 (SingleLevel' Term)
bs

  forall (m :: * -> *).
MonadConstraint m =>
Constraint -> m () -> m ()
catchConstraint (Comparison -> Level -> Level -> Constraint
LevelCmp Comparison
CmpEq Level
a Level
b) forall a b. (a -> b) -> a -> b
$ case (List1 (SingleLevel' (Blocked Term))
as, List1 (SingleLevel' (Blocked Term))
bs) of

        -- closed == closed
        (SingleClosed Integer
m :| [], SingleClosed Integer
n :| [])
          | Integer
m forall a. Eq a => a -> a -> Bool
== Integer
n    -> forall (m :: * -> *) a. Monad m => a -> m a
return ()
          | Bool
otherwise -> m ()
notok

        -- closed == neutral
        (SingleClosed Integer
m :| [] , List1 (SingleLevel' (Blocked Term))
bs) | forall (t :: * -> *) a. Foldable t => (a -> Bool) -> t a -> Bool
any forall {t} {a}. SingleLevel' (Blocked' t a) -> Bool
isNeutral List1 (SingleLevel' (Blocked Term))
bs -> m ()
notok
        (List1 (SingleLevel' (Blocked Term))
as , SingleClosed Integer
n :| []) | forall (t :: * -> *) a. Foldable t => (a -> Bool) -> t a -> Bool
any forall {t} {a}. SingleLevel' (Blocked' t a) -> Bool
isNeutral List1 (SingleLevel' (Blocked Term))
as -> m ()
notok

        -- closed == b
        (SingleClosed Integer
m :| [] , List1 (SingleLevel' (Blocked Term))
_) | Integer
m forall a. Ord a => a -> a -> Bool
< Level -> Integer
levelLowerBound Level
b -> m ()
notok
        (List1 (SingleLevel' (Blocked Term))
_ , SingleClosed Integer
n :| []) | Integer
n forall a. Ord a => a -> a -> Bool
< Level -> Integer
levelLowerBound Level
a -> m ()
notok

        -- 0 == a ⊔ b
        (SingleClosed Integer
0 :| [] , bs :: List1 (SingleLevel' (Blocked Term))
bs@(SingleLevel' (Blocked Term)
_:|SingleLevel' (Blocked Term)
_:[SingleLevel' (Blocked Term)]
_)) ->
          forall (t :: * -> *) (m :: * -> *) a b.
(Foldable t, Monad m) =>
t a -> (a -> m b) -> m ()
forM_ List1 (SingleLevel' (Blocked Term))
bs forall a b. (a -> b) -> a -> b
$ \ SingleLevel' (Blocked Term)
b' ->  forall (m :: * -> *). MonadConversion m => Level -> Level -> m ()
equalLevel (Integer -> Level
ClosedLevel Integer
0) (forall t. SingleLevel' t -> Level' t
unSingleLevel forall a b. (a -> b) -> a -> b
$ forall t a. Blocked' t a -> a
ignoreBlocking forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> SingleLevel' (Blocked Term)
b')
        (as :: List1 (SingleLevel' (Blocked Term))
as@(SingleLevel' (Blocked Term)
_:|SingleLevel' (Blocked Term)
_:[SingleLevel' (Blocked Term)]
_) , SingleClosed Integer
0 :| []) ->
          forall (t :: * -> *) (m :: * -> *) a b.
(Foldable t, Monad m) =>
t a -> (a -> m b) -> m ()
forM_ List1 (SingleLevel' (Blocked Term))
as forall a b. (a -> b) -> a -> b
$ \ SingleLevel' (Blocked Term)
a' -> forall (m :: * -> *). MonadConversion m => Level -> Level -> m ()
equalLevel (forall t. SingleLevel' t -> Level' t
unSingleLevel forall a b. (a -> b) -> a -> b
$ forall t a. Blocked' t a -> a
ignoreBlocking forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> SingleLevel' (Blocked Term)
a') (Integer -> Level
ClosedLevel Integer
0)

        -- meta == any
        (SinglePlus (Plus Integer
k Blocked Term
a) :| [] , SinglePlus (Plus Integer
l Blocked Term
b) :| [])
          -- there is only a potential choice when k == l
          | MetaV MetaId
x Elims
as' <- forall t a. Blocked' t a -> a
ignoreBlocking Blocked Term
a
          , MetaV MetaId
y Elims
bs' <- forall t a. Blocked' t a -> a
ignoreBlocking Blocked Term
b
          , Integer
k forall a. Eq a => a -> a -> Bool
== Integer
l -> do
              Type
lvl <- forall (m :: * -> *). HasBuiltins m => m Type
levelType'
              forall (m :: * -> *).
MonadConversion m =>
Comparison
-> CompareAs -> MetaId -> Elims -> MetaId -> Elims -> m ()
compareMetas Comparison
CmpEq (Type -> CompareAs
AsTermsOf Type
lvl) MetaId
x Elims
as' MetaId
y Elims
bs'
        (SinglePlus (Plus Integer
k Blocked Term
a) :| [] , List1 (SingleLevel' (Blocked Term))
_)
          | MetaV MetaId
x Elims
as' <- forall t a. Blocked' t a -> a
ignoreBlocking Blocked Term
a
          , Just Level
b' <- Integer -> Level -> Maybe Level
subLevel Integer
k Level
b -> forall {m :: * -> *}.
(MonadMetaSolver m, MonadWarning m, MonadStatistics m,
 MonadFresh ProblemId m, MonadFresh Int m) =>
MetaId -> Elims -> Level -> m ()
meta MetaId
x Elims
as' Level
b'
        (List1 (SingleLevel' (Blocked Term))
_ , SinglePlus (Plus Integer
l Blocked Term
b) :| [])
          | MetaV MetaId
y Elims
bs' <- forall t a. Blocked' t a -> a
ignoreBlocking Blocked Term
b
          , Just Level
a' <- Integer -> Level -> Maybe Level
subLevel Integer
l Level
a -> forall {m :: * -> *}.
(MonadMetaSolver m, MonadWarning m, MonadStatistics m,
 MonadFresh ProblemId m, MonadFresh Int m) =>
MetaId -> Elims -> Level -> m ()
meta MetaId
y Elims
bs' Level
a'

        -- a' ⊔ b == b
        (List1 (SingleLevel' (Blocked Term)),
 List1 (SingleLevel' (Blocked Term)))
_ | Just Level
a' <- Level -> Level -> Maybe Level
levelMaxDiff Level
a Level
b
          , Level
b forall a. Eq a => a -> a -> Bool
/= Integer -> Level
ClosedLevel Integer
0 -> forall (m :: * -> *). MonadConversion m => Level -> Level -> m ()
leqLevel Level
a' Level
b

        -- a == b' ⊔ a
        (List1 (SingleLevel' (Blocked Term)),
 List1 (SingleLevel' (Blocked Term)))
_ | Just Level
b' <- Level -> Level -> Maybe Level
levelMaxDiff Level
b Level
a
          , Level
a forall a. Eq a => a -> a -> Bool
/= Integer -> Level
ClosedLevel Integer
0 -> forall (m :: * -> *). MonadConversion m => Level -> Level -> m ()
leqLevel Level
b' Level
a

        -- neutral/closed == neutral/closed
        (List1 (SingleLevel' (Blocked Term))
as , List1 (SingleLevel' (Blocked Term))
bs)
          | forall (t :: * -> *) a. Foldable t => (a -> Bool) -> t a -> Bool
all forall {t} {a}. SingleLevel' (Blocked' t a) -> Bool
isNeutralOrClosed (List1 (SingleLevel' (Blocked Term))
as forall a. Semigroup a => a -> a -> a
<> List1 (SingleLevel' (Blocked Term))
bs)
          -- Andreas, 2013-10-31: There could be metas in neutral levels (see Issue 930).
          -- Should not we postpone there as well?  Yes!
          , Bool -> Bool
not (forall (t :: * -> *) a. Foldable t => (a -> Bool) -> t a -> Bool
any forall {a} {t}. AllMetas a => SingleLevel' (Blocked' t a) -> Bool
hasMeta (List1 (SingleLevel' (Blocked Term))
as forall a. Semigroup a => a -> a -> a
<> List1 (SingleLevel' (Blocked Term))
bs))
          , forall (t :: * -> *) a. Foldable t => t a -> Int
length List1 (SingleLevel' (Blocked Term))
as forall a. Eq a => a -> a -> Bool
== forall (t :: * -> *) a. Foldable t => t a -> Int
length List1 (SingleLevel' (Blocked Term))
bs -> do
              forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> VerboseKey -> m ()
reportSLn VerboseKey
"tc.conv.level" Int
60 forall a b. (a -> b) -> a -> b
$ VerboseKey
"equalLevel: all are neutral or closed"
              forall (m :: * -> *) a b c.
Applicative m =>
(a -> b -> m c) -> List1 a -> List1 b -> m ()
List1.zipWithM_ (forall {m :: * -> *}.
(MonadMetaSolver m, MonadWarning m, MonadStatistics m,
 MonadFresh ProblemId m, MonadFresh Int m) =>
Term -> Term -> m ()
(===) forall b c a. (b -> b -> c) -> (a -> b) -> a -> a -> c
`on` Level -> Term
levelTm forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall t. SingleLevel' t -> Level' t
unSingleLevel forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
fmap forall t a. Blocked' t a -> a
ignoreBlocking) List1 (SingleLevel' (Blocked Term))
as List1 (SingleLevel' (Blocked Term))
bs

        -- more cases?
        (List1 (SingleLevel' (Blocked Term)),
 List1 (SingleLevel' (Blocked Term)))
_ | forall a. AllMetas a => a -> Bool
noMetas (Level
a , Level
b) -> m ()
notok
          | Bool
otherwise       -> m ()
postpone

      where
        Term
a === :: Term -> Term -> m ()
=== Term
b = forall (m :: * -> *). Monad m => m Bool -> m () -> m ()
unlessM forall (m :: * -> *). HasOptions m => m Bool
typeInType forall a b. (a -> b) -> a -> b
$ do
          Type
lvl <- forall (m :: * -> *). HasBuiltins m => m Type
levelType'
          forall (m :: * -> *).
MonadConversion m =>
CompareAs -> Term -> Term -> m ()
equalAtom (Type -> CompareAs
AsTermsOf Type
lvl) Term
a Term
b

        -- perform assignment (MetaV x as) := b
        meta :: MetaId -> Elims -> Level -> m ()
meta MetaId
x Elims
as Level
b = do
          forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> VerboseKey -> m ()
reportSLn VerboseKey
"tc.meta.level" Int
30 forall a b. (a -> b) -> a -> b
$ VerboseKey
"Assigning meta level"
          forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"tc.meta.level" Int
50 forall a b. (a -> b) -> a -> b
$ TCMT IO Doc
"meta" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
sep [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 (m :: * -> *) a. (Applicative m, Pretty a) => a -> m Doc
pretty Elims
as, forall (m :: * -> *) a. (Applicative m, Pretty a) => a -> m Doc
pretty Level
b]
          Type
lvl <- forall (m :: * -> *). HasBuiltins m => m Type
levelType'
          forall (m :: * -> *).
MonadConversion m =>
CompareDirection
-> MetaId
-> Elims
-> Term
-> CompareAs
-> (Term -> Term -> m ())
-> m ()
assignE CompareDirection
DirEq MetaId
x Elims
as (Level -> Term
levelTm Level
b) (Type -> CompareAs
AsTermsOf Type
lvl) forall {m :: * -> *}.
(MonadMetaSolver m, MonadWarning m, MonadStatistics m,
 MonadFresh ProblemId m, MonadFresh Int m) =>
Term -> Term -> m ()
(===) -- fallback: check equality as atoms

        isNeutral :: SingleLevel' (Blocked' t a) -> Bool
isNeutral (SinglePlus (Plus Integer
_ NotBlocked{})) = Bool
True
        isNeutral SingleLevel' (Blocked' t a)
_                                  = Bool
False

        isNeutralOrClosed :: SingleLevel' (Blocked' t a) -> Bool
isNeutralOrClosed (SingleClosed Integer
_)                   = Bool
True
        isNeutralOrClosed (SinglePlus (Plus Integer
_ NotBlocked{})) = Bool
True
        isNeutralOrClosed SingleLevel' (Blocked' t a)
_                                  = Bool
False

        hasMeta :: SingleLevel' (Blocked' t a) -> Bool
hasMeta (SinglePlus (Plus Integer
_ Blocked{})) = Bool
True
        hasMeta (SinglePlus (Plus Integer
_ Blocked' t a
a))         = forall a. Maybe a -> Bool
isJust forall a b. (a -> b) -> a -> b
$ forall a. AllMetas a => a -> Maybe MetaId
firstMeta forall a b. (a -> b) -> a -> b
$ forall t a. Blocked' t a -> a
ignoreBlocking Blocked' t a
a
        hasMeta (SingleClosed Integer
_)                = Bool
False

        removeSubsumed :: Level -> Level -> (Level, Level)
removeSubsumed Level
a Level
b =
          let as :: [Item (List1 (SingleLevel' Term))]
as = forall l. IsList l => l -> [Item l]
List1.toList forall a b. (a -> b) -> a -> b
$ forall t. Level' t -> List1 (SingleLevel' t)
levelMaxView Level
a
              bs :: [Item (List1 (SingleLevel' Term))]
bs = forall l. IsList l => l -> [Item l]
List1.toList forall a b. (a -> b) -> a -> b
$ forall t. Level' t -> List1 (SingleLevel' t)
levelMaxView Level
b
              a' :: Level
a' = [SingleLevel' Term] -> Level
unSingleLevels forall a b. (a -> b) -> a -> b
$ forall a. (a -> Bool) -> [a] -> [a]
filter (Bool -> Bool
not forall b c a. (b -> c) -> (a -> b) -> a -> c
. (forall {t :: * -> *} {a}.
(Foldable t, Eq a) =>
SingleLevel' a -> t (SingleLevel' a) -> Bool
`isStrictlySubsumedBy` [Item (List1 (SingleLevel' Term))]
bs)) [Item (List1 (SingleLevel' Term))]
as
              b' :: Level
b' = [SingleLevel' Term] -> Level
unSingleLevels forall a b. (a -> b) -> a -> b
$ forall a. (a -> Bool) -> [a] -> [a]
filter (Bool -> Bool
not forall b c a. (b -> c) -> (a -> b) -> a -> c
. (forall {t :: * -> *} {a}.
(Foldable t, Eq a) =>
SingleLevel' a -> t (SingleLevel' a) -> Bool
`isStrictlySubsumedBy` [Item (List1 (SingleLevel' Term))]
as)) [Item (List1 (SingleLevel' Term))]
bs
          in (Level
a',Level
b')

        SingleLevel' a
x isStrictlySubsumedBy :: SingleLevel' a -> t (SingleLevel' a) -> Bool
`isStrictlySubsumedBy` t (SingleLevel' a)
ys = forall (t :: * -> *) a. Foldable t => (a -> Bool) -> t a -> Bool
any (forall {a}. Eq a => SingleLevel' a -> SingleLevel' a -> Bool
`strictlySubsumes` SingleLevel' a
x) t (SingleLevel' a)
ys

        SingleClosed Integer
m        strictlySubsumes :: SingleLevel' a -> SingleLevel' a -> Bool
`strictlySubsumes` SingleClosed Integer
n        = Integer
m forall a. Ord a => a -> a -> Bool
> Integer
n
        SinglePlus (Plus Integer
m a
a) `strictlySubsumes` SingleClosed Integer
n        = Integer
m forall a. Ord a => a -> a -> Bool
> Integer
n
        SinglePlus (Plus Integer
m a
a) `strictlySubsumes` SinglePlus (Plus Integer
n a
b) = a
a forall a. Eq a => a -> a -> Bool
== a
b Bool -> Bool -> Bool
&& Integer
m forall a. Ord a => a -> a -> Bool
> Integer
n
        SingleLevel' a
_                     `strictlySubsumes` SingleLevel' a
_                     = Bool
False


-- | Check that the first sort equal to the second.
equalSort :: forall m. MonadConversion m => Sort -> Sort -> m ()
equalSort :: forall (m :: * -> *). MonadConversion m => Sort -> Sort -> m ()
equalSort Sort
s1 Sort
s2 = do
  forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"tc.conv.sort" Int
30 forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
sep
    [ TCMT IO Doc
"equalSort"
    , 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 :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
fsep [ forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Sort
s1 forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> TCMT IO Doc
"=="
                           , forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Sort
s2 ]
           ]
    ]
  forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"tc.conv.sort" Int
60 forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
sep
    [ TCMT IO Doc
"equalSort"
    , 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 :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
fsep [ forall (m :: * -> *) a. (Applicative m, Pretty a) => a -> m Doc
pretty Sort
s1 forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> TCMT IO Doc
"=="
                           , forall (m :: * -> *) a. (Applicative m, Pretty a) => a -> m Doc
pretty Sort
s2 ]
           ]
    ]
  forall (m :: * -> *). MonadDebug m => ProfileOption -> m () -> m ()
whenProfile ProfileOption
Profile.Conversion forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *). MonadStatistics m => VerboseKey -> m ()
tick VerboseKey
"compare sorts"

  forall a (m :: * -> *) b.
(Instantiate a, SynEq a, MonadReduce m) =>
a -> a -> (a -> a -> m b) -> (a -> a -> m b) -> m b
SynEq.checkSyntacticEquality Sort
s1 Sort
s2 (\Sort
_ Sort
_ -> forall (m :: * -> *) a. Monad m => a -> m a
return ()) forall a b. (a -> b) -> a -> b
$ \Sort
s1 Sort
s2 -> do

    Blocked Sort
s1b <- forall a (m :: * -> *).
(Reduce a, MonadReduce m) =>
a -> m (Blocked a)
reduceB Sort
s1
    Blocked Sort
s2b <- forall a (m :: * -> *).
(Reduce a, MonadReduce m) =>
a -> m (Blocked a)
reduceB Sort
s2

    let (Sort
s1,Sort
s2) = (forall t a. Blocked' t a -> a
ignoreBlocking Blocked Sort
s1b, forall t a. Blocked' t a -> a
ignoreBlocking Blocked Sort
s2b)
        blocker :: Blocker
blocker = Blocker -> Blocker -> Blocker
unblockOnEither (forall t a. Blocked' t a -> Blocker
getBlocker Blocked Sort
s1b) (forall t a. Blocked' t a -> Blocker
getBlocker Blocked Sort
s2b)

    let postponeIfBlocked :: m () -> m ()
postponeIfBlocked = forall (m :: * -> *) a.
MonadBlock m =>
(Blocker -> m a) -> m a -> m a
catchPatternErr forall a b. (a -> b) -> a -> b
$ \Blocker
blocker ->
          if | Blocker
blocker forall a. Eq a => a -> a -> Bool
== Blocker
neverUnblock -> forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
TypeError -> m a
typeError forall a b. (a -> b) -> a -> b
$ Sort -> Sort -> TypeError
UnequalSorts Sort
s1 Sort
s2
             | Bool
otherwise -> do
                 forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"tc.conv.sort" 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
"Postponing constraint"
                   , 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
fsep [ forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Sort
s1 forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> TCMT IO Doc
"=="
                                   , forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Sort
s2 ]
                   ]
                 forall (m :: * -> *).
MonadConstraint m =>
Blocker -> Constraint -> m ()
addConstraint Blocker
blocker forall a b. (a -> b) -> a -> b
$ Comparison -> Sort -> Sort -> Constraint
SortCmp Comparison
CmpEq Sort
s1 Sort
s2

    Bool
propEnabled <- forall (m :: * -> *). HasOptions m => m Bool
isPropEnabled
    Bool
typeInTypeEnabled <- forall (m :: * -> *). HasOptions m => m Bool
typeInType
    Bool
omegaInOmegaEnabled <- PragmaOptions -> Bool
optOmegaInOmega forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> forall (m :: * -> *). HasOptions m => m PragmaOptions
pragmaOptions

    m () -> m ()
postponeIfBlocked forall a b. (a -> b) -> a -> b
$ case (Sort
s1, Sort
s2) of

            -- Andreas, 2018-09-03: crash on dummy sort
            (DummyS VerboseKey
s, Sort
_) -> forall {m :: * -> *} {a} {b}.
(ReportS [a], MonadDebug m, IsString a) =>
a -> m b
impossibleSort VerboseKey
s
            (Sort
_, DummyS VerboseKey
s) -> forall {m :: * -> *} {a} {b}.
(ReportS [a], MonadDebug m, IsString a) =>
a -> m b
impossibleSort VerboseKey
s

            -- one side is a meta sort: try to instantiate
            -- In case both sides are meta sorts, instantiate the
            -- bigger (i.e. more recent) one.
            (MetaS MetaId
x Elims
es , MetaS MetaId
y Elims
es') -> forall (m :: * -> *).
MonadConversion m =>
Comparison
-> CompareAs -> MetaId -> Elims -> MetaId -> Elims -> m ()
compareMetas Comparison
CmpEq CompareAs
AsTypes MetaId
x Elims
es MetaId
y Elims
es'
            (MetaS MetaId
x Elims
es , Sort
_          ) -> MetaId -> Elims -> Sort -> m ()
meta MetaId
x Elims
es Sort
s2
            (Sort
_          , MetaS MetaId
x Elims
es ) -> MetaId -> Elims -> Sort -> m ()
meta MetaId
x Elims
es Sort
s1

            -- diagonal cases for rigid sorts
            (Type Level
a     , Type Level
b     ) -> forall (m :: * -> *). MonadConversion m => Level -> Level -> m ()
equalLevel Level
a Level
b forall {m :: * -> *} {a}. MonadError TCErr m => m a -> m a -> m a
`catchInequalLevel` forall {a}. m a
no
            (Sort
SizeUniv   , Sort
SizeUniv   ) -> m ()
yes
            (Sort
LockUniv   , Sort
LockUniv   ) -> m ()
yes
            (Sort
IntervalUniv , Sort
IntervalUniv) -> m ()
yes
            (Prop Level
a     , Prop Level
b     ) -> forall (m :: * -> *). MonadConversion m => Level -> Level -> m ()
equalLevel Level
a Level
b forall {m :: * -> *} {a}. MonadError TCErr m => m a -> m a -> m a
`catchInequalLevel` forall {a}. m a
no
            (Inf IsFibrant
f Integer
m    , Inf IsFibrant
f' Integer
n   ) ->
              if IsFibrant
f forall a. Eq a => a -> a -> Bool
== IsFibrant
f' Bool -> Bool -> Bool
&& (Integer
m forall a. Eq a => a -> a -> Bool
== Integer
n Bool -> Bool -> Bool
|| Bool
typeInTypeEnabled Bool -> Bool -> Bool
|| Bool
omegaInOmegaEnabled) then m ()
yes else forall {a}. m a
no
            (SSet Level
a     , SSet Level
b     ) -> forall (m :: * -> *). MonadConversion m => Level -> Level -> m ()
equalLevel Level
a Level
b

            -- if --type-in-type is enabled, Setωᵢ is equal to any Set ℓ (see #3439)
            (Type{}     , Inf{}      )
              | Bool
typeInTypeEnabled      -> m ()
yes
            (Inf{}      , Type{}     )
              | Bool
typeInTypeEnabled      -> m ()
yes

            -- equating @PiSort a b@ to another sort
            (Sort
s1 , PiSort Dom' Term Term
a Sort
b Abs Sort
c) -> Sort -> Dom' Term Term -> Sort -> Abs Sort -> Blocker -> m ()
piSortEquals Sort
s1 Dom' Term Term
a Sort
b Abs Sort
c Blocker
blocker
            (PiSort Dom' Term Term
a Sort
b Abs Sort
c , Sort
s2) -> Sort -> Dom' Term Term -> Sort -> Abs Sort -> Blocker -> m ()
piSortEquals Sort
s2 Dom' Term Term
a Sort
b Abs Sort
c Blocker
blocker

            -- equating @FunSort a b@ to another sort
            (Sort
s1 , FunSort Sort
a Sort
b) -> Sort -> Sort -> Sort -> Blocker -> m ()
funSortEquals Sort
s1 Sort
a Sort
b Blocker
blocker
            (FunSort Sort
a Sort
b , Sort
s2) -> Sort -> Sort -> Sort -> Blocker -> m ()
funSortEquals Sort
s2 Sort
a Sort
b Blocker
blocker

            -- equating @UnivSort s@ to another sort
            (Sort
s1          , UnivSort Sort
s2) -> Sort -> Sort -> Blocker -> m ()
univSortEquals Sort
s1 Sort
s2 Blocker
blocker
            (UnivSort Sort
s1 , Sort
s2         ) -> Sort -> Sort -> Blocker -> m ()
univSortEquals Sort
s2 Sort
s1 Blocker
blocker

            -- postulated sorts can only be equal if they have the same head
            (DefS QName
d Elims
es  , DefS QName
d' Elims
es')
              | QName
d forall a. Eq a => a -> a -> Bool
== QName
d'                -> do
                  [Polarity]
pol <- forall (m :: * -> *).
HasConstInfo m =>
Comparison -> QName -> m [Polarity]
getPolarity' Comparison
CmpEq QName
d
                  Type
a <- forall (m :: * -> *).
MonadConversion m =>
QName -> Elims -> Elims -> m Type
computeElimHeadType QName
d Elims
es Elims
es'
                  forall (m :: * -> *).
MonadConversion m =>
[Polarity] -> [IsForced] -> Type -> Term -> Elims -> Elims -> m ()
compareElims [Polarity]
pol [] Type
a (QName -> Elims -> Term
Def QName
d []) Elims
es Elims
es'
              | Bool
otherwise              -> forall {a}. m a
no

            -- any other combinations of sorts are not equal
            (Sort
_          , Sort
_          ) -> forall {a}. m a
no

    where
      yes :: m ()
yes = forall (m :: * -> *) a. Monad m => a -> m a
return ()
      no :: m a
no  = forall (m :: * -> *) a. MonadBlock m => Blocker -> m a
patternViolation Blocker
neverUnblock

      -- perform assignment (MetaS x es) := s
      meta :: MetaId -> [Elim' Term] -> Sort -> m ()
      meta :: MetaId -> Elims -> Sort -> m ()
meta MetaId
x Elims
es Sort
s = do
        forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> VerboseKey -> m ()
reportSLn VerboseKey
"tc.meta.sort" Int
30 forall a b. (a -> b) -> a -> b
$ VerboseKey
"Assigning meta sort"
        forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"tc.meta.sort" Int
50 forall a b. (a -> b) -> a -> b
$ TCMT IO Doc
"meta" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
sep [forall (m :: * -> *) a. (Applicative m, Pretty a) => a -> m Doc
pretty MetaId
x, 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 (m :: * -> *) a. (Applicative m, Pretty a) => a -> m Doc
pretty Elims
es, forall (m :: * -> *) a. (Applicative m, Pretty a) => a -> m Doc
pretty Sort
s]
        forall (m :: * -> *).
MonadConversion m =>
CompareDirection
-> MetaId
-> Elims
-> Term
-> CompareAs
-> (Term -> Term -> m ())
-> m ()
assignE CompareDirection
DirEq MetaId
x Elims
es (Sort -> Term
Sort Sort
s) CompareAs
AsTypes forall a. HasCallStack => a
__IMPOSSIBLE__

      -- Equate a sort @s1@ to @univSort s2@
      -- Precondition: @s1@ and @univSort s2@ are already reduced.
      univSortEquals :: Sort -> Sort -> Blocker -> m ()
      univSortEquals :: Sort -> Sort -> Blocker -> m ()
univSortEquals Sort
s1 Sort
s2 Blocker
blocker = do
        forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"tc.conv.sort" Int
35 forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
vcat
          [ TCMT IO Doc
"univSortEquals"
          , TCMT IO Doc
"  s1 =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Sort
s1
          , TCMT IO Doc
"  s2 =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Sort
s2
          ]
        let postpone :: m ()
postpone = forall (m :: * -> *) a. MonadBlock m => Blocker -> m a
patternViolation Blocker
blocker
        case Sort
s1 of
          -- @Set l1@ is the successor sort of either @Set l2@ or
          -- @Prop l2@ where @l1 == lsuc l2@.
          Type Level
l1 -> do
            Bool
propEnabled <- forall (m :: * -> *). HasOptions m => m Bool
isPropEnabled
               -- @s2@ is definitely not @Inf n@ or @SizeUniv@
            if | Inf IsFibrant
_ Integer
n  <- Sort
s2 -> forall {a}. m a
no
               | Sort
SizeUniv <- Sort
s2 -> forall {a}. m a
no
               -- If @Prop@ is not used, then @s2@ must be of the form
               -- @Set l2@
               | Bool -> Bool
not Bool
propEnabled -> do
                   Level
l2 <- case Integer -> Level -> Maybe Level
subLevel Integer
1 Level
l1 of
                     Just Level
l2 -> forall (m :: * -> *) a. Monad m => a -> m a
return Level
l2
                     Maybe Level
Nothing -> do
                       Level
l2 <- forall (m :: * -> *). MonadMetaSolver m => m Level
newLevelMeta
                       forall (m :: * -> *). MonadConversion m => Level -> Level -> m ()
equalLevel Level
l1 (Level -> Level
levelSuc Level
l2)
                       forall (m :: * -> *) a. Monad m => a -> m a
return Level
l2
                   forall (m :: * -> *). MonadConversion m => Sort -> Sort -> m ()
equalSort (forall t. Level' t -> Sort' t
Type Level
l2) Sort
s2
               -- Otherwise we postpone
               | Bool
otherwise -> m ()
postpone
          -- @Setωᵢ@ is a successor sort if n > 0, or if
          -- --type-in-type or --omega-in-omega is enabled.
          Inf IsFibrant
f Integer
n | Integer
n forall a. Ord a => a -> a -> Bool
> Integer
0 -> forall (m :: * -> *). MonadConversion m => Sort -> Sort -> m ()
equalSort (forall t. IsFibrant -> Integer -> Sort' t
Inf IsFibrant
f forall a b. (a -> b) -> a -> b
$ Integer
n forall a. Num a => a -> a -> a
- Integer
1) Sort
s2
          Inf IsFibrant
f Integer
0 -> do
            Bool
infInInf <- (PragmaOptions -> Bool
optOmegaInOmega forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> forall (m :: * -> *). HasOptions m => m PragmaOptions
pragmaOptions) forall (m :: * -> *). Monad m => m Bool -> m Bool -> m Bool
`or2M` forall (m :: * -> *). HasOptions m => m Bool
typeInType
            if | Bool
infInInf  -> forall (m :: * -> *). MonadConversion m => Sort -> Sort -> m ()
equalSort (forall t. IsFibrant -> Integer -> Sort' t
Inf IsFibrant
f Integer
0) Sort
s2
               | Bool
otherwise -> forall {a}. m a
no
          -- @Prop l@ and @SizeUniv@ are not successor sorts
          Prop{}     -> forall {a}. m a
no
          SizeUniv{} -> forall {a}. m a
no
          -- Anything else: postpone
          Sort
_          -> m ()
postpone


      -- Equate a sort @s@ to @piSort a s1 s2@
      -- Precondition: @s@ and @piSort a s1 s2@ are already reduced.
      piSortEquals :: Sort -> Dom Term -> Sort -> Abs Sort -> Blocker -> m ()
      piSortEquals :: Sort -> Dom' Term Term -> Sort -> Abs Sort -> Blocker -> m ()
piSortEquals Sort
s Dom' Term Term
a Sort
s1 NoAbs{} Blocker
blocker = forall a. HasCallStack => a
__IMPOSSIBLE__
      piSortEquals Sort
s Dom' Term Term
a Sort
s1 s2Abs :: Abs Sort
s2Abs@(Abs VerboseKey
x Sort
s2) Blocker
blocker = do
        let adom :: Dom Type
adom = forall t a. Sort' t -> a -> Type'' t a
El Sort
s1 forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> Dom' Term Term
a
        forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"tc.conv.sort" Int
35 forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
vcat
          [ TCMT IO Doc
"piSortEquals"
          , TCMT IO Doc
"  s  =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Sort
s
          , TCMT IO Doc
"  a  =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Dom Type
adom
          , TCMT IO Doc
"  s1 =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Sort
s1
          , TCMT IO Doc
"  s2 =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall b (m :: * -> *) a.
(AddContext b, MonadAddContext m) =>
b -> m a -> m a
addContext (VerboseKey
x,Dom Type
adom) (forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Sort
s2)
          ]
        Bool
propEnabled <- forall (m :: * -> *). HasOptions m => m Bool
isPropEnabled
        let postpone :: m ()
postpone = forall (m :: * -> *) a. MonadBlock m => Blocker -> m a
patternViolation Blocker
blocker
           -- If @s2@ is dependent, then @piSort a s1 s2@ computes to
           -- @Setωi@. Hence, if @s@ is small, then @s2@
           -- cannot be dependent.
        if | Sort -> Bool
isSmallSort Sort
s -> do
               -- We force @s2@ to be non-dependent by unifying it with
               -- a fresh meta that does not depend on @x : a@
               Sort
s2' <- forall (m :: * -> *). MonadMetaSolver m => m Sort
newSortMeta
               forall b (m :: * -> *) a.
(AddContext b, MonadAddContext m) =>
b -> m a -> m a
addContext (VerboseKey
x , Dom Type
adom) forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *). MonadConversion m => Sort -> Sort -> m ()
equalSort Sort
s2 (forall a. Subst a => Int -> a -> a
raise Int
1 Sort
s2')
               Sort -> Sort -> Sort -> Blocker -> m ()
funSortEquals Sort
s Sort
s1 Sort
s2' Blocker
blocker
           -- Otherwise: postpone
           | Bool
otherwise -> m ()
postpone

      -- Equate a sort @s@ to @funSort s1 s2@
      -- Precondition: @s@ and @funSort s1 s2@ are already reduced
      funSortEquals :: Sort -> Sort -> Sort -> Blocker -> m ()
      funSortEquals :: Sort -> Sort -> Sort -> Blocker -> m ()
funSortEquals Sort
s0 Sort
s1 Sort
s2 Blocker
blocker = do
        forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"tc.conv.sort" Int
35 forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
vcat
          [ TCMT IO Doc
"funSortEquals"
          , TCMT IO Doc
"  s0 =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Sort
s0
          , TCMT IO Doc
"  s1 =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Sort
s1
          , TCMT IO Doc
"  s2 =" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Sort
s2
          ]
        Bool
propEnabled <- forall (m :: * -> *). HasOptions m => m Bool
isPropEnabled
        Bool
sizedTypesEnabled <- forall (m :: * -> *). HasOptions m => m Bool
sizedTypesOption
        Bool
cubicalEnabled <- forall a. Maybe a -> Bool
isJust forall b c a. (b -> c) -> (a -> b) -> a -> c
. PragmaOptions -> Maybe Cubical
optCubical forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> forall (m :: * -> *). HasOptions m => m PragmaOptions
pragmaOptions
        let postpone :: m ()
postpone = forall (m :: * -> *) a. MonadBlock m => Blocker -> m a
patternViolation Blocker
blocker
        case Sort
s0 of
          -- If @Setωᵢ == funSort s1 s2@, then either @s1@ or @s2@ must
          -- be @Setωᵢ@.
          Inf IsFibrant
f Integer
n | Sort -> Bool
isSmallSort Sort
s1, Sort -> Bool
isSmallSort Sort
s2 -> do
                    forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
TypeError -> m a
typeError forall a b. (a -> b) -> a -> b
$ Sort -> Sort -> TypeError
UnequalSorts Sort
s0 (forall t. Sort' t -> Sort' t -> Sort' t
FunSort Sort
s1 Sort
s2)
                  | Right (SmallSort IsFibrant
IsFibrant) <- Sort -> Either Blocker SizeOfSort
sizeOfSort Sort
s1 -> forall (m :: * -> *). MonadConversion m => Sort -> Sort -> m ()
equalSort (forall t. IsFibrant -> Integer -> Sort' t
Inf IsFibrant
f Integer
n) Sort
s2
                  | Right (SmallSort IsFibrant
IsFibrant) <- Sort -> Either Blocker SizeOfSort
sizeOfSort Sort
s2 -> forall (m :: * -> *). MonadConversion m => Sort -> Sort -> m ()
equalSort (forall t. IsFibrant -> Integer -> Sort' t
Inf IsFibrant
f Integer
n) Sort
s1
                  -- TODO 2ltt: handle IsStrict cases.
                  | Bool
otherwise -> m ()
postpone
          -- If @Set l == funSort s1 s2@, then @s2@ must be of the
          -- form @Set l2@. @s1@ can be one of @Set l1@, @Prop l1@,
          -- @SizeUniv@, or @IUniv@.
          Type Level
l -> do
            Level
l2 <- Sort -> m Level
forceType Sort
s2
            -- We must have @l2 =< l@, this might help us to solve
            -- more constraints (in particular when @l == 0@).
            forall (m :: * -> *). MonadConversion m => Level -> Level -> m ()
leqLevel Level
l2 Level
l
            -- Jesper, 2022-10-22, #6211: the operations `forceType`
            -- and `leqLevel` above might have instantiated some
            -- metas, so we need to reduce s1 again to get an
            -- up-to-date Blocker.
            Blocked Sort
s1b <- forall a (m :: * -> *).
(Reduce a, MonadReduce m) =>
a -> m (Blocked a)
reduceB Sort
s1
            let s1 :: Sort
s1 = forall t a. Blocked' t a -> a
ignoreBlocking Blocked Sort
s1b
                blocker :: Blocker
blocker = forall t a. Blocked' t a -> Blocker
getBlocker Blocked Sort
s1b
            -- Jesper, 2019-12-27: SizeUniv is disabled at the moment.
            if | {- sizedTypesEnabled || -} Bool
propEnabled Bool -> Bool -> Bool
|| Bool
cubicalEnabled ->
                case Sort -> Sort -> Either Blocker Sort
funSort' Sort
s1 (forall t. Level' t -> Sort' t
Type Level
l2) of
                   -- If the work we did makes the @funSort@ compute,
                   -- continue working.
                   Right Sort
s -> forall (m :: * -> *). MonadConversion m => Sort -> Sort -> m ()
equalSort (forall t. Level' t -> Sort' t
Type Level
l) Sort
s
                   -- Otherwise: postpone
                   Left{}  -> forall (m :: * -> *) a. MonadBlock m => Blocker -> m a
patternViolation Blocker
blocker
               -- If both Prop and sized types are disabled, only the
               -- case @s1 == Set l1@ remains.
               | Bool
otherwise -> do
                   Level
l1 <- Sort -> m Level
forceType Sort
s1
                   forall (m :: * -> *). MonadConversion m => Level -> Level -> m ()
equalLevel Level
l (Level -> Level -> Level
levelLub Level
l1 Level
l2)
          -- If @Prop l == funSort s1 s2@, then @s2@ must be of the
          -- form @Prop l2@, and @s1@ can be one of @Set l1@, Prop
          -- l1@, or @SizeUniv@.
          Prop Level
l -> do
            Level
l2 <- Sort -> m Level
forceProp Sort
s2
            forall (m :: * -> *). MonadConversion m => Level -> Level -> m ()
leqLevel Level
l2 Level
l
            Blocked Sort
s1b <- forall a (m :: * -> *).
(Reduce a, MonadReduce m) =>
a -> m (Blocked a)
reduceB Sort
s1
            let s1 :: Sort
s1 = forall t a. Blocked' t a -> a
ignoreBlocking Blocked Sort
s1b
                blocker :: Blocker
blocker = forall t a. Blocked' t a -> Blocker
getBlocker Blocked Sort
s1b
            case Sort -> Sort -> Either Blocker Sort
funSort' Sort
s1 (forall t. Level' t -> Sort' t
Prop Level
l2) of
                   -- If the work we did makes the @funSort@ compute,
                   -- continue working.
                   Right Sort
s -> forall (m :: * -> *). MonadConversion m => Sort -> Sort -> m ()
equalSort (forall t. Level' t -> Sort' t
Prop Level
l) Sort
s
                   -- Otherwise: postpone
                   Left Blocker
_  -> forall (m :: * -> *) a. MonadBlock m => Blocker -> m a
patternViolation Blocker
blocker
          -- We have @SizeUniv == funSort s1 s2@ iff @s2 == SizeUniv@
          Sort
SizeUniv -> forall (m :: * -> *). MonadConversion m => Sort -> Sort -> m ()
equalSort forall t. Sort' t
SizeUniv Sort
s2
          -- Anything else: postpone
          Sort
_        -> m ()
postpone

      -- check if the given sort @s0@ is a (closed) bottom sort
      -- i.e. @piSort a b == s0@ implies @b == s0@.
      isBottomSort :: Bool -> Sort -> Bool
      isBottomSort :: Bool -> Sort -> Bool
isBottomSort Bool
propEnabled (Prop (ClosedLevel Integer
0)) = Bool
True
      isBottomSort Bool
propEnabled (Type (ClosedLevel Integer
0)) = Bool -> Bool
not Bool
propEnabled
      isBottomSort Bool
propEnabled Sort
_                      = Bool
False
      -- (NB: Defined but not currently used)

      forceType :: Sort -> m Level
      forceType :: Sort -> m Level
forceType (Type Level
l) = forall (m :: * -> *) a. Monad m => a -> m a
return Level
l
      forceType Sort
s = do
        Level
l <- forall (m :: * -> *). MonadMetaSolver m => m Level
newLevelMeta
        forall (m :: * -> *). MonadConversion m => Sort -> Sort -> m ()
equalSort Sort
s (forall t. Level' t -> Sort' t
Type Level
l)
        forall (m :: * -> *) a. Monad m => a -> m a
return Level
l

      forceProp :: Sort -> m Level
      forceProp :: Sort -> m Level
forceProp (Prop Level
l) = forall (m :: * -> *) a. Monad m => a -> m a
return Level
l
      forceProp Sort
s = do
        Level
l <- forall (m :: * -> *). MonadMetaSolver m => m Level
newLevelMeta
        forall (m :: * -> *). MonadConversion m => Sort -> Sort -> m ()
equalSort Sort
s (forall t. Level' t -> Sort' t
Prop Level
l)
        forall (m :: * -> *) a. Monad m => a -> m a
return Level
l

      impossibleSort :: a -> m b
impossibleSort a
s = do
        forall a (m :: * -> *).
(ReportS a, MonadDebug m) =>
VerboseKey -> Int -> a -> m ()
reportS VerboseKey
"impossible" Int
10
          [ a
"equalSort: found dummy sort with description:"
          , a
s
          ]
        forall a. HasCallStack => a
__IMPOSSIBLE__

      catchInequalLevel :: m a -> m a -> m a
catchInequalLevel m a
m m a
fail = m a
m forall e (m :: * -> *) a.
MonadError e m =>
m a -> (e -> m a) -> m a
`catchError` \case
        TypeError{} -> m a
fail
        TCErr
err         -> forall e (m :: * -> *) a. MonadError e m => e -> m a
throwError TCErr
err


-- -- This should probably represent face maps with a more precise type
-- toFaceMaps :: Term -> TCM [[(Int,Term)]]
-- toFaceMaps t = do
--   view <- intervalView'
--   iz <- primIZero
--   io <- primIOne
--   ineg <- (\ q t -> Def q [Apply $ Arg defaultArgInfo t]) <$> fromMaybe __IMPOSSIBLE__ <$> getPrimitiveName' "primINeg"

--   let f IZero = mzero
--       f IOne  = return []
--       f (IMin x y) = do xs <- (f . view . unArg) x; ys <- (f . view . unArg) y; return (xs ++ ys)
--       f (IMax x y) = msum $ map (f . view . unArg) [x,y]
--       f (INeg x)   = map (id -*- not) <$> (f . view . unArg) x
--       f (OTerm (Var i [])) = return [(i,True)]
--       f (OTerm _) = return [] -- what about metas? we should suspend? maybe no metas is a precondition?
--       isConsistent xs = all (\ xs -> length xs == 1) . map nub . Map.elems $ xs  -- optimize by not doing generate + filter
--       as = map (map (id -*- head) . Map.toAscList) . filter isConsistent . map (Map.fromListWith (++) . map (id -*- (:[]))) $ (f (view t))
--   xs <- mapM (mapM (\ (i,b) -> (,) i <$> intervalUnview (if b then IOne else IZero))) as
--   return xs

forallFaceMaps
  :: MonadConversion m
  => Term
  -> (IntMap Bool -> Blocker -> Term -> m a)
  -> (IntMap Bool -> Substitution -> m a)
  -> m [a]
forallFaceMaps :: forall (m :: * -> *) a.
MonadConversion m =>
Term
-> (IntMap Bool -> Blocker -> Term -> m a)
-> (IntMap Bool -> Substitution -> m a)
-> m [a]
forallFaceMaps Term
t IntMap Bool -> Blocker -> Term -> m a
kb IntMap Bool -> Substitution -> m a
k = do
  forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"conv.forall" Int
20 forall a b. (a -> b) -> a -> b
$
      forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
fsep [TCMT IO Doc
"forallFaceMaps"
           , forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Term
t
           ]
  [(IntMap Bool, [Term])]
as <- forall (m :: * -> *).
HasBuiltins m =>
Term -> m [(IntMap Bool, [Term])]
decomposeInterval Term
t
  Bool -> Term
boolToI <- do
    Term
io <- forall (m :: * -> *).
(HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) =>
m Term
primIOne
    Term
iz <- forall (m :: * -> *).
(HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) =>
m Term
primIZero
    forall (m :: * -> *) a. Monad m => a -> m a
return (\Bool
b -> if Bool
b then Term
io else Term
iz)
  forall (t :: * -> *) (m :: * -> *) a b.
(Traversable t, Monad m) =>
t a -> (a -> m b) -> m (t b)
forM [(IntMap Bool, [Term])]
as forall a b. (a -> b) -> a -> b
$ \ (IntMap Bool
ms,[Term]
ts) -> do
   forall {m :: * -> *} {t :: * -> *} {b}.
(HasBuiltins m, MonadError TCErr m, Foldable t, MonadReduce m) =>
t Term
-> (Blocker -> Term -> m b) -> (NotBlocked -> Term -> m b) -> m b
ifBlockeds [Term]
ts (IntMap Bool -> Blocker -> Term -> m a
kb IntMap Bool
ms) forall a b. (a -> b) -> a -> b
$ \ NotBlocked
_ Term
_ -> do
    let xs :: [(Int, Term)]
xs = forall a b. (a -> b) -> [a] -> [b]
map (forall (a :: * -> * -> *) b c d.
Arrow a =>
a b c -> a (d, b) (d, c)
second Bool -> Term
boolToI) forall a b. (a -> b) -> a -> b
$ forall a. IntMap a -> [(Int, a)]
IntMap.toAscList IntMap Bool
ms
    Context
cxt <- forall (m :: * -> *). MonadTCEnv m => m Context
getContext
    forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"conv.forall" Int
20 forall a b. (a -> b) -> a -> b
$
      forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
fsep [TCMT IO Doc
"substContextN"
           , forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Context
cxt
           , forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM [(Int, Term)]
xs
           ]
    (Context
cxt',Substitution
sigma) <- forall (m :: * -> *).
MonadConversion m =>
Context -> [(Int, Term)] -> m (Context, Substitution)
substContextN Context
cxt [(Int, Term)]
xs
    [(Dom' Term (Name, Type), Term)]
resolved <- forall (t :: * -> *) (m :: * -> *) a b.
(Traversable t, Monad m) =>
t a -> (a -> m b) -> m (t b)
forM [(Int, Term)]
xs (\ (Int
i,Term
t) -> (,) forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> forall (m :: * -> *).
(MonadFail m, MonadTCEnv m) =>
Int -> m (Dom' Term (Name, Type))
lookupBV Int
i forall (f :: * -> *) a b. Applicative f => f (a -> b) -> f a -> f b
<*> forall (m :: * -> *) a. Monad m => a -> m a
return (forall a. Subst a => Substitution' (SubstArg a) -> a -> a
applySubst Substitution
sigma Term
t))
    forall (m :: * -> *) a.
MonadAddContext m =>
Substitution -> (Context -> Context) -> m a -> m a
updateContext Substitution
sigma (forall a b. a -> b -> a
const Context
cxt') forall a b. (a -> b) -> a -> b
$
      forall {m :: * -> *} {t} {a}.
MonadAddContext m =>
[(Dom' t (Name, Type), Term)] -> m a -> m a
addBindings [(Dom' Term (Name, Type), Term)]
resolved forall a b. (a -> b) -> a -> b
$ do
        Closure ()
cl <- forall (m :: * -> *) a.
(MonadTCEnv m, ReadTCState m) =>
a -> m (Closure a)
buildClosure ()
        Telescope
tel <- forall (m :: * -> *). (Applicative m, MonadTCEnv m) => m Telescope
getContextTelescope
        ModuleName
m <- forall (m :: * -> *). MonadTCEnv m => m ModuleName
currentModule
        Maybe Substitution
sub <- forall (m :: * -> *).
(MonadTCEnv m, ReadTCState m) =>
ModuleName -> m (Maybe Substitution)
getModuleParameterSub ModuleName
m
        forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"conv.forall" Int
30 forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
vcat
          [ forall (m :: * -> *). Applicative m => VerboseKey -> m Doc
text (forall a. Int -> a -> [a]
replicate Int
10 Char
'-')
          , forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM (TCEnv -> ModuleName
envCurrentModule forall a b. (a -> b) -> a -> b
$ forall a. Closure a -> TCEnv
clEnv Closure ()
cl)
          -- , prettyTCM (envLetBindings $ clEnv cl)
          , forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Telescope
tel -- (toTelescope $ envContext $ clEnv cl)
          , forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Substitution
sigma
          , forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM ModuleName
m
          , forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Maybe Substitution
sub
          ]
        IntMap Bool -> Substitution -> m a
k IntMap Bool
ms Substitution
sigma
  where
    -- TODO Andrea: inefficient because we try to reduce the ts which we know are in whnf
    ifBlockeds :: t Term
-> (Blocker -> Term -> m b) -> (NotBlocked -> Term -> m b) -> m b
ifBlockeds t Term
ts Blocker -> Term -> m b
blocked NotBlocked -> Term -> m b
unblocked = do
      Term
and <- forall (m :: * -> *).
(HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) =>
VerboseKey -> m Term
getPrimitiveTerm VerboseKey
"primIMin"
      Term
io  <- forall (m :: * -> *).
(HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) =>
m Term
primIOne
      let t :: Term
t = forall (t :: * -> *) a b.
Foldable t =>
(a -> b -> b) -> b -> t a -> b
foldr (\ Term
x Term
r -> Term
and forall t. Apply t => t -> [Arg Term] -> t
`apply` [forall e. e -> Arg e
argN Term
x,forall e. e -> Arg e
argN Term
r]) Term
io t Term
ts
      forall t (m :: * -> *) a.
(Reduce t, IsMeta t, MonadReduce m) =>
t -> (Blocker -> t -> m a) -> (NotBlocked -> t -> m a) -> m a
ifBlocked Term
t Blocker -> Term -> m b
blocked NotBlocked -> Term -> m b
unblocked
    addBindings :: [(Dom' t (Name, Type), Term)] -> m a -> m a
addBindings [] m a
m = m a
m
    addBindings ((Dom{domInfo :: forall t e. Dom' t e -> ArgInfo
domInfo = ArgInfo
info,unDom :: forall t e. Dom' t e -> e
unDom = (Name
nm,Type
ty)},Term
t):[(Dom' t (Name, Type), Term)]
bs) m a
m = forall (m :: * -> *) a.
MonadAddContext m =>
ArgInfo -> Name -> Term -> Type -> m a -> m a
addLetBinding ArgInfo
info Name
nm Term
t Type
ty ([(Dom' t (Name, Type), Term)] -> m a -> m a
addBindings [(Dom' t (Name, Type), Term)]
bs m a
m)

    substContextN :: MonadConversion m => Context -> [(Int,Term)] -> m (Context , Substitution)
    substContextN :: forall (m :: * -> *).
MonadConversion m =>
Context -> [(Int, Term)] -> m (Context, Substitution)
substContextN Context
c [] = forall (m :: * -> *) a. Monad m => a -> m a
return (Context
c, forall a. Substitution' a
idS)
    substContextN Context
c ((Int
i,Term
t):[(Int, Term)]
xs) = do
      (Context
c', Substitution
sigma) <- forall (m :: * -> *).
MonadConversion m =>
Int -> Term -> Context -> m (Context, Substitution)
substContext Int
i Term
t Context
c
      (Context
c'', Substitution
sigma')  <- forall (m :: * -> *).
MonadConversion m =>
Context -> [(Int, Term)] -> m (Context, Substitution)
substContextN Context
c' (forall a b. (a -> b) -> [a] -> [b]
map (forall a. Num a => a -> a -> a
subtract Int
1 forall a c b d. (a -> c) -> (b -> d) -> (a, b) -> (c, d)
-*- forall a. Subst a => Substitution' (SubstArg a) -> a -> a
applySubst Substitution
sigma) [(Int, Term)]
xs)
      forall (m :: * -> *) a. Monad m => a -> m a
return (Context
c'', forall a. Subst a => Substitution' (SubstArg a) -> a -> a
applySubst Substitution
sigma' Substitution
sigma)


    -- assumes the term can be typed in the shorter telescope
    -- the terms we get from toFaceMaps are closed.
    substContext :: MonadConversion m => Int -> Term -> Context -> m (Context , Substitution)
    substContext :: forall (m :: * -> *).
MonadConversion m =>
Int -> Term -> Context -> m (Context, Substitution)
substContext Int
i Term
t [] = forall a. HasCallStack => a
__IMPOSSIBLE__
    substContext Int
i Term
t (Dom' Term (Name, Type)
x:Context
xs) | Int
i forall a. Eq a => a -> a -> Bool
== Int
0 = forall (m :: * -> *) a. Monad m => a -> m a
return forall a b. (a -> b) -> a -> b
$ (Context
xs , forall a. DeBruijn a => Int -> a -> Substitution' a
singletonS Int
0 Term
t)
    substContext Int
i Term
t (Dom' Term (Name, Type)
x:Context
xs) | Int
i forall a. Ord a => a -> a -> Bool
> Int
0 = do
                                  forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"conv.forall" Int
20 forall a b. (a -> b) -> a -> b
$
                                    forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
fsep [TCMT IO Doc
"substContext"
                                        , forall (m :: * -> *). Applicative m => VerboseKey -> m Doc
text (forall a. Show a => a -> VerboseKey
show (Int
iforall a. Num a => a -> a -> a
-Int
1))
                                        , forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Term
t
                                        , forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Context
xs
                                        ]
                                  (Context
c,Substitution
sigma) <- forall (m :: * -> *).
MonadConversion m =>
Int -> Term -> Context -> m (Context, Substitution)
substContext (Int
iforall a. Num a => a -> a -> a
-Int
1) Term
t Context
xs
                                  let e :: Dom' Term (Name, Type)
e = forall a. Subst a => Substitution' (SubstArg a) -> a -> a
applySubst Substitution
sigma Dom' Term (Name, Type)
x
                                  forall (m :: * -> *) a. Monad m => a -> m a
return (Dom' Term (Name, Type)
eforall a. a -> [a] -> [a]
:Context
c, forall a. Int -> Substitution' a -> Substitution' a
liftS Int
1 Substitution
sigma)
    substContext Int
i Term
t (Dom' Term (Name, Type)
x:Context
xs) = forall a. HasCallStack => a
__IMPOSSIBLE__

compareInterval :: MonadConversion m => Comparison -> Type -> Term -> Term -> m ()
compareInterval :: forall (m :: * -> *).
MonadConversion m =>
Comparison -> Type -> Term -> Term -> m ()
compareInterval Comparison
cmp Type
i Term
t Term
u = do
  forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"tc.conv.interval" Int
15 forall a b. (a -> b) -> a -> b
$
    forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
sep [ TCMT IO Doc
"{ compareInterval" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Term
t 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
u ]
  forall (m :: * -> *). MonadDebug m => ProfileOption -> m () -> m ()
whenProfile ProfileOption
Profile.Conversion forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *). MonadStatistics m => VerboseKey -> m ()
tick VerboseKey
"compare at interval type"
  Blocked Term
tb <- forall a (m :: * -> *).
(Reduce a, MonadReduce m) =>
a -> m (Blocked a)
reduceB Term
t
  Blocked Term
ub <- forall a (m :: * -> *).
(Reduce a, MonadReduce m) =>
a -> m (Blocked a)
reduceB Term
u
  let t :: Term
t = forall t a. Blocked' t a -> a
ignoreBlocking Blocked Term
tb
      u :: Term
u = forall t a. Blocked' t a -> a
ignoreBlocking Blocked Term
ub
  [(IntMap BoolSet, [Term])]
it <- forall (m :: * -> *).
HasBuiltins m =>
Term -> m [(IntMap BoolSet, [Term])]
decomposeInterval' Term
t
  [(IntMap BoolSet, [Term])]
iu <- forall (m :: * -> *).
HasBuiltins m =>
Term -> m [(IntMap BoolSet, [Term])]
decomposeInterval' Term
u
  case () of
    ()
_ | forall {t} {a}. Blocked' t a -> Bool
isBlocked Blocked Term
tb Bool -> Bool -> Bool
|| forall {t} {a}. Blocked' t a -> Bool
isBlocked Blocked Term
ub -> do
      -- in case of metas we wouldn't be able to make progress by how we deal with de morgan laws.
      -- (because the constraints generated by decomposition are sufficient but not necessary).
      -- but we could still prune/solve some metas by comparing the terms as atoms.
      -- also if blocked we won't find the terms conclusively unequal(?) so compareAtom
      -- won't report type errors when we should accept.
      Type
interval <- forall (m :: * -> *).
(HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) =>
m Type
primIntervalType
      forall (m :: * -> *).
MonadConversion m =>
Comparison -> CompareAs -> Term -> Term -> m ()
compareAtom Comparison
CmpEq (Type -> CompareAs
AsTermsOf Type
interval) Term
t Term
u
    ()
_ | Bool
otherwise -> do
      Bool
x <- forall (m :: * -> *).
MonadConversion m =>
[(IntMap BoolSet, [Term])] -> [(IntMap BoolSet, [Term])] -> m Bool
leqInterval [(IntMap BoolSet, [Term])]
it [(IntMap BoolSet, [Term])]
iu
      Bool
y <- forall (m :: * -> *).
MonadConversion m =>
[(IntMap BoolSet, [Term])] -> [(IntMap BoolSet, [Term])] -> m Bool
leqInterval [(IntMap BoolSet, [Term])]
iu [(IntMap BoolSet, [Term])]
it
      let final :: Bool
final = [(IntMap BoolSet, [Term])] -> Bool
isCanonical [(IntMap BoolSet, [Term])]
it Bool -> Bool -> Bool
&& [(IntMap BoolSet, [Term])] -> Bool
isCanonical [(IntMap BoolSet, [Term])]
iu
      if Bool
x Bool -> Bool -> Bool
&& Bool
y then forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"tc.conv.interval" Int
15 forall a b. (a -> b) -> a -> b
$ TCMT IO Doc
"Ok! }" else
        if Bool
final then forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
TypeError -> m a
typeError forall a b. (a -> b) -> a -> b
$ Comparison -> Term -> Term -> CompareAs -> TypeError
UnequalTerms Comparison
cmp Term
t Term
u (Type -> CompareAs
AsTermsOf Type
i)
                 else do
                   forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"tc.conv.interval" Int
15 forall a b. (a -> b) -> a -> b
$ TCMT IO Doc
"Giving up! }"
                   forall (m :: * -> *) a. MonadBlock m => Blocker -> m a
patternViolation (forall t. AllMetas t => t -> Blocker
unblockOnAnyMetaIn (Term
t, Term
u))
 where
   isBlocked :: Blocked' t a -> Bool
isBlocked Blocked{}    = Bool
True
   isBlocked NotBlocked{} = Bool
False


type Conj = (IntMap BoolSet, [Term])

isCanonical :: [Conj] -> Bool
isCanonical :: [(IntMap BoolSet, [Term])] -> Bool
isCanonical = forall (t :: * -> *) a. Foldable t => (a -> Bool) -> t a -> Bool
all (forall (t :: * -> *) a. Foldable t => t a -> Bool
null forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall a b. (a, b) -> b
snd)

-- | leqInterval r q = r ≤ q in the I lattice.
-- (∨ r_i) ≤ (∨ q_j)  iff  ∀ i. ∃ j. r_i ≤ q_j
leqInterval :: MonadConversion m => [Conj] -> [Conj] -> m Bool
leqInterval :: forall (m :: * -> *).
MonadConversion m =>
[(IntMap BoolSet, [Term])] -> [(IntMap BoolSet, [Term])] -> m Bool
leqInterval [(IntMap BoolSet, [Term])]
r [(IntMap BoolSet, [Term])]
q =
  forall (t :: * -> *). Foldable t => t Bool -> Bool
and forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> forall (t :: * -> *) (m :: * -> *) a b.
(Traversable t, Monad m) =>
t a -> (a -> m b) -> m (t b)
forM [(IntMap BoolSet, [Term])]
r (\ (IntMap BoolSet, [Term])
r_i ->
   forall (t :: * -> *). Foldable t => t Bool -> Bool
or forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> forall (t :: * -> *) (m :: * -> *) a b.
(Traversable t, Monad m) =>
t a -> (a -> m b) -> m (t b)
forM [(IntMap BoolSet, [Term])]
q (\ (IntMap BoolSet, [Term])
q_j -> forall (m :: * -> *).
MonadConversion m =>
(IntMap BoolSet, [Term]) -> (IntMap BoolSet, [Term]) -> m Bool
leqConj (IntMap BoolSet, [Term])
r_i (IntMap BoolSet, [Term])
q_j))  -- TODO shortcut

-- | leqConj r q = r ≤ q in the I lattice, when r and q are conjuctions.
-- ' (∧ r_i)   ≤ (∧ q_j)               iff
-- ' (∧ r_i)   ∧ (∧ q_j)   = (∧ r_i)   iff
-- ' {r_i | i} ∪ {q_j | j} = {r_i | i} iff
-- ' {q_j | j} ⊆ {r_i | i}
leqConj :: MonadConversion m => Conj -> Conj -> m Bool
leqConj :: forall (m :: * -> *).
MonadConversion m =>
(IntMap BoolSet, [Term]) -> (IntMap BoolSet, [Term]) -> m Bool
leqConj (IntMap BoolSet
rs, [Term]
rst) (IntMap BoolSet
qs, [Term]
qst) = do
  if forall a b. (a -> b -> Bool) -> IntMap a -> IntMap b -> Bool
IntMap.isSubmapOfBy BoolSet -> BoolSet -> Bool
BoolSet.isSubsetOf IntMap BoolSet
qs IntMap BoolSet
rs
    then do
      Type
interval <-
        forall t a. Sort' t -> a -> Type'' t a
El forall t. Sort' t
IntervalUniv forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall a. a -> Maybe a -> a
fromMaybe forall a. HasCallStack => a
__IMPOSSIBLE__ forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> forall (m :: * -> *). HasBuiltins m => VerboseKey -> m (Maybe Term)
getBuiltin' VerboseKey
builtinInterval
      -- we don't want to generate new constraints here because
      -- 1. in some situations the same constraint would get generated twice.
      -- 2. unless things are completely accepted we are going to
      --    throw patternViolation in compareInterval.
      let eqT :: Term -> Term -> m Bool
eqT Term
t Term
u = forall (m :: * -> *).
(MonadConstraint m, MonadWarning m, MonadError TCErr m,
 MonadFresh ProblemId m) =>
m () -> m Bool
tryConversion (forall (m :: * -> *).
MonadConversion m =>
Comparison -> CompareAs -> Term -> Term -> m ()
compareAtom Comparison
CmpEq (Type -> CompareAs
AsTermsOf Type
interval) Term
t Term
u)
      let listSubset :: [Term] -> [Term] -> m Bool
listSubset [Term]
ts [Term]
us =
            forall (t :: * -> *). Foldable t => t Bool -> Bool
and forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> forall (t :: * -> *) (m :: * -> *) a b.
(Traversable t, Monad m) =>
t a -> (a -> m b) -> m (t b)
forM [Term]
ts (\Term
t -> forall (t :: * -> *). Foldable t => t Bool -> Bool
or forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> forall (t :: * -> *) (m :: * -> *) a b.
(Traversable t, Monad m) =>
t a -> (a -> m b) -> m (t b)
forM [Term]
us (\Term
u -> Term -> Term -> m Bool
eqT Term
t Term
u)) -- TODO shortcut
      [Term] -> [Term] -> m Bool
listSubset [Term]
qst [Term]
rst
    else
      forall (m :: * -> *) a. Monad m => a -> m a
return Bool
False

-- | equalTermOnFace φ A u v = _ , φ ⊢ u = v : A
equalTermOnFace :: MonadConversion m => Term -> Type -> Term -> Term -> m ()
equalTermOnFace :: forall (m :: * -> *).
MonadConversion m =>
Term -> Type -> Term -> Term -> m ()
equalTermOnFace = forall (m :: * -> *).
MonadConversion m =>
Comparison -> Term -> Type -> Term -> Term -> m ()
compareTermOnFace Comparison
CmpEq

compareTermOnFace :: MonadConversion m => Comparison -> Term -> Type -> Term -> Term -> m ()
compareTermOnFace :: forall (m :: * -> *).
MonadConversion m =>
Comparison -> Term -> Type -> Term -> Term -> m ()
compareTermOnFace = forall (m :: * -> *).
MonadConversion m =>
(Substitution -> Comparison -> Type -> Term -> Term -> m ())
-> Comparison -> Term -> Type -> Term -> Term -> m ()
compareTermOnFace' (forall a b. a -> b -> a
const forall (m :: * -> *).
MonadConversion m =>
Comparison -> Type -> Term -> Term -> m ()
compareTerm)

compareTermOnFace'
  :: MonadConversion m
  => (Substitution -> Comparison -> Type -> Term -> Term -> m ())
  -> Comparison -> Term -> Type -> Term -> Term -> m ()
compareTermOnFace' :: forall (m :: * -> *).
MonadConversion m =>
(Substitution -> Comparison -> Type -> Term -> Term -> m ())
-> Comparison -> Term -> Type -> Term -> Term -> m ()
compareTermOnFace' Substitution -> Comparison -> Type -> Term -> Term -> m ()
k Comparison
cmp Term
phi Type
ty Term
u Term
v = do
  forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"tc.conv.face" Int
40 forall a b. (a -> b) -> a -> b
$
    forall (m :: * -> *). Applicative m => VerboseKey -> m Doc
text VerboseKey
"compareTermOnFace:" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall (m :: * -> *) a. (Applicative m, Pretty a) => a -> m Doc
pretty Term
phi forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> TCMT IO Doc
"|-" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall (m :: * -> *) a. (Applicative m, Pretty a) => a -> m Doc
pretty Term
u forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> TCMT IO Doc
"==" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall (m :: * -> *) a. (Applicative m, Pretty a) => a -> m Doc
pretty 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 (m :: * -> *) a. (Applicative m, Pretty a) => a -> m Doc
pretty Type
ty
  forall (m :: * -> *). MonadDebug m => ProfileOption -> m () -> m ()
whenProfile ProfileOption
Profile.Conversion forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *). MonadStatistics m => VerboseKey -> m ()
tick VerboseKey
"compare at face type"

  Term
phi <- forall a (m :: * -> *). (Reduce a, MonadReduce m) => a -> m a
reduce Term
phi
  [()]
_ <- forall (m :: * -> *) a.
MonadConversion m =>
Term
-> (IntMap Bool -> Blocker -> Term -> m a)
-> (IntMap Bool -> Substitution -> m a)
-> m [a]
forallFaceMaps Term
phi IntMap Bool -> Blocker -> Term -> m ()
postponed forall a b. (a -> b) -> a -> b
$ \ IntMap Bool
faces Substitution
alpha ->
      Substitution -> Comparison -> Type -> Term -> Term -> m ()
k Substitution
alpha Comparison
cmp (forall a. Subst a => Substitution' (SubstArg a) -> a -> a
applySubst Substitution
alpha Type
ty) (forall a. Subst a => Substitution' (SubstArg a) -> a -> a
applySubst Substitution
alpha Term
u) (forall a. Subst a => Substitution' (SubstArg a) -> a -> a
applySubst Substitution
alpha Term
v)
  forall (m :: * -> *) a. Monad m => a -> m a
return ()
 where
  postponed :: IntMap Bool -> Blocker -> Term -> m ()
postponed IntMap Bool
ms Blocker
blocker Term
psi = do
    Term
phi <- forall (m :: * -> *) a. Names -> NamesT m a -> m a
runNamesT [] forall a b. (a -> b) -> a -> b
$ do
             Term
imin <- forall (m :: * -> *) a. Monad m => m a -> NamesT m a
cl forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *).
(HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) =>
VerboseKey -> m Term
getPrimitiveTerm VerboseKey
"primIMin"
             Term
ineg <- forall (m :: * -> *) a. Monad m => m a -> NamesT m a
cl forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *).
(HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) =>
VerboseKey -> m Term
getPrimitiveTerm VerboseKey
"primINeg"
             NamesT m Term
psi <- forall (m :: * -> *) a.
(MonadFail m, Subst a) =>
a -> NamesT m (NamesT m a)
open Term
psi
             let phi :: NamesT m Term
phi = forall (t :: * -> *) a b.
Foldable t =>
(a -> b -> b) -> b -> t a -> b
foldr (\ (Int
i,Bool
b) NamesT m Term
r -> do NamesT m Term
i <- forall (m :: * -> *) a.
(MonadFail m, Subst a) =>
a -> NamesT m (NamesT m a)
open (Int -> Term
var Int
i); forall (f :: * -> *) a. Applicative f => a -> f a
pure Term
imin forall (m :: * -> *). Applicative m => m Term -> m Term -> m Term
<@> (if Bool
b then NamesT m Term
i else forall (f :: * -> *) a. Applicative f => a -> f a
pure Term
ineg forall (m :: * -> *). Applicative m => m Term -> m Term -> m Term
<@> NamesT m Term
i) forall (m :: * -> *). Applicative m => m Term -> m Term -> m Term
<@> NamesT m Term
r)
                          NamesT m Term
psi (forall a. IntMap a -> [(Int, a)]
IntMap.toList IntMap Bool
ms) -- TODO Andrea: make a view?
             NamesT m Term
phi
    forall (m :: * -> *).
MonadConstraint m =>
Blocker -> Constraint -> m ()
addConstraint Blocker
blocker (Comparison -> Term -> Type -> Term -> Term -> Constraint
ValueCmpOnFace Comparison
cmp Term
phi Type
ty Term
u Term
v)

---------------------------------------------------------------------------
-- * Definitions
---------------------------------------------------------------------------

bothAbsurd :: MonadConversion m => QName -> QName -> m Bool
bothAbsurd :: forall (m :: * -> *). MonadConversion m => QName -> QName -> m Bool
bothAbsurd QName
f QName
f'
  | QName -> Bool
isAbsurdLambdaName QName
f, QName -> Bool
isAbsurdLambdaName QName
f' = do
      -- Double check we are really dealing with absurd lambdas:
      -- Their functions should not have bodies.
      Definition
def  <- forall (m :: * -> *). HasConstInfo m => QName -> m Definition
getConstInfo QName
f
      Definition
def' <- forall (m :: * -> *). HasConstInfo m => QName -> m Definition
getConstInfo QName
f'
      case (Definition -> Defn
theDef Definition
def, Definition -> Defn
theDef Definition
def') of
        (Function{ funClauses :: Defn -> [Clause]
funClauses = [Clause{ clauseBody :: Clause -> Maybe Term
clauseBody = Maybe Term
Nothing }] },
         Function{ funClauses :: Defn -> [Clause]
funClauses = [Clause{ clauseBody :: Clause -> Maybe Term
clauseBody = Maybe Term
Nothing }] }) -> forall (m :: * -> *) a. Monad m => a -> m a
return Bool
True
        (Defn, Defn)
_ -> forall (m :: * -> *) a. Monad m => a -> m a
return Bool
False
  | Bool
otherwise = forall (m :: * -> *) a. Monad m => a -> m a
return Bool
False