{-# LANGUAGE NondecreasingIndentation #-}
module Agda.TypeChecking.Conversion where
import Control.Arrow (first, second)
import Control.Monad
import Control.Monad.Fail (MonadFail)
import Data.Function
import qualified Data.List as List
import Data.List.NonEmpty (NonEmpty(..))
import qualified Data.List.NonEmpty as NonEmpty
import qualified Data.Map as Map
import qualified Data.Set as Set
import qualified Data.IntSet as IntSet
import Agda.Syntax.Abstract.Views (isSet)
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.Monad.Builtin
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)
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.Except ( MonadError(catchError, throwError) )
import Agda.Utils.Functor
import Agda.Utils.Monad
import Agda.Utils.Maybe
import Agda.Utils.Permutation
import Agda.Utils.Size
import Agda.Utils.Tuple
import Agda.Utils.WithDefault
import Agda.Utils.Impossible
type MonadConversion m =
( MonadReduce m
, MonadAddContext m
, MonadConstraint m
, MonadMetaSolver m
, MonadError TCErr m
, MonadWarning m
, MonadDebug m
, MonadStatistics m
, MonadFresh ProblemId m
, MonadFresh Int m
, HasBuiltins m
, HasConstInfo m
, HasOptions m
, MonadFail m
)
tryConversion
:: (MonadConstraint m, MonadWarning m, MonadError TCErr m, MonadFresh ProblemId m)
=> m () -> m Bool
tryConversion :: m () -> m Bool
tryConversion = Maybe () -> Bool
forall a. Maybe a -> Bool
isJust (Maybe () -> Bool) -> (m () -> m (Maybe ())) -> m () -> m Bool
forall (m :: * -> *) b c a.
Functor m =>
(b -> c) -> (a -> m b) -> a -> m c
<.> m () -> m (Maybe ())
forall (m :: * -> *) a.
(MonadConstraint m, MonadWarning m, MonadError TCErr m,
MonadFresh ProblemId m) =>
m a -> m (Maybe a)
tryConversion'
tryConversion'
:: (MonadConstraint m, MonadWarning m, MonadError TCErr m, MonadFresh ProblemId m)
=> m a -> m (Maybe a)
tryConversion' :: m a -> m (Maybe a)
tryConversion' m a
m = m a -> m (Maybe a)
forall e (m :: * -> *) a.
(MonadError e m, Functor m) =>
m a -> m (Maybe a)
tryMaybe (m a -> m (Maybe a)) -> m a -> m (Maybe a)
forall a b. (a -> b) -> a -> b
$ m a -> m a
forall (m :: * -> *) a.
(MonadConstraint m, MonadWarning m, MonadError TCErr m,
MonadFresh ProblemId m) =>
m a -> m a
noConstraints m a
m
sameVars :: Elims -> Elims -> Bool
sameVars :: Elims -> Elims -> Bool
sameVars Elims
xs Elims
ys = [Bool] -> Bool
forall (t :: * -> *). Foldable t => t Bool -> Bool
and ([Bool] -> Bool) -> [Bool] -> Bool
forall a b. (a -> b) -> a -> b
$ (Elim' Term -> Elim' Term -> Bool) -> Elims -> Elims -> [Bool]
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 Int -> Int -> Bool
forall a. Eq a => a -> a -> Bool
== Int
m
same Elim' Term
_ Elim' Term
_ = Bool
False
intersectVars :: Elims -> Elims -> Maybe [Bool]
intersectVars :: Elims -> Elims -> Maybe [Bool]
intersectVars = (Elim' Term -> Elim' Term -> Maybe Bool)
-> Elims -> Elims -> Maybe [Bool]
forall (m :: * -> *) a b c.
Applicative m =>
(a -> b -> m c) -> [a] -> [b] -> m [c]
zipWithM Elim' Term -> Elim' Term -> Maybe Bool
areVars where
areVars :: Elim' Term -> Elim' Term -> Maybe Bool
areVars (Apply Arg Term
u) Elim' Term
v | Arg Term -> Bool
forall a. LensRelevance a => a -> Bool
isIrrelevant Arg Term
u = Bool -> Maybe Bool
forall a. a -> Maybe a
Just Bool
False
areVars (Apply (Arg ArgInfo
_ (Var Int
n []))) (Apply (Arg ArgInfo
_ (Var Int
m []))) = Bool -> Maybe Bool
forall a. a -> Maybe a
Just (Bool -> Maybe Bool) -> Bool -> Maybe Bool
forall a b. (a -> b) -> a -> b
$ Int
n Int -> Int -> Bool
forall a. Eq a => a -> a -> Bool
/= Int
m
areVars Elim' Term
_ Elim' Term
_ = Maybe Bool
forall a. Maybe a
Nothing
equalTerm :: MonadConversion m => Type -> Term -> Term -> m ()
equalTerm :: Type -> Term -> Term -> m ()
equalTerm = Comparison -> Type -> Term -> Term -> m ()
forall (m :: * -> *).
MonadConversion m =>
Comparison -> Type -> Term -> Term -> m ()
compareTerm Comparison
CmpEq
equalAtom :: MonadConversion m => CompareAs -> Term -> Term -> m ()
equalAtom :: CompareAs -> Term -> Term -> m ()
equalAtom = Comparison -> CompareAs -> Term -> Term -> m ()
forall (m :: * -> *).
MonadConversion m =>
Comparison -> CompareAs -> Term -> Term -> m ()
compareAtom Comparison
CmpEq
equalType :: MonadConversion m => Type -> Type -> m ()
equalType :: Type -> Type -> m ()
equalType = Comparison -> Type -> Type -> m ()
forall (m :: * -> *).
MonadConversion m =>
Comparison -> Type -> Type -> m ()
compareType Comparison
CmpEq
convError :: TypeError -> TCM ()
convError :: TypeError -> TCM ()
convError TypeError
err = TCMT IO Bool -> TCM () -> TCM () -> TCM ()
forall (m :: * -> *) a. Monad m => m Bool -> m a -> m a -> m a
ifM (Relevance -> Relevance -> Bool
forall a. Eq a => a -> a -> Bool
(==) Relevance
Irrelevant (Relevance -> Bool) -> TCMT IO Relevance -> TCMT IO Bool
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> (TCEnv -> Relevance) -> TCMT IO Relevance
forall (m :: * -> *) a. MonadTCEnv m => (TCEnv -> a) -> m a
asksTC TCEnv -> Relevance
forall a. LensRelevance a => a -> Relevance
getRelevance) (() -> TCM ()
forall (m :: * -> *) a. Monad m => a -> m a
return ()) (TCM () -> TCM ()) -> TCM () -> TCM ()
forall a b. (a -> b) -> a -> b
$ TypeError -> TCM ()
forall (m :: * -> *) a.
(MonadTCEnv m, ReadTCState m, MonadError TCErr m) =>
TypeError -> m a
typeError TypeError
err
compareTerm :: forall m. MonadConversion m => Comparison -> Type -> Term -> Term -> m ()
compareTerm :: Comparison -> Type -> Term -> Term -> m ()
compareTerm Comparison
cmp Type
a Term
u Term
v = Comparison -> CompareAs -> Term -> Term -> m ()
forall (m :: * -> *).
MonadConversion m =>
Comparison -> CompareAs -> Term -> Term -> m ()
compareAs Comparison
cmp (Type -> CompareAs
AsTermsOf Type
a) Term
u Term
v
compareAs :: forall m. MonadConversion m => Comparison -> CompareAs -> Term -> Term -> m ()
compareAs :: Comparison -> CompareAs -> Term -> Term -> m ()
compareAs Comparison
cmp CompareAs
a Term
u Term
v = do
VerboseKey -> Int -> TCM Doc -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCM Doc -> m ()
reportSDoc VerboseKey
"tc.conv.term" Int
10 (TCM Doc -> m ()) -> TCM Doc -> m ()
forall a b. (a -> b) -> a -> b
$ [TCM Doc] -> TCM Doc
forall (m :: * -> *). Monad m => [m Doc] -> m Doc
sep ([TCM Doc] -> TCM Doc) -> [TCM Doc] -> TCM Doc
forall a b. (a -> b) -> a -> b
$
[ TCM Doc
"compareTerm"
, Int -> TCM Doc -> TCM Doc
forall (m :: * -> *). Functor m => Int -> m Doc -> m Doc
nest Int
2 (TCM Doc -> TCM Doc) -> TCM Doc -> TCM Doc
forall a b. (a -> b) -> a -> b
$ Term -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Term
u TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Comparison -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Comparison
cmp TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Term -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Term
v
, Int -> TCM Doc -> TCM Doc
forall (m :: * -> *). Functor m => Int -> m Doc -> m Doc
nest Int
2 (TCM Doc -> TCM Doc) -> TCM Doc -> TCM Doc
forall a b. (a -> b) -> a -> b
$ CompareAs -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM CompareAs
a
]
((Term
u, Term
v), Bool
equal) <- Term -> Term -> m ((Term, Term), Bool)
forall a (m :: * -> *).
(Instantiate a, SynEq a, MonadReduce m) =>
a -> a -> m ((a, a), Bool)
SynEq.checkSyntacticEquality Term
u Term
v
if Bool
equal then VerboseKey -> Int -> m () -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> m () -> m ()
verboseS VerboseKey
"profile.sharing" Int
20 (m () -> m ()) -> m () -> m ()
forall a b. (a -> b) -> a -> b
$ VerboseKey -> m ()
forall (m :: * -> *). MonadStatistics m => VerboseKey -> m ()
tick VerboseKey
"equal terms" else do
VerboseKey -> Int -> m () -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> m () -> m ()
verboseS VerboseKey
"profile.sharing" Int
20 (m () -> m ()) -> m () -> m ()
forall a b. (a -> b) -> a -> b
$ VerboseKey -> m ()
forall (m :: * -> *). MonadStatistics m => VerboseKey -> m ()
tick VerboseKey
"unequal terms"
VerboseKey -> Int -> TCM Doc -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCM Doc -> m ()
reportSDoc VerboseKey
"tc.conv.term" Int
15 (TCM Doc -> m ()) -> TCM Doc -> m ()
forall a b. (a -> b) -> a -> b
$ [TCM Doc] -> TCM Doc
forall (m :: * -> *). Monad m => [m Doc] -> m Doc
sep ([TCM Doc] -> TCM Doc) -> [TCM Doc] -> TCM Doc
forall a b. (a -> b) -> a -> b
$
[ TCM Doc
"compareTerm (not syntactically equal)"
, Int -> TCM Doc -> TCM Doc
forall (m :: * -> *). Functor m => Int -> m Doc -> m Doc
nest Int
2 (TCM Doc -> TCM Doc) -> TCM Doc -> TCM Doc
forall a b. (a -> b) -> a -> b
$ Term -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Term
u TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Comparison -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Comparison
cmp TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Term -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Term
v
, Int -> TCM Doc -> TCM Doc
forall (m :: * -> *). Functor m => Int -> m Doc -> m Doc
nest Int
2 (TCM Doc -> TCM Doc) -> TCM Doc -> TCM Doc
forall a b. (a -> b) -> a -> b
$ CompareAs -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM CompareAs
a
]
let fallback :: m ()
fallback = Comparison -> CompareAs -> Term -> Term -> m ()
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 Comparison -> Comparison -> Bool
forall a. Eq a => a -> a -> Bool
== Comparison
CmpEq then m ()
cont else do
CompareAs
-> (MetaId -> CompareAs -> m ())
-> (NotBlocked -> CompareAs -> m ())
-> m ()
forall t (m :: * -> *) a.
(Reduce t, IsMeta t, MonadReduce m, HasBuiltins m) =>
t -> (MetaId -> t -> m a) -> (NotBlocked -> t -> m a) -> m a
ifBlocked CompareAs
a (\ MetaId
_ CompareAs
_ -> m ()
fallback) ((NotBlocked -> CompareAs -> m ()) -> m ())
-> (NotBlocked -> CompareAs -> m ()) -> m ()
forall a b. (a -> b) -> a -> b
$ \ NotBlocked
_ CompareAs
a -> do
m (Maybe BoundedSize) -> m () -> (BoundedSize -> m ()) -> m ()
forall (m :: * -> *) a b.
Monad m =>
m (Maybe a) -> m b -> (a -> m b) -> m b
caseMaybeM (CompareAs -> m (Maybe BoundedSize)
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
case (Term
u, Term
v) of
(MetaV MetaId
x Elims
us, MetaV MetaId
y Elims
vs)
| MetaId
x MetaId -> MetaId -> Bool
forall a. Eq a => a -> a -> Bool
/= MetaId
y -> m () -> m ()
unlessSubtyping (m () -> m ()) -> m () -> m ()
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 MetaId -> MetaId -> Bool
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 (m () -> m ()) -> m () -> m ()
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 (m () -> m ()) -> m () -> m ()
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
(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
VerboseKey -> Int -> TCM Doc -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCM Doc -> m ()
reportSDoc VerboseKey
"tc.conv.term.shortcut" Int
20 (TCM Doc -> m ()) -> TCM Doc -> m ()
forall a b. (a -> b) -> a -> b
$ [TCM Doc] -> TCM Doc
forall (m :: * -> *). Monad m => [m Doc] -> m Doc
sep
[ TCM Doc
"attempting shortcut"
, Int -> TCM Doc -> TCM Doc
forall (m :: * -> *). Functor m => Int -> m Doc -> m Doc
nest Int
2 (TCM Doc -> TCM Doc) -> TCM Doc -> TCM Doc
forall a b. (a -> b) -> a -> b
$ Term -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM (MetaId -> Elims -> Term
MetaV MetaId
x Elims
es) TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> TCM Doc
":=" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Term -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Term
v
]
m Bool -> m () -> m ()
forall (m :: * -> *). Monad m => m Bool -> m () -> m ()
whenM (MetaId -> m Bool
forall a (m :: * -> *).
(IsInstantiatedMeta a, MonadFail m, ReadTCState m) =>
a -> m Bool
isInstantiatedMeta MetaId
x) m ()
forall (m :: * -> *) a. MonadError TCErr m => m a
patternViolation
CompareDirection
-> MetaId
-> Elims
-> Term
-> CompareAs
-> (Term -> Term -> m ())
-> m ()
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 ()) -> m ()) -> (Term -> Term -> m ()) -> m ()
forall a b. (a -> b) -> a -> b
$ CompareDirection -> CompareAs -> Term -> Term -> m ()
forall (m :: * -> *).
MonadConversion m =>
CompareDirection -> CompareAs -> Term -> Term -> m ()
compareAsDir CompareDirection
dir CompareAs
a
VerboseKey -> Int -> TCM Doc -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCM Doc -> m ()
reportSDoc VerboseKey
"tc.conv.term.shortcut" Int
50 (TCM Doc -> m ()) -> TCM Doc -> m ()
forall a b. (a -> b) -> a -> b
$
TCM Doc
"shortcut successful" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
$$ Int -> TCM Doc -> TCM Doc
forall (m :: * -> *). Functor m => Int -> m Doc -> m Doc
nest Int
2 (TCM Doc
"result:" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> (Term -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty (Term -> TCM Doc) -> TCMT IO Term -> TCM Doc
forall (m :: * -> *) a b. Monad m => (a -> m b) -> m a -> m b
=<< Term -> TCMT IO Term
forall a (m :: * -> *). (Instantiate a, MonadReduce m) => a -> m a
instantiate (MetaId -> Elims -> Term
MetaV MetaId
x Elims
es)))
orelse :: m () -> m () -> m ()
orelse :: m () -> m () -> m ()
orelse m ()
m m ()
h = m () -> (TCErr -> m ()) -> m ()
forall e (m :: * -> *) a.
MonadError e m =>
m a -> (e -> m a) -> m a
catchError m ()
m (\TCErr
_ -> m ()
h)
assignE :: (MonadConversion m)
=> CompareDirection -> MetaId -> Elims -> Term -> CompareAs -> (Term -> Term -> m ()) -> m ()
assignE :: CompareDirection
-> MetaId
-> Elims
-> Term
-> CompareAs
-> (Term -> Term -> m ())
-> m ()
assignE CompareDirection
dir MetaId
x Elims
es Term
v CompareAs
a Term -> Term -> m ()
comp = CompareDirection -> MetaId -> Elims -> Term -> m () -> m ()
forall (m :: * -> *).
(MonadMetaSolver m, MonadConstraint m, MonadError TCErr m,
MonadDebug m, HasOptions m) =>
CompareDirection -> MetaId -> Elims -> Term -> m () -> m ()
assignWrapper CompareDirection
dir MetaId
x Elims
es Term
v (m () -> m ()) -> m () -> m ()
forall a b. (a -> b) -> a -> b
$ do
case Elims -> Maybe [Arg Term]
forall a. [Elim' a] -> Maybe [Arg a]
allApplyElims Elims
es of
Just [Arg Term]
vs -> CompareDirection
-> MetaId -> [Arg Term] -> Term -> CompareAs -> m ()
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
VerboseKey -> Int -> TCM Doc -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCM Doc -> m ()
reportSDoc VerboseKey
"tc.conv.assign" Int
30 (TCM Doc -> m ()) -> TCM Doc -> m ()
forall a b. (a -> b) -> a -> b
$ [TCM Doc] -> TCM Doc
forall (m :: * -> *). Monad m => [m Doc] -> m Doc
sep
[ TCM Doc
"assigning to projected meta "
, MetaId -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM MetaId
x TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> [TCM Doc] -> TCM Doc
forall (m :: * -> *). Monad m => [m Doc] -> m Doc
sep ((Elim' Term -> TCM Doc) -> Elims -> [TCM Doc]
forall a b. (a -> b) -> [a] -> [b]
map Elim' Term -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Elims
es) TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> VerboseKey -> TCM Doc
forall (m :: * -> *). Monad m => VerboseKey -> m Doc
text (VerboseKey
":" VerboseKey -> VerboseKey -> VerboseKey
forall a. [a] -> [a] -> [a]
++ CompareDirection -> VerboseKey
forall a. Show a => a -> VerboseKey
show CompareDirection
dir) TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Term -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Term
v
]
[MetaKind] -> MetaId -> m ()
forall (m :: * -> *).
MonadMetaSolver m =>
[MetaKind] -> MetaId -> m ()
etaExpandMeta [MetaKind
Records] MetaId
x
Maybe Term
res <- MetaId -> m (Maybe Term)
forall (m :: * -> *).
(MonadFail m, ReadTCState m) =>
MetaId -> m (Maybe Term)
isInstantiatedMeta' MetaId
x
case Maybe Term
res of
Just Term
u -> do
VerboseKey -> Int -> TCM Doc -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCM Doc -> m ()
reportSDoc VerboseKey
"tc.conv.assign" Int
30 (TCM Doc -> m ()) -> TCM Doc -> m ()
forall a b. (a -> b) -> a -> b
$ [TCM Doc] -> TCM Doc
forall (m :: * -> *). Monad m => [m Doc] -> m Doc
sep
[ TCM Doc
"seems like eta expansion instantiated meta "
, MetaId -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM MetaId
x TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> VerboseKey -> TCM Doc
forall (m :: * -> *). Monad m => VerboseKey -> m Doc
text (VerboseKey
":" VerboseKey -> VerboseKey -> VerboseKey
forall a. [a] -> [a] -> [a]
++ CompareDirection -> VerboseKey
forall a. Show a => a -> VerboseKey
show CompareDirection
dir) TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Term -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Term
u
]
let w :: Term
w = Term
u Term -> Elims -> Term
forall t. Apply t => t -> Elims -> t
`applyE` Elims
es
Term -> Term -> m ()
comp Term
w Term
v
Maybe Term
Nothing -> do
VerboseKey -> Int -> VerboseKey -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> VerboseKey -> m ()
reportSLn VerboseKey
"tc.conv.assign" Int
30 VerboseKey
"eta expansion did not instantiate meta"
m ()
forall (m :: * -> *) a. MonadError TCErr m => m a
patternViolation
compareAsDir :: MonadConversion m => CompareDirection -> CompareAs -> Term -> Term -> m ()
compareAsDir :: CompareDirection -> CompareAs -> Term -> Term -> m ()
compareAsDir CompareDirection
dir CompareAs
a = (Comparison -> Term -> Term -> m ())
-> CompareDirection -> Term -> Term -> m ()
forall a c.
(Comparison -> a -> a -> c) -> CompareDirection -> a -> a -> c
dirToCmp (Comparison -> CompareAs -> Term -> Term -> m ()
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' :: Comparison -> CompareAs -> Term -> Term -> m ()
compareAs' Comparison
cmp CompareAs
tt Term
m Term
n = case CompareAs
tt of
AsTermsOf Type
a -> Comparison -> Type -> Term -> Term -> m ()
forall (m :: * -> *).
MonadConversion m =>
Comparison -> Type -> Term -> Term -> m ()
compareTerm' Comparison
cmp Type
a Term
m Term
n
CompareAs
AsSizes -> Comparison -> Term -> Term -> m ()
forall (m :: * -> *).
MonadConversion m =>
Comparison -> Term -> Term -> m ()
compareSizes Comparison
cmp Term
m Term
n
CompareAs
AsTypes -> Comparison -> CompareAs -> Term -> Term -> m ()
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' :: Comparison -> Type -> Term -> Term -> m ()
compareTerm' Comparison
cmp Type
a Term
m Term
n =
VerboseKey -> Int -> VerboseKey -> m () -> m ()
forall (m :: * -> *) a.
MonadDebug m =>
VerboseKey -> Int -> VerboseKey -> m a -> m a
verboseBracket VerboseKey
"tc.conv.term" Int
20 VerboseKey
"compareTerm" (m () -> m ()) -> m () -> m ()
forall a b. (a -> b) -> a -> b
$ do
Type
a' <- Type -> m Type
forall a (m :: * -> *). (Reduce a, MonadReduce m) => a -> m a
reduce Type
a
(Constraint -> m () -> m ()
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 ()) (m () -> m ()) -> m () -> m ()
forall a b. (a -> b) -> a -> b
$ do
VerboseKey -> Int -> TCM Doc -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCM Doc -> m ()
reportSDoc VerboseKey
"tc.conv.term" Int
30 (TCM Doc -> m ()) -> TCM Doc -> m ()
forall a b. (a -> b) -> a -> b
$ [TCM Doc] -> TCM Doc
forall (m :: * -> *). Monad m => [m Doc] -> m Doc
fsep
[ TCM Doc
"compareTerm", Term -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Term
m, Comparison -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Comparison
cmp, Term -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Term
n, TCM Doc
":", Type -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Type
a' ]
Bool
propIrr <- m Bool
forall (m :: * -> *). HasOptions m => m Bool
isPropEnabled
Bool
isSize <- Maybe BoundedSize -> Bool
forall a. Maybe a -> Bool
isJust (Maybe BoundedSize -> Bool) -> m (Maybe BoundedSize) -> m Bool
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> Type -> m (Maybe BoundedSize)
forall a (m :: * -> *).
(IsSizeType a, HasOptions m, HasBuiltins m) =>
a -> m (Maybe BoundedSize)
isSizeType Type
a'
Sort
s <- Sort -> m Sort
forall a (m :: * -> *). (Reduce a, MonadReduce m) => a -> m a
reduce (Sort -> m Sort) -> Sort -> m Sort
forall a b. (a -> b) -> a -> b
$ Type -> Sort
forall a. LensSort a => a -> Sort
getSort Type
a'
Maybe Term
mlvl <- VerboseKey -> m (Maybe Term)
forall (m :: * -> *). HasBuiltins m => VerboseKey -> m (Maybe Term)
getBuiltin' VerboseKey
builtinLevel
VerboseKey -> Int -> TCM Doc -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCM Doc -> m ()
reportSDoc VerboseKey
"tc.conv.level" Int
60 (TCM Doc -> m ()) -> TCM Doc -> m ()
forall a b. (a -> b) -> a -> b
$ Int -> TCM Doc -> TCM Doc
forall (m :: * -> *). Functor m => Int -> m Doc -> m Doc
nest Int
2 (TCM Doc -> TCM Doc) -> TCM Doc -> TCM Doc
forall a b. (a -> b) -> a -> b
$ [TCM Doc] -> TCM Doc
forall (m :: * -> *). Monad m => [m Doc] -> m Doc
sep
[ TCM Doc
"a' =" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Type -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty Type
a'
, TCM Doc
"mlvl =" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Maybe Term -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty Maybe Term
mlvl
, VerboseKey -> TCM Doc
forall (m :: * -> *). Monad m => VerboseKey -> m Doc
text (VerboseKey -> TCM Doc) -> VerboseKey -> TCM Doc
forall a b. (a -> b) -> a -> b
$ VerboseKey
"(Just (unEl a') == mlvl) = " VerboseKey -> VerboseKey -> VerboseKey
forall a. [a] -> [a] -> [a]
++ Bool -> VerboseKey
forall a. Show a => a -> VerboseKey
show (Term -> Maybe Term
forall a. a -> Maybe a
Just (Type -> Term
forall t a. Type'' t a -> a
unEl Type
a') Maybe Term -> Maybe Term -> Bool
forall a. Eq a => a -> a -> Bool
== Maybe Term
mlvl)
]
case Sort
s of
Prop{} | Bool
propIrr -> Type -> Term -> Term -> m ()
forall (m :: * -> *).
MonadConversion m =>
Type -> Term -> Term -> m ()
compareIrrelevant Type
a' Term
m Term
n
Sort
_ | Bool
isSize -> Comparison -> Term -> Term -> m ()
forall (m :: * -> *).
MonadConversion m =>
Comparison -> Term -> Term -> m ()
compareSizes Comparison
cmp Term
m Term
n
Sort
_ -> case Type -> Term
forall t a. Type'' t a -> a
unEl Type
a' of
Term
a | Term -> Maybe Term
forall a. a -> Maybe a
Just Term
a Maybe Term -> Maybe Term -> Bool
forall a. Eq a => a -> a -> Bool
== Maybe Term
mlvl -> do
Level
a <- Term -> m Level
forall (m :: * -> *).
(HasBuiltins m, MonadReduce m, MonadDebug m) =>
Term -> m Level
levelView Term
m
Level
b <- Term -> m Level
forall (m :: * -> *).
(HasBuiltins m, MonadReduce m, MonadDebug m) =>
Term -> m Level
levelView Term
n
Level -> Level -> m ()
forall (m :: * -> *). MonadConversion m => Level -> Level -> m ()
equalLevel Level
a Level
b
a :: Term
a@Pi{} -> MonadConversion m => Sort -> Term -> Term -> Term -> m ()
Sort -> Term -> Term -> Term -> m ()
equalFun Sort
s Term
a Term
m Term
n
Lam ArgInfo
_ Abs Term
_ -> m ()
forall a. HasCallStack => a
__IMPOSSIBLE__
Def QName
r Elims
es -> do
Bool
isrec <- QName -> m Bool
forall (m :: * -> *). HasConstInfo m => QName -> m Bool
isEtaRecord QName
r
if Bool
isrec
then do
Signature
sig <- m Signature
forall (m :: * -> *). ReadTCState m => m Signature
getSignature
let ps :: [Arg Term]
ps = [Arg Term] -> Maybe [Arg Term] -> [Arg Term]
forall a. a -> Maybe a -> a
fromMaybe [Arg Term]
forall a. HasCallStack => a
__IMPOSSIBLE__ (Maybe [Arg Term] -> [Arg Term]) -> Maybe [Arg Term] -> [Arg Term]
forall a b. (a -> b) -> a -> b
$ Elims -> Maybe [Arg Term]
forall a. [Elim' a] -> Maybe [Arg a]
allApplyElims Elims
es
isNeutral :: Blocked Term -> m Bool
isNeutral (NotBlocked NotBlocked
_ Con{}) = Bool -> m Bool
forall (m :: * -> *) a. Monad m => a -> m a
return Bool
False
isNeutral (NotBlocked NotBlocked
r (Def QName
q Elims
_)) = do
Bool -> Bool
not (Bool -> Bool) -> m Bool -> m Bool
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> QName -> m Bool
forall (m :: * -> *). HasConstInfo m => QName -> m Bool
usesCopatterns QName
q
isNeutral Blocked Term
_ = Bool -> m Bool
forall (m :: * -> *) a. Monad m => a -> m a
return Bool
True
isMeta :: Blocked Term -> Bool
isMeta (NotBlocked NotBlocked
_ MetaV{}) = Bool
True
isMeta Blocked Term
_ = Bool
False
VerboseKey -> Int -> TCM Doc -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCM Doc -> m ()
reportSDoc VerboseKey
"tc.conv.term" Int
30 (TCM Doc -> m ()) -> TCM Doc -> m ()
forall a b. (a -> b) -> a -> b
$ Type -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Type
a TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> TCM Doc
"is eta record type"
Blocked Term
m <- Term -> m (Blocked Term)
forall a (m :: * -> *).
(Reduce a, MonadReduce m) =>
a -> m (Blocked a)
reduceB Term
m
Bool
mNeutral <- Blocked Term -> m Bool
forall (m :: * -> *). HasConstInfo m => Blocked Term -> m Bool
isNeutral Blocked Term
m
Blocked Term
n <- Term -> m (Blocked Term)
forall a (m :: * -> *).
(Reduce a, MonadReduce m) =>
a -> m (Blocked a)
reduceB Term
n
Bool
nNeutral <- Blocked Term -> m Bool
forall (m :: * -> *). HasConstInfo m => Blocked Term -> m Bool
isNeutral Blocked Term
n
case (Blocked Term
m, Blocked Term
n) of
(Blocked Term, Blocked Term)
_ | Blocked Term -> Bool
isMeta Blocked Term
m Bool -> Bool -> Bool
|| Blocked Term -> Bool
isMeta Blocked Term
n ->
Comparison -> CompareAs -> Term -> Term -> m ()
forall (m :: * -> *).
MonadConversion m =>
Comparison -> CompareAs -> Term -> Term -> m ()
compareAtom Comparison
cmp (Type -> CompareAs
AsTermsOf Type
a') (Blocked Term -> Term
forall t. Blocked t -> t
ignoreBlocking Blocked Term
m) (Blocked Term -> Term
forall t. Blocked t -> t
ignoreBlocking Blocked Term
n)
(Blocked Term, Blocked Term)
_ | Bool
mNeutral Bool -> Bool -> Bool
&& Bool
nNeutral -> do
Either MetaId Bool
isSing <- QName -> [Arg Term] -> m (Either MetaId Bool)
forall (m :: * -> *).
(MonadReduce m, MonadAddContext m, HasConstInfo m, HasBuiltins m,
ReadTCState m) =>
QName -> [Arg Term] -> m (Either MetaId Bool)
isSingletonRecordModuloRelevance QName
r [Arg Term]
ps
case Either MetaId Bool
isSing of
Right Bool
True -> () -> m ()
forall (m :: * -> *) a. Monad m => a -> m a
return ()
Either MetaId Bool
_ -> Comparison -> CompareAs -> Term -> Term -> m ()
forall (m :: * -> *).
MonadConversion m =>
Comparison -> CompareAs -> Term -> Term -> m ()
compareAtom Comparison
cmp (Type -> CompareAs
AsTermsOf Type
a') (Blocked Term -> Term
forall t. Blocked t -> t
ignoreBlocking Blocked Term
m) (Blocked Term -> Term
forall t. Blocked t -> t
ignoreBlocking Blocked Term
n)
(Blocked Term, Blocked Term)
_ -> do
(Telescope
tel, [Arg Term]
m') <- QName -> [Arg Term] -> Term -> m (Telescope, [Arg Term])
forall (m :: * -> *).
(HasConstInfo m, MonadDebug m, ReadTCState m,
MonadError TCErr m) =>
QName -> [Arg Term] -> Term -> m (Telescope, [Arg Term])
etaExpandRecord QName
r [Arg Term]
ps (Term -> m (Telescope, [Arg Term]))
-> Term -> m (Telescope, [Arg Term])
forall a b. (a -> b) -> a -> b
$ Blocked Term -> Term
forall t. Blocked t -> t
ignoreBlocking Blocked Term
m
(Telescope
_ , [Arg Term]
n') <- QName -> [Arg Term] -> Term -> m (Telescope, [Arg Term])
forall (m :: * -> *).
(HasConstInfo m, MonadDebug m, ReadTCState m,
MonadError TCErr m) =>
QName -> [Arg Term] -> Term -> m (Telescope, [Arg Term])
etaExpandRecord QName
r [Arg Term]
ps (Term -> m (Telescope, [Arg Term]))
-> Term -> m (Telescope, [Arg Term])
forall a b. (a -> b) -> a -> b
$ Blocked Term -> Term
forall t. Blocked t -> t
ignoreBlocking Blocked Term
n
ConHead
c <- QName -> m ConHead
forall (m :: * -> *).
(HasConstInfo m, ReadTCState m, MonadError TCErr m) =>
QName -> m ConHead
getRecordConstructor QName
r
[Polarity]
-> [IsForced] -> Type -> Term -> [Arg Term] -> [Arg Term] -> m ()
forall (m :: * -> *).
MonadConversion m =>
[Polarity]
-> [IsForced] -> Type -> Term -> [Arg Term] -> [Arg Term] -> m ()
compareArgs (Polarity -> [Polarity]
forall a. a -> [a]
repeat (Polarity -> [Polarity]) -> Polarity -> [Polarity]
forall a b. (a -> b) -> a -> b
$ Comparison -> Polarity
polFromCmp Comparison
cmp) [] (Telescope -> Type -> Type
telePi_ Telescope
tel Type
HasCallStack => Type
__DUMMY_TYPE__) (ConHead -> ConInfo -> Elims -> Term
Con ConHead
c ConInfo
ConOSystem []) [Arg Term]
m' [Arg Term]
n'
else (do PathView
pathview <- Type -> m PathView
forall (m :: * -> *). HasBuiltins m => Type -> m PathView
pathView Type
a'
MonadConversion m => PathView -> Type -> Term -> Term -> m ()
PathView -> Type -> Term -> Term -> m ()
equalPath PathView
pathview Type
a' Term
m Term
n)
Term
_ -> Comparison -> CompareAs -> Term -> Term -> m ()
forall (m :: * -> *).
MonadConversion m =>
Comparison -> CompareAs -> Term -> Term -> m ()
compareAtom Comparison
cmp (Type -> CompareAs
AsTermsOf Type
a') Term
m Term
n
where
equalFun :: (MonadConversion m) => Sort -> Term -> Term -> Term -> m ()
equalFun :: Sort -> Term -> Term -> Term -> m ()
equalFun Sort
s a :: Term
a@(Pi Dom Type
dom Abs Type
b) Term
m Term
n | Dom Type -> Bool
forall t e. Dom' t e -> Bool
domFinite Dom Type
dom = do
Maybe QName
mp <- (Term -> QName) -> Maybe Term -> Maybe QName
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
fmap Term -> QName
getPrimName (Maybe Term -> Maybe QName) -> m (Maybe Term) -> m (Maybe QName)
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> VerboseKey -> m (Maybe Term)
forall (m :: * -> *). HasBuiltins m => VerboseKey -> m (Maybe Term)
getBuiltin' VerboseKey
builtinIsOne
case Type -> Term
forall t a. Type'' t a -> a
unEl (Type -> Term) -> Type -> Term
forall a b. (a -> b) -> a -> b
$ Dom Type -> Type
forall t e. Dom' t e -> e
unDom Dom Type
dom of
Def QName
q [Apply Arg Term
phi]
| QName -> Maybe QName
forall a. a -> Maybe a
Just QName
q Maybe QName -> Maybe QName -> Bool
forall a. Eq a => a -> a -> Bool
== Maybe QName
mp -> Comparison -> Term -> Type -> Term -> Term -> m ()
forall (m :: * -> *).
MonadConversion m =>
Comparison -> Term -> Type -> Term -> Term -> m ()
compareTermOnFace Comparison
cmp (Arg Term -> Term
forall e. Arg e -> e
unArg Arg Term
phi) (Sort -> Term -> Type
forall t a. Sort' t -> a -> Type'' t a
El Sort
s (Dom Type -> Abs Type -> Term
Pi (Dom Type
dom {domFinite :: Bool
domFinite = Bool
False}) Abs Type
b)) Term
m Term
n
Term
_ -> MonadConversion m => Sort -> Term -> Term -> Term -> m ()
Sort -> Term -> Term -> Term -> m ()
equalFun Sort
s (Dom Type -> Abs Type -> Term
Pi (Dom Type
dom{domFinite :: Bool
domFinite = Bool
False}) Abs Type
b) 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 | Bool -> Bool
not (Bool -> Bool) -> Bool -> Bool
forall a b. (a -> b) -> a -> b
$ Dom Type -> Bool
forall t e. Dom' t e -> Bool
domFinite Dom Type
dom = do
let name :: VerboseKey
name = [Suggestion] -> VerboseKey
suggests [ Abs Type -> Suggestion
forall a. Suggest a => a -> Suggestion
Suggestion Abs Type
b , Term -> Suggestion
forall a. Suggest a => a -> Suggestion
Suggestion Term
m , Term -> Suggestion
forall a. Suggest a => a -> Suggestion
Suggestion Term
n ]
(VerboseKey, Dom Type) -> m () -> m ()
forall b (m :: * -> *) a.
(AddContext b, MonadAddContext m) =>
b -> m a -> m a
addContext (VerboseKey
name, Dom Type
dom) (m () -> m ()) -> m () -> m ()
forall a b. (a -> b) -> a -> b
$ Comparison -> Type -> Term -> Term -> m ()
forall (m :: * -> *).
MonadConversion m =>
Comparison -> Type -> Term -> Term -> m ()
compareTerm Comparison
cmp (Abs Type -> Type
forall t a. Subst t a => Abs a -> a
absBody Abs Type
b) Term
m' Term
n'
where
(Term
m',Term
n') = Int -> (Term, Term) -> (Term, Term)
forall t a. Subst t a => Int -> a -> a
raise Int
1 (Term
m,Term
n) (Term, Term) -> [Arg Term] -> (Term, Term)
forall t. Apply t => t -> [Arg Term] -> t
`apply` [ArgInfo -> Term -> Arg Term
forall e. ArgInfo -> e -> Arg e
Arg ArgInfo
info (Term -> Arg Term) -> Term -> Arg Term
forall a b. (a -> b) -> a -> b
$ Int -> Term
var Int
0]
equalFun Sort
_ Term
_ Term
_ Term
_ = m ()
forall a. HasCallStack => a
__IMPOSSIBLE__
equalPath :: (MonadConversion m) => PathView -> Type -> Term -> Term -> m ()
equalPath :: 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
let name :: VerboseKey
name = VerboseKey
"i" :: String
Type
interval <- m Term -> m Type
forall (tcm :: * -> *). Functor tcm => tcm Term -> tcm Type
el m Term
forall (m :: * -> *).
(HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) =>
m Term
primInterval
let (Term
m',Term
n') = Int -> (Term, Term) -> (Term, Term)
forall t a. Subst t a => Int -> a -> a
raise Int
1 (Term
m, Term
n) (Term, Term) -> Elims -> (Term, Term)
forall t. Apply t => t -> Elims -> t
`applyE` [Term -> Term -> Term -> Elim' Term
forall a. a -> a -> a -> Elim' a
IApply (Int -> Term -> Term
forall t a. Subst t a => Int -> a -> a
raise Int
1 (Term -> Term) -> Term -> Term
forall a b. (a -> b) -> a -> b
$ Arg Term -> Term
forall e. Arg e -> e
unArg Arg Term
x) (Int -> Term -> Term
forall t a. Subst t a => Int -> a -> a
raise Int
1 (Term -> Term) -> Term -> Term
forall a b. (a -> b) -> a -> b
$ Arg Term -> Term
forall e. Arg e -> e
unArg Arg Term
y) (Int -> Term
var Int
0)]
(VerboseKey, Dom Type) -> m () -> m ()
forall b (m :: * -> *) a.
(AddContext b, MonadAddContext m) =>
b -> m a -> m a
addContext (VerboseKey
name, Type -> Dom Type
forall a. a -> Dom a
defaultDom Type
interval) (m () -> m ()) -> m () -> m ()
forall a b. (a -> b) -> a -> b
$ Comparison -> Type -> Term -> Term -> m ()
forall (m :: * -> *).
MonadConversion m =>
Comparison -> Type -> Term -> Term -> m ()
compareTerm Comparison
cmp (Sort -> Term -> Type
forall t a. Sort' t -> a -> Type'' t a
El (Int -> Sort -> Sort
forall t a. Subst t a => Int -> a -> a
raise Int
1 Sort
s) (Term -> Type) -> Term -> Type
forall a b. (a -> b) -> a -> b
$ (Int -> Term -> Term
forall t a. Subst t a => Int -> a -> a
raise Int
1 (Term -> Term) -> Term -> Term
forall a b. (a -> b) -> a -> b
$ Arg Term -> Term
forall e. Arg e -> e
unArg Arg Term
a) Term -> [Arg Term] -> Term
forall t. Apply t => t -> [Arg Term] -> t
`apply` [Term -> Arg Term
forall e. e -> Arg e
argN (Term -> Arg Term) -> Term -> Arg Term
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 ()
forall (m :: * -> *).
(MonadMetaSolver m, MonadWarning m, MonadStatistics m,
MonadFresh ProblemId m, MonadFresh Int m) =>
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 <- VerboseKey -> m (Maybe QName)
forall (m :: * -> *).
HasBuiltins m =>
VerboseKey -> m (Maybe QName)
getBuiltinName' VerboseKey
builtinInterval
Maybe QName
mIsOne <- VerboseKey -> m (Maybe QName)
forall (m :: * -> *).
HasBuiltins m =>
VerboseKey -> m (Maybe QName)
getBuiltinName' VerboseKey
builtinIsOne
Maybe QName
mGlue <- VerboseKey -> m (Maybe QName)
forall (m :: * -> *).
HasBuiltins m =>
VerboseKey -> m (Maybe QName)
getPrimitiveName' VerboseKey
builtinGlue
Maybe QName
mHComp <- VerboseKey -> m (Maybe QName)
forall (m :: * -> *).
HasBuiltins m =>
VerboseKey -> m (Maybe QName)
getPrimitiveName' VerboseKey
builtinHComp
Maybe QName
mSub <- VerboseKey -> m (Maybe QName)
forall (m :: * -> *).
HasBuiltins m =>
VerboseKey -> m (Maybe QName)
getBuiltinName' VerboseKey
builtinSub
case Term
ty of
Def QName
q Elims
es | QName -> Maybe QName
forall a. a -> Maybe a
Just QName
q Maybe QName -> Maybe QName -> Bool
forall a. Eq a => a -> a -> Bool
== Maybe QName
mIsOne -> () -> m ()
forall (m :: * -> *) a. Monad m => a -> m a
return ()
Def QName
q Elims
es | QName -> Maybe QName
forall a. a -> Maybe a
Just QName
q Maybe QName -> Maybe QName -> Bool
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]
_) <- Elims -> Maybe [Arg Term]
forall a. [Elim' a] -> Maybe [Arg a]
allApplyElims Elims
es -> do
Type
ty <- m Term -> m Term -> m Type
forall (m :: * -> *). Monad m => m Term -> m Term -> m Type
el' (Term -> m Term
forall (f :: * -> *) a. Applicative f => a -> f a
pure (Term -> m Term) -> Term -> m Term
forall a b. (a -> b) -> a -> b
$ Arg Term -> Term
forall e. Arg e -> e
unArg Arg Term
l) (Term -> m Term
forall (f :: * -> *) a. Applicative f => a -> f a
pure (Term -> m Term) -> Term -> m Term
forall a b. (a -> b) -> a -> b
$ Arg Term -> Term
forall e. Arg e -> e
unArg Arg Term
a)
Term
unglue <- m Term
forall (m :: * -> *).
(HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) =>
m Term
prim_unglue
let mkUnglue :: Term -> Term
mkUnglue Term
m = Term -> [Arg Term] -> Term
forall t. Apply t => t -> [Arg Term] -> t
apply Term
unglue ([Arg Term] -> Term) -> [Arg Term] -> Term
forall a b. (a -> b) -> a -> b
$ (Arg Term -> Arg Term) -> [Arg Term] -> [Arg Term]
forall a b. (a -> b) -> [a] -> [b]
map (Hiding -> Arg Term -> Arg Term
forall a. LensHiding a => Hiding -> a -> a
setHiding Hiding
Hidden) [Arg Term]
args [Arg Term] -> [Arg Term] -> [Arg Term]
forall a. [a] -> [a] -> [a]
++ [Term -> Arg Term
forall e. e -> Arg e
argN Term
m]
VerboseKey -> Int -> TCM Doc -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCM Doc -> m ()
reportSDoc VerboseKey
"conv.glue" Int
20 (TCM Doc -> m ()) -> TCM Doc -> m ()
forall a b. (a -> b) -> a -> b
$ (Type, Term, Term) -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM (Type
ty,Term -> Term
mkUnglue Term
m,Term -> Term
mkUnglue Term
n)
Comparison -> Term -> Type -> Term -> Term -> m ()
forall (m :: * -> *).
MonadConversion m =>
Comparison -> Term -> Type -> Term -> Term -> m ()
compareTermOnFace Comparison
cmp (Arg Term -> Term
forall e. Arg e -> e
unArg Arg Term
phi) Type
ty Term
m Term
n
Comparison -> Type -> Term -> Term -> m ()
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 | QName -> Maybe QName
forall a. a -> Maybe a
Just QName
q Maybe QName -> Maybe QName -> Bool
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]) <- Elims -> Maybe [Arg Term]
forall a. [Elim' a] -> Maybe [Arg a]
allApplyElims Elims
es
, Sort (Type Level
lvl) <- Arg Term -> Term
forall e. Arg e -> e
unArg Arg Term
s -> do
let l :: Term
l = Level -> Term
Level Level
lvl
Type
ty <- m Term -> m Term -> m Type
forall (m :: * -> *). Monad m => m Term -> m Term -> m Type
el' (Term -> m Term
forall (f :: * -> *) a. Applicative f => a -> f a
pure (Term -> m Term) -> Term -> m Term
forall a b. (a -> b) -> a -> b
$ Term
l) (Term -> m Term
forall (f :: * -> *) a. Applicative f => a -> f a
pure (Term -> m Term) -> Term -> m Term
forall a b. (a -> b) -> a -> b
$ Arg Term -> Term
forall e. Arg e -> e
unArg Arg Term
u0)
Term
unglueU <- m Term
forall (m :: * -> *).
(HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) =>
m Term
prim_unglueU
Term
subIn <- m Term
forall (m :: * -> *).
(HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) =>
m Term
primSubIn
let bA :: Term
bA = Term
subIn Term -> [Arg Term] -> Term
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 = Term -> [Arg Term] -> Term
forall t. Apply t => t -> [Arg Term] -> t
apply Term
unglueU ([Arg Term] -> Term) -> [Arg Term] -> Term
forall a b. (a -> b) -> a -> b
$ [Term -> Arg Term
forall e. e -> Arg e
argH Term
l] [Arg Term] -> [Arg Term] -> [Arg Term]
forall a. [a] -> [a] -> [a]
++ (Arg Term -> Arg Term) -> [Arg Term] -> [Arg Term]
forall a b. (a -> b) -> [a] -> [b]
map (Hiding -> Arg Term -> Arg Term
forall a. LensHiding a => Hiding -> a -> a
setHiding Hiding
Hidden) [Arg Term
phi,Arg Term
u] [Arg Term] -> [Arg Term] -> [Arg Term]
forall a. [a] -> [a] -> [a]
++ [Term -> Arg Term
forall e. e -> Arg e
argH Term
bA,Term -> Arg Term
forall e. e -> Arg e
argN Term
m]
VerboseKey -> Int -> TCM Doc -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCM Doc -> m ()
reportSDoc VerboseKey
"conv.hcompU" Int
20 (TCM Doc -> m ()) -> TCM Doc -> m ()
forall a b. (a -> b) -> a -> b
$ (Type, Term, Term) -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM (Type
ty,Term -> Term
mkUnglue Term
m,Term -> Term
mkUnglue Term
n)
Comparison -> Term -> Type -> Term -> Term -> m ()
forall (m :: * -> *).
MonadConversion m =>
Comparison -> Term -> Type -> Term -> Term -> m ()
compareTermOnFace Comparison
cmp (Arg Term -> Term
forall e. Arg e -> e
unArg Arg Term
phi) Type
ty Term
m Term
n
Comparison -> Type -> Term -> Term -> m ()
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 | QName -> Maybe QName
forall a. a -> Maybe a
Just QName
q Maybe QName -> Maybe QName -> Bool
forall a. Eq a => a -> a -> Bool
== Maybe QName
mSub, Just args :: [Arg Term]
args@(Arg Term
l:Arg Term
a:[Arg Term]
_) <- Elims -> Maybe [Arg Term]
forall a. [Elim' a] -> Maybe [Arg a]
allApplyElims Elims
es -> do
Type
ty <- m Term -> m Term -> m Type
forall (m :: * -> *). Monad m => m Term -> m Term -> m Type
el' (Term -> m Term
forall (f :: * -> *) a. Applicative f => a -> f a
pure (Term -> m Term) -> Term -> m Term
forall a b. (a -> b) -> a -> b
$ Arg Term -> Term
forall e. Arg e -> e
unArg Arg Term
l) (Term -> m Term
forall (f :: * -> *) a. Applicative f => a -> f a
pure (Term -> m Term) -> Term -> m Term
forall a b. (a -> b) -> a -> b
$ Arg Term -> Term
forall e. Arg e -> e
unArg Arg Term
a)
Term
out <- m Term
forall (m :: * -> *).
(HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) =>
m Term
primSubOut
let mkOut :: Term -> Term
mkOut Term
m = Term -> [Arg Term] -> Term
forall t. Apply t => t -> [Arg Term] -> t
apply Term
out ([Arg Term] -> Term) -> [Arg Term] -> Term
forall a b. (a -> b) -> a -> b
$ (Arg Term -> Arg Term) -> [Arg Term] -> [Arg Term]
forall a b. (a -> b) -> [a] -> [b]
map (Hiding -> Arg Term -> Arg Term
forall a. LensHiding a => Hiding -> a -> a
setHiding Hiding
Hidden) [Arg Term]
args [Arg Term] -> [Arg Term] -> [Arg Term]
forall a. [a] -> [a] -> [a]
++ [Term -> Arg Term
forall e. e -> Arg e
argN Term
m]
Comparison -> Type -> Term -> 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 [] | QName -> Maybe QName
forall a. a -> Maybe a
Just QName
q Maybe QName -> Maybe QName -> Bool
forall a. Eq a => a -> a -> Bool
== Maybe QName
mI -> Comparison -> Type -> Term -> Term -> m ()
forall (m :: * -> *).
MonadConversion m =>
Comparison -> Type -> Term -> Term -> m ()
compareInterval Comparison
cmp Type
a' Term
m Term
n
Term
_ -> Comparison -> CompareAs -> Term -> Term -> m ()
forall (m :: * -> *).
MonadConversion m =>
Comparison -> CompareAs -> Term -> Term -> m ()
compareAtom Comparison
cmp (Type -> CompareAs
AsTermsOf Type
a') Term
m Term
n
compareTel :: MonadConversion m => Type -> Type ->
Comparison -> Telescope -> Telescope -> m ()
compareTel :: Type -> Type -> Comparison -> Telescope -> Telescope -> m ()
compareTel Type
t1 Type
t2 Comparison
cmp Telescope
tel1 Telescope
tel2 =
VerboseKey -> Int -> VerboseKey -> m () -> m ()
forall (m :: * -> *) a.
MonadDebug m =>
VerboseKey -> Int -> VerboseKey -> m a -> m a
verboseBracket VerboseKey
"tc.conv.tel" Int
20 VerboseKey
"compareTel" (m () -> m ()) -> m () -> m ()
forall a b. (a -> b) -> a -> b
$
Constraint -> m () -> m ()
forall (m :: * -> *).
MonadConstraint m =>
Constraint -> m () -> m ()
catchConstraint (Type -> Type -> Comparison -> Telescope -> Telescope -> Constraint
TelCmp Type
t1 Type
t2 Comparison
cmp Telescope
tel1 Telescope
tel2) (m () -> m ()) -> m () -> m ()
forall a b. (a -> b) -> a -> b
$ case (Telescope
tel1, Telescope
tel2) of
(Telescope
EmptyTel, Telescope
EmptyTel) -> () -> m ()
forall (m :: * -> *) a. Monad m => a -> m a
return ()
(Telescope
EmptyTel, Telescope
_) -> m ()
forall a. m a
bad
(Telescope
_, Telescope
EmptyTel) -> m ()
forall a. m a
bad
(ExtendTel Dom Type
dom1 Abs Telescope
tel1, ExtendTel Dom Type
dom2 Abs Telescope
tel2) -> do
Comparison
-> Dom Type
-> Dom Type
-> Abs Telescope
-> Abs Telescope
-> m ()
-> m ()
-> m ()
-> m ()
-> m ()
-> m ()
forall (m :: * -> *) c b.
(MonadConversion m, Free c) =>
Comparison
-> Dom Type
-> Dom Type
-> Abs b
-> Abs c
-> m ()
-> m ()
-> m ()
-> m ()
-> m ()
-> m ()
compareDom Comparison
cmp Dom Type
dom1 Dom Type
dom2 Abs Telescope
tel1 Abs Telescope
tel2 m ()
forall a. m a
bad m ()
forall a. m a
bad m ()
forall a. m a
bad m ()
forall a. m a
bad (m () -> m ()) -> m () -> m ()
forall a b. (a -> b) -> a -> b
$
Type -> Type -> Comparison -> Telescope -> Telescope -> m ()
forall (m :: * -> *).
MonadConversion m =>
Type -> Type -> Comparison -> Telescope -> Telescope -> m ()
compareTel Type
t1 Type
t2 Comparison
cmp (Abs Telescope -> Telescope
forall t a. Subst t a => Abs a -> a
absBody Abs Telescope
tel1) (Abs Telescope -> Telescope
forall t a. Subst t a => Abs a -> a
absBody Abs Telescope
tel2)
where
bad :: m a
bad = TypeError -> m a
forall (m :: * -> *) a.
(MonadTCEnv m, ReadTCState m, MonadError TCErr m) =>
TypeError -> m a
typeError (TypeError -> m a) -> TypeError -> m a
forall a b. (a -> b) -> a -> b
$ Comparison -> Type -> Type -> TypeError
UnequalTypes Comparison
cmp Type
t2 Type
t1
compareAtomDir :: MonadConversion m => CompareDirection -> CompareAs -> Term -> Term -> m ()
compareAtomDir :: CompareDirection -> CompareAs -> Term -> Term -> m ()
compareAtomDir CompareDirection
dir CompareAs
a = (Comparison -> Term -> Term -> m ())
-> CompareDirection -> Term -> Term -> m ()
forall a c.
(Comparison -> a -> a -> c) -> CompareDirection -> a -> a -> c
dirToCmp (Comparison -> CompareAs -> Term -> Term -> m ()
forall (m :: * -> *).
MonadConversion m =>
Comparison -> CompareAs -> Term -> Term -> m ()
`compareAtom` CompareAs
a) CompareDirection
dir
computeElimHeadType :: MonadConversion m => QName -> Elims -> Elims -> m Type
computeElimHeadType :: QName -> Elims -> Elims -> m Type
computeElimHeadType QName
f Elims
es Elims
es' = do
Definition
def <- QName -> m Definition
forall (m :: * -> *). HasConstInfo m => QName -> m Definition
getConstInfo QName
f
if Defn -> Int
projectionArgs (Definition -> Defn
theDef Definition
def) Int -> Int -> Bool
forall a. Ord a => a -> a -> Bool
<= Int
0 then Type -> m Type
forall (m :: * -> *) a. Monad m => a -> m a
return (Type -> m Type) -> Type -> m Type
forall a b. (a -> b) -> a -> b
$ Definition -> Type
defType Definition
def else do
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)
_ -> Arg Term
forall a. HasCallStack => a
__IMPOSSIBLE__
VerboseKey -> Int -> TCM Doc -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCM Doc -> m ()
reportSDoc VerboseKey
"tc.conv.infer" Int
30 (TCM Doc -> m ()) -> TCM Doc -> m ()
forall a b. (a -> b) -> a -> b
$
TCM Doc
"inferring type of internal arg: " TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Arg Term -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Arg Term
arg
Type
targ <- Term -> m Type
forall (m :: * -> *). MonadCheckInternal m => Term -> m Type
infer (Term -> m Type) -> Term -> m Type
forall a b. (a -> b) -> a -> b
$ Arg Term -> Term
forall e. Arg e -> e
unArg Arg Term
arg
VerboseKey -> Int -> TCM Doc -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCM Doc -> m ()
reportSDoc VerboseKey
"tc.conv.infer" Int
30 (TCM Doc -> m ()) -> TCM Doc -> m ()
forall a b. (a -> b) -> a -> b
$
TCM Doc
"inferred type: " TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Type -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Type
targ
m Type -> m (Maybe Type) -> m Type
forall (m :: * -> *) a. Monad m => m a -> m (Maybe a) -> m a
fromMaybeM m Type
forall (m :: * -> *) a. MonadError TCErr m => m a
patternViolation (m (Maybe Type) -> m Type) -> m (Maybe Type) -> m Type
forall a b. (a -> b) -> a -> b
$ QName -> Type -> m (Maybe Type)
forall (m :: * -> *).
(HasConstInfo m, MonadReduce m, MonadDebug m) =>
QName -> Type -> m (Maybe Type)
getDefType QName
f (Type -> m (Maybe Type)) -> m Type -> m (Maybe Type)
forall (m :: * -> *) a b. Monad m => (a -> m b) -> m a -> m b
=<< Type -> m Type
forall a (m :: * -> *). (Reduce a, MonadReduce m) => a -> m a
reduce Type
targ
compareAtom :: forall m. MonadConversion m => Comparison -> CompareAs -> Term -> Term -> m ()
compareAtom :: Comparison -> CompareAs -> Term -> Term -> m ()
compareAtom Comparison
cmp CompareAs
t Term
m Term
n =
VerboseKey -> Int -> VerboseKey -> m () -> m ()
forall (m :: * -> *) a.
MonadDebug m =>
VerboseKey -> Int -> VerboseKey -> m a -> m a
verboseBracket VerboseKey
"tc.conv.atom" Int
20 VerboseKey
"compareAtom" (m () -> m ()) -> m () -> m ()
forall a b. (a -> b) -> a -> b
$
(Constraint -> m () -> m ()
forall (m :: * -> *).
MonadConstraint m =>
Constraint -> m () -> m ()
catchConstraint (Comparison -> CompareAs -> Term -> Term -> Constraint
ValueCmp Comparison
cmp CompareAs
t Term
m Term
n) :: m () -> m ()) (m () -> m ()) -> m () -> m ()
forall a b. (a -> b) -> a -> b
$ do
VerboseKey -> Int -> TCM Doc -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCM Doc -> m ()
reportSDoc VerboseKey
"tc.conv.atom" Int
50 (TCM Doc -> m ()) -> TCM Doc -> m ()
forall a b. (a -> b) -> a -> b
$
TCM Doc
"compareAtom" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> [TCM Doc] -> TCM Doc
forall (m :: * -> *). Monad m => [m Doc] -> m Doc
fsep [ Term -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Term
m TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Comparison -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Comparison
cmp
, Term -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Term
n
, CompareAs -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM CompareAs
t
]
(Blocked Term
mb',Blocked Term
nb') <- m Bool
-> m (Blocked Term, Blocked Term)
-> m (Blocked Term, Blocked Term)
-> m (Blocked Term, Blocked Term)
forall (m :: * -> *) a. Monad m => m Bool -> m a -> m a -> m a
ifM ((TCEnv -> Bool) -> m Bool
forall (m :: * -> *) a. MonadTCEnv m => (TCEnv -> a) -> m a
asksTC TCEnv -> Bool
envCompareBlocked) ((Term -> Blocked Term
forall a. a -> Blocked a
notBlocked (Term -> Blocked Term)
-> (Term -> Blocked Term)
-> (Term, Term)
-> (Blocked Term, Blocked Term)
forall a c b d. (a -> c) -> (b -> d) -> (a, b) -> (c, d)
-*- Term -> Blocked Term
forall a. a -> Blocked a
notBlocked) ((Term, Term) -> (Blocked Term, Blocked Term))
-> m (Term, Term) -> m (Blocked Term, Blocked Term)
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> (Term, Term) -> m (Term, Term)
forall a (m :: * -> *). (Reduce a, MonadReduce m) => a -> m a
reduce (Term
m,Term
n)) (m (Blocked Term, Blocked Term) -> m (Blocked Term, Blocked Term))
-> m (Blocked Term, Blocked Term) -> m (Blocked Term, Blocked Term)
forall a b. (a -> b) -> a -> b
$ do
Blocked Term
mb' <- Blocked Term -> m (Blocked Term)
forall (m :: * -> *) t.
(MonadReduce m, MonadMetaSolver m, Reduce t) =>
Blocked t -> m (Blocked t)
etaExpandBlocked (Blocked Term -> m (Blocked Term))
-> m (Blocked Term) -> m (Blocked Term)
forall (m :: * -> *) a b. Monad m => (a -> m b) -> m a -> m b
=<< Term -> m (Blocked Term)
forall a (m :: * -> *).
(Reduce a, MonadReduce m) =>
a -> m (Blocked a)
reduceB Term
m
Blocked Term
nb' <- Blocked Term -> m (Blocked Term)
forall (m :: * -> *) t.
(MonadReduce m, MonadMetaSolver m, Reduce t) =>
Blocked t -> m (Blocked t)
etaExpandBlocked (Blocked Term -> m (Blocked Term))
-> m (Blocked Term) -> m (Blocked Term)
forall (m :: * -> *) a b. Monad m => (a -> m b) -> m a -> m b
=<< Term -> m (Blocked Term)
forall a (m :: * -> *).
(Reduce a, MonadReduce m) =>
a -> m (Blocked a)
reduceB Term
n
(Blocked Term, Blocked Term) -> m (Blocked Term, Blocked Term)
forall (m :: * -> *) a. Monad m => a -> m a
return (Blocked Term
mb', Blocked Term
nb')
(Blocked Term
mb'', Blocked Term
nb'') <- case (Blocked Term -> Term
forall t. Blocked t -> t
ignoreBlocking Blocked Term
mb', Blocked Term -> Term
forall t. Blocked t -> t
ignoreBlocking Blocked Term
nb') of
(Lit Literal
_, Lit Literal
_) -> (Blocked Term, Blocked Term) -> m (Blocked Term, Blocked Term)
forall (m :: * -> *) a. Monad m => a -> m a
return (Blocked Term
mb', Blocked Term
nb')
(Term, Term)
_ -> (,) (Blocked Term -> Blocked Term -> (Blocked Term, Blocked Term))
-> m (Blocked Term)
-> m (Blocked Term -> (Blocked Term, Blocked Term))
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> (Term -> m Term) -> Blocked Term -> m (Blocked Term)
forall (t :: * -> *) (f :: * -> *) a b.
(Traversable t, Applicative f) =>
(a -> f b) -> t a -> f (t b)
traverse Term -> m Term
forall (m :: * -> *).
(HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) =>
Term -> m Term
constructorForm Blocked Term
mb'
m (Blocked Term -> (Blocked Term, Blocked Term))
-> m (Blocked Term) -> m (Blocked Term, Blocked Term)
forall (f :: * -> *) a b. Applicative f => f (a -> b) -> f a -> f b
<*> (Term -> m Term) -> Blocked Term -> m (Blocked Term)
forall (t :: * -> *) (f :: * -> *) a b.
(Traversable t, Applicative f) =>
(a -> f b) -> t a -> f (t b)
traverse Term -> m Term
forall (m :: * -> *).
(HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) =>
Term -> m Term
constructorForm Blocked Term
nb'
Blocked Term
mb <- (Term -> m Term) -> Blocked Term -> m (Blocked Term)
forall (t :: * -> *) (f :: * -> *) a b.
(Traversable t, Applicative f) =>
(a -> f b) -> t a -> f (t b)
traverse Term -> m Term
forall (m :: * -> *). HasBuiltins m => Term -> m Term
unLevel Blocked Term
mb''
Blocked Term
nb <- (Term -> m Term) -> Blocked Term -> m (Blocked Term)
forall (t :: * -> *) (f :: * -> *) a b.
(Traversable t, Applicative f) =>
(a -> f b) -> t a -> f (t b)
traverse Term -> m Term
forall (m :: * -> *). HasBuiltins m => Term -> m Term
unLevel Blocked Term
nb''
Bool
cmpBlocked <- Lens' Bool TCEnv -> m Bool
forall (m :: * -> *) a. MonadTCEnv m => Lens' a TCEnv -> m a
viewTC Lens' Bool TCEnv
eCompareBlocked
let m :: Term
m = Blocked Term -> Term
forall t. Blocked t -> t
ignoreBlocking Blocked Term
mb
n :: Term
n = Blocked Term -> Term
forall t. Blocked t -> t
ignoreBlocking Blocked Term
nb
postpone :: m ()
postpone = Constraint -> m ()
forall (m :: * -> *). MonadConstraint m => Constraint -> m ()
addConstraint (Constraint -> m ()) -> Constraint -> m ()
forall a b. (a -> b) -> a -> b
$ Comparison -> CompareAs -> Term -> Term -> Constraint
ValueCmp Comparison
cmp CompareAs
t Term
m Term
n
postponeIfBlockedAs :: CompareAs -> (Blocked CompareAs -> m ()) -> m ()
postponeIfBlockedAs :: CompareAs -> (Blocked CompareAs -> m ()) -> m ()
postponeIfBlockedAs CompareAs
AsTypes Blocked CompareAs -> m ()
f = Blocked CompareAs -> m ()
f (Blocked CompareAs -> m ()) -> Blocked CompareAs -> m ()
forall a b. (a -> b) -> a -> b
$ NotBlocked -> CompareAs -> Blocked CompareAs
forall t. NotBlocked -> t -> Blocked t
NotBlocked NotBlocked
ReallyNotBlocked CompareAs
AsTypes
postponeIfBlockedAs CompareAs
AsSizes Blocked CompareAs -> m ()
f = Blocked CompareAs -> m ()
f (Blocked CompareAs -> m ()) -> Blocked CompareAs -> m ()
forall a b. (a -> b) -> a -> b
$ NotBlocked -> CompareAs -> Blocked CompareAs
forall t. NotBlocked -> t -> Blocked t
NotBlocked NotBlocked
ReallyNotBlocked CompareAs
AsSizes
postponeIfBlockedAs (AsTermsOf Type
t) Blocked CompareAs -> m ()
f = Type
-> (MetaId -> Type -> m ()) -> (NotBlocked -> Type -> m ()) -> m ()
forall t (m :: * -> *) a.
(Reduce t, IsMeta t, MonadReduce m, HasBuiltins m) =>
t -> (MetaId -> t -> m a) -> (NotBlocked -> t -> m a) -> m a
ifBlocked Type
t
(\MetaId
m Type
t -> (Blocked CompareAs -> m ()
f (Blocked CompareAs -> m ()) -> Blocked CompareAs -> m ()
forall a b. (a -> b) -> a -> b
$ MetaId -> CompareAs -> Blocked CompareAs
forall t. MetaId -> t -> Blocked t
Blocked MetaId
m (CompareAs -> Blocked CompareAs) -> CompareAs -> Blocked CompareAs
forall a b. (a -> b) -> a -> b
$ Type -> CompareAs
AsTermsOf Type
t) m () -> (TCErr -> m ()) -> m ()
forall e (m :: * -> *) a.
MonadError e m =>
m a -> (e -> m a) -> m a
`catchError` \case
TypeError{} -> m ()
postpone
TCErr
err -> TCErr -> m ()
forall e (m :: * -> *) a. MonadError e m => e -> m a
throwError TCErr
err)
(\NotBlocked
nb Type
t -> Blocked CompareAs -> m ()
f (Blocked CompareAs -> m ()) -> Blocked CompareAs -> m ()
forall a b. (a -> b) -> a -> b
$ NotBlocked -> CompareAs -> Blocked CompareAs
forall t. NotBlocked -> t -> Blocked t
NotBlocked NotBlocked
nb (CompareAs -> Blocked CompareAs) -> CompareAs -> Blocked CompareAs
forall a b. (a -> b) -> a -> b
$ Type -> CompareAs
AsTermsOf Type
t)
checkDefinitionalEquality :: m ()
checkDefinitionalEquality = m Bool -> m () -> m ()
forall (m :: * -> *). Monad m => m Bool -> m () -> m ()
unlessM (Comparison -> CompareAs -> Term -> Term -> m Bool
forall (m :: * -> *).
(MonadReduce m, MonadAddContext m, HasBuiltins m,
HasConstInfo m) =>
Comparison -> CompareAs -> Term -> Term -> m Bool
pureCompareAs Comparison
CmpEq CompareAs
t Term
m Term
n) m ()
postpone
dir :: CompareDirection
dir = Comparison -> CompareDirection
fromCmp Comparison
cmp
rid :: CompareDirection
rid = CompareDirection -> CompareDirection
flipCmp CompareDirection
dir
assign :: CompareDirection -> MetaId -> Elims -> Term -> m ()
assign CompareDirection
dir MetaId
x Elims
es Term
v = CompareDirection
-> MetaId
-> Elims
-> Term
-> CompareAs
-> (Term -> Term -> m ())
-> m ()
forall (m :: * -> *).
MonadConversion m =>
CompareDirection
-> MetaId
-> Elims
-> Term
-> CompareAs
-> (Term -> Term -> m ())
-> m ()
assignE CompareDirection
dir MetaId
x Elims
es Term
v CompareAs
t ((Term -> Term -> m ()) -> m ()) -> (Term -> Term -> m ()) -> m ()
forall a b. (a -> b) -> a -> b
$ CompareDirection -> CompareAs -> Term -> Term -> m ()
forall (m :: * -> *).
MonadConversion m =>
CompareDirection -> CompareAs -> Term -> Term -> m ()
compareAtomDir CompareDirection
dir CompareAs
t
VerboseKey -> Int -> TCM Doc -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCM Doc -> m ()
reportSDoc VerboseKey
"tc.conv.atom" Int
30 (TCM Doc -> m ()) -> TCM Doc -> m ()
forall a b. (a -> b) -> a -> b
$
TCM Doc
"compareAtom" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> [TCM Doc] -> TCM Doc
forall (m :: * -> *). Monad m => [m Doc] -> m Doc
fsep [ Blocked Term -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Blocked Term
mb TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Comparison -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Comparison
cmp
, Blocked Term -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Blocked Term
nb
, CompareAs -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM CompareAs
t
]
VerboseKey -> Int -> TCM Doc -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCM Doc -> m ()
reportSDoc VerboseKey
"tc.conv.atom" Int
80 (TCM Doc -> m ()) -> TCM Doc -> m ()
forall a b. (a -> b) -> a -> b
$
TCM Doc
"compareAtom" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> [TCM Doc] -> TCM Doc
forall (m :: * -> *). Monad m => [m Doc] -> m Doc
fsep [ (VerboseKey -> TCM Doc
forall (m :: * -> *). Monad m => VerboseKey -> m Doc
text (VerboseKey -> TCM Doc)
-> (Blocked Term -> VerboseKey) -> Blocked Term -> TCM Doc
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Blocked Term -> VerboseKey
forall a. Show a => a -> VerboseKey
show) Blocked Term
mb TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Comparison -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Comparison
cmp
, (VerboseKey -> TCM Doc
forall (m :: * -> *). Monad m => VerboseKey -> m Doc
text (VerboseKey -> TCM Doc)
-> (Blocked Term -> VerboseKey) -> Blocked Term -> TCM Doc
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Blocked Term -> VerboseKey
forall a. Show a => a -> VerboseKey
show) Blocked Term
nb
, TCM Doc
":" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> (VerboseKey -> TCM Doc
forall (m :: * -> *). Monad m => VerboseKey -> m Doc
text (VerboseKey -> TCM Doc)
-> (CompareAs -> VerboseKey) -> CompareAs -> TCM Doc
forall b c a. (b -> c) -> (a -> b) -> a -> c
. CompareAs -> VerboseKey
forall a. Show a => a -> VerboseKey
show) CompareAs
t ]
case (Blocked Term
mb, Blocked Term
nb) of
(NotBlocked NotBlocked
_ (MetaV MetaId
x Elims
xArgs), NotBlocked NotBlocked
_ (MetaV MetaId
y Elims
yArgs))
| MetaId
x MetaId -> MetaId -> Bool
forall a. Eq a => a -> a -> Bool
== MetaId
y , Bool
cmpBlocked -> do
Type
a <- MetaId -> m Type
forall (m :: * -> *).
(MonadFail m, ReadTCState m) =>
MetaId -> m Type
metaType MetaId
x
[Polarity] -> [IsForced] -> Type -> Term -> Elims -> Elims -> m ()
forall (m :: * -> *).
MonadConversion m =>
[Polarity] -> [IsForced] -> Type -> Term -> Elims -> Elims -> m ()
compareElims [] [] Type
a (MetaId -> Elims -> Term
MetaV MetaId
x []) Elims
xArgs Elims
yArgs
| MetaId
x MetaId -> MetaId -> Bool
forall a. Eq a => a -> a -> Bool
== MetaId
y ->
case Elims -> Elims -> Maybe [Bool]
intersectVars Elims
xArgs Elims
yArgs of
Just [Bool]
kills -> do
PruneResult
killResult <- [Bool] -> MetaId -> m PruneResult
forall (m :: * -> *).
MonadMetaSolver m =>
[Bool] -> MetaId -> m PruneResult
killArgs [Bool]
kills MetaId
x
case PruneResult
killResult of
PruneResult
NothingToPrune -> () -> m ()
forall (m :: * -> *) a. Monad m => a -> m a
return ()
PruneResult
PrunedEverything -> () -> m ()
forall (m :: * -> *) a. Monad m => a -> m a
return ()
PruneResult
PrunedNothing -> m ()
postpone
PruneResult
PrunedSomething -> m ()
postpone
Maybe [Bool]
Nothing -> m ()
checkDefinitionalEquality
| Bool
otherwise -> do
[MetaPriority
p1, MetaPriority
p2] <- (MetaId -> m MetaPriority) -> [MetaId] -> m [MetaPriority]
forall (t :: * -> *) (m :: * -> *) a b.
(Traversable t, Monad m) =>
(a -> m b) -> t a -> m (t b)
mapM MetaId -> m MetaPriority
forall (m :: * -> *).
(MonadFail m, ReadTCState m) =>
MetaId -> m MetaPriority
getMetaPriority [MetaId
x,MetaId
y]
let (m ()
solve1, m ()
solve2)
| (MetaPriority
p1, MetaId
x) (MetaPriority, MetaId) -> (MetaPriority, MetaId) -> Bool
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 = CompareDirection -> MetaId -> Elims -> Term -> m ()
forall (m :: * -> *).
(MonadMetaSolver m, MonadWarning m, MonadStatistics m,
MonadFresh ProblemId m, MonadFresh Int m) =>
CompareDirection -> MetaId -> Elims -> Term -> m ()
assign CompareDirection
dir MetaId
x Elims
xArgs Term
n
r1 :: m ()
r1 = CompareDirection -> MetaId -> Elims -> Term -> m ()
forall (m :: * -> *).
(MonadMetaSolver m, MonadWarning m, MonadStatistics m,
MonadFresh ProblemId m, MonadFresh Int m) =>
CompareDirection -> MetaId -> Elims -> Term -> m ()
assign CompareDirection
rid MetaId
y Elims
yArgs Term
m
l2 :: m ()
l2 = m Bool -> m () -> m () -> m ()
forall (m :: * -> *) a. Monad m => m Bool -> m a -> m a -> m a
ifM (MetaId -> m Bool
forall a (m :: * -> *).
(IsInstantiatedMeta a, MonadFail m, ReadTCState m) =>
a -> m Bool
isInstantiatedMeta MetaId
x) (CompareDirection -> CompareAs -> Term -> Term -> m ()
forall (m :: * -> *).
MonadConversion m =>
CompareDirection -> CompareAs -> Term -> Term -> m ()
compareAsDir CompareDirection
dir CompareAs
t Term
m Term
n) m ()
l1
r2 :: m ()
r2 = m Bool -> m () -> m () -> m ()
forall (m :: * -> *) a. Monad m => m Bool -> m a -> m a -> m a
ifM (MetaId -> m Bool
forall a (m :: * -> *).
(IsInstantiatedMeta a, MonadFail m, ReadTCState m) =>
a -> m Bool
isInstantiatedMeta MetaId
y) (CompareDirection -> CompareAs -> Term -> Term -> m ()
forall (m :: * -> *).
MonadConversion m =>
CompareDirection -> CompareAs -> Term -> Term -> m ()
compareAsDir CompareDirection
rid CompareAs
t Term
n Term
m) m ()
r1
m () -> m () -> m ()
forall (m :: * -> *) a. MonadConstraint m => m a -> m a -> m a
catchPatternErr m ()
solve2 m ()
solve1
(NotBlocked NotBlocked
_ (MetaV MetaId
x Elims
es), Blocked Term
_) -> CompareDirection -> MetaId -> Elims -> Term -> m ()
forall (m :: * -> *).
(MonadMetaSolver m, MonadWarning m, MonadStatistics m,
MonadFresh ProblemId m, MonadFresh Int m) =>
CompareDirection -> MetaId -> Elims -> Term -> m ()
assign CompareDirection
dir MetaId
x Elims
es Term
n
(Blocked Term
_, NotBlocked NotBlocked
_ (MetaV MetaId
x Elims
es)) -> CompareDirection -> MetaId -> Elims -> Term -> m ()
forall (m :: * -> *).
(MonadMetaSolver m, MonadWarning m, MonadStatistics m,
MonadFresh ProblemId m, MonadFresh Int m) =>
CompareDirection -> MetaId -> Elims -> Term -> m ()
assign CompareDirection
rid MetaId
x Elims
es Term
m
(Blocked{}, Blocked{}) -> m ()
checkDefinitionalEquality
(Blocked{}, Blocked Term
_) -> CompareDirection -> CompareAs -> Term -> Term -> m ()
forall (m :: * -> *).
MonadConversion m =>
CompareDirection -> CompareAs -> Term -> Term -> m ()
useInjectivity (Comparison -> CompareDirection
fromCmp Comparison
cmp) CompareAs
t Term
m Term
n
(Blocked Term
_, Blocked{}) -> CompareDirection -> CompareAs -> Term -> Term -> m ()
forall (m :: * -> *).
MonadConversion m =>
CompareDirection -> CompareAs -> Term -> Term -> m ()
useInjectivity (CompareDirection -> CompareDirection
flipCmp (CompareDirection -> CompareDirection)
-> CompareDirection -> CompareDirection
forall a b. (a -> b) -> a -> b
$ Comparison -> CompareDirection
fromCmp Comparison
cmp) CompareAs
t Term
n Term
m
(Blocked Term, Blocked Term)
_ -> CompareAs -> (Blocked CompareAs -> m ()) -> m ()
postponeIfBlockedAs CompareAs
t ((Blocked CompareAs -> m ()) -> m ())
-> (Blocked CompareAs -> m ()) -> m ()
forall a b. (a -> b) -> a -> b
$ \Blocked CompareAs
bt -> do
case (Term
m, Term
n) of
(Pi{}, Pi{}) -> Term -> Term -> m ()
forall (m :: * -> *).
(MonadMetaSolver m, MonadWarning m, MonadStatistics m,
MonadFresh ProblemId m, MonadFresh Int m) =>
Term -> Term -> m ()
equalFun Term
m Term
n
(Sort Sort
s1, Sort Sort
s2) ->
m Bool -> m () -> m () -> m ()
forall (m :: * -> *) a. Monad m => m Bool -> m a -> m a -> m a
ifM (PragmaOptions -> Bool
optCumulativity (PragmaOptions -> Bool) -> m PragmaOptions -> m Bool
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> m PragmaOptions
forall (m :: * -> *). HasOptions m => m PragmaOptions
pragmaOptions)
(Comparison -> Sort -> Sort -> m ()
forall (m :: * -> *).
MonadConversion m =>
Comparison -> Sort -> Sort -> m ()
compareSort Comparison
cmp Sort
s1 Sort
s2)
(Sort -> Sort -> m ()
forall (m :: * -> *). MonadConversion m => Sort -> Sort -> m ()
equalSort Sort
s1 Sort
s2)
(Lit Literal
l1, Lit Literal
l2) | Literal
l1 Literal -> Literal -> Bool
forall a. Eq a => a -> a -> Bool
== Literal
l2 -> () -> m ()
forall (m :: * -> *) a. Monad m => a -> m a
return ()
(Var Int
i Elims
es, Var Int
i' Elims
es') | Int
i Int -> Int -> Bool
forall a. Eq a => a -> a -> Bool
== Int
i' -> do
Type
a <- Int -> m Type
forall (m :: * -> *).
(Applicative m, MonadFail m, MonadTCEnv m) =>
Int -> m Type
typeOfBV Int
i
[Polarity] -> [IsForced] -> Type -> Term -> Elims -> Elims -> m ()
forall (m :: * -> *).
MonadConversion m =>
[Polarity] -> [IsForced] -> Type -> Term -> Elims -> Elims -> m ()
compareElims [] [] Type
a (Int -> Term
var Int
i) Elims
es Elims
es'
(Def QName
f Elims
es, Def QName
f' Elims
es') -> do
m Bool -> m () -> m ()
forall (m :: * -> *). Monad m => m Bool -> m () -> m ()
unlessM (QName -> QName -> m Bool
forall (m :: * -> *). MonadConversion m => QName -> QName -> m Bool
bothAbsurd QName
f QName
f') (m () -> m ()) -> m () -> m ()
forall a b. (a -> b) -> a -> b
$ do
if QName
f QName -> QName -> Bool
forall a. Eq a => a -> a -> Bool
/= QName
f' then Comparison
-> CompareAs
-> Term
-> Term
-> QName
-> Elims
-> QName
-> Elims
-> m ()
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
Bool -> m () -> m ()
forall (f :: * -> *). Applicative f => Bool -> f () -> f ()
unless (Elims -> Bool
forall (t :: * -> *) a. Foldable t => t a -> Bool
null Elims
es Bool -> Bool -> Bool
&& Elims -> Bool
forall (t :: * -> *) a. Foldable t => t a -> Bool
null Elims
es') (m () -> m ()) -> m () -> m ()
forall a b. (a -> b) -> a -> b
$ do
m Bool -> m () -> m ()
forall (m :: * -> *). Monad m => m Bool -> m () -> m ()
unlessM (MonadConversion m => QName -> Elims -> Elims -> m Bool
QName -> Elims -> Elims -> m Bool
compareEtaPrims QName
f Elims
es Elims
es') (m () -> m ()) -> m () -> m ()
forall a b. (a -> b) -> a -> b
$ do
Type
a <- QName -> Elims -> Elims -> m Type
forall (m :: * -> *).
MonadConversion m =>
QName -> Elims -> Elims -> m Type
computeElimHeadType QName
f Elims
es Elims
es'
[Polarity]
pol <- Comparison -> QName -> m [Polarity]
forall (m :: * -> *).
HasConstInfo m =>
Comparison -> QName -> m [Polarity]
getPolarity' Comparison
cmp QName
f
[Polarity] -> [IsForced] -> Type -> Term -> Elims -> Elims -> m ()
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'
(Con ConHead
x ConInfo
ci Elims
xArgs, Con ConHead
y ConInfo
_ Elims
yArgs)
| ConHead
x ConHead -> ConHead -> Bool
forall a. Eq a => a -> a -> Bool
== ConHead
y -> do
Type
a' <- case CompareAs
t of
AsTermsOf Type
a -> ConHead -> Type -> m Type
forall (m :: * -> *).
(MonadReduce m, HasBuiltins m, HasConstInfo m,
MonadError TCErr m) =>
ConHead -> Type -> m Type
conType ConHead
x Type
a
CompareAs
AsSizes -> m Type
forall a. HasCallStack => a
__IMPOSSIBLE__
CompareAs
AsTypes -> m Type
forall a. HasCallStack => a
__IMPOSSIBLE__
[IsForced]
forcedArgs <- QName -> m [IsForced]
forall (m :: * -> *). HasConstInfo m => QName -> m [IsForced]
getForcedArgs (QName -> m [IsForced]) -> QName -> m [IsForced]
forall a b. (a -> b) -> a -> b
$ ConHead -> QName
conName ConHead
x
[Polarity] -> [IsForced] -> Type -> Term -> Elims -> Elims -> m ()
forall (m :: * -> *).
MonadConversion m =>
[Polarity] -> [IsForced] -> Type -> Term -> Elims -> Elims -> m ()
compareElims (Polarity -> [Polarity]
forall a. a -> [a]
repeat (Polarity -> [Polarity]) -> Polarity -> [Polarity]
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)
_ -> TypeError -> m ()
forall (m :: * -> *) a.
(MonadTCEnv m, ReadTCState m, MonadError TCErr m) =>
TypeError -> m a
typeError (TypeError -> m ()) -> TypeError -> m ()
forall a b. (a -> b) -> a -> b
$ Comparison -> Term -> Term -> CompareAs -> TypeError
UnequalTerms Comparison
cmp Term
m Term
n (CompareAs -> TypeError) -> CompareAs -> TypeError
forall a b. (a -> b) -> a -> b
$ Blocked CompareAs -> CompareAs
forall t. Blocked t -> t
ignoreBlocking Blocked CompareAs
bt
where
compareEtaPrims :: MonadConversion m => QName -> Elims -> Elims -> m Bool
compareEtaPrims :: QName -> Elims -> Elims -> m Bool
compareEtaPrims QName
q Elims
es Elims
es' = do
Maybe QName
munglue <- VerboseKey -> m (Maybe QName)
forall (m :: * -> *).
HasBuiltins m =>
VerboseKey -> m (Maybe QName)
getPrimitiveName' VerboseKey
builtin_unglue
Maybe QName
munglueU <- VerboseKey -> m (Maybe QName)
forall (m :: * -> *).
HasBuiltins m =>
VerboseKey -> m (Maybe QName)
getPrimitiveName' VerboseKey
builtin_unglueU
Maybe QName
msubout <- VerboseKey -> m (Maybe QName)
forall (m :: * -> *).
HasBuiltins m =>
VerboseKey -> m (Maybe QName)
getPrimitiveName' VerboseKey
builtinSubOut
case () of
()
_ | QName -> Maybe QName
forall a. a -> Maybe a
Just QName
q Maybe QName -> Maybe QName -> Bool
forall a. Eq a => a -> a -> Bool
== Maybe QName
munglue -> QName -> Elims -> Elims -> m Bool
forall (m :: * -> *).
(MonadMetaSolver m, MonadWarning m, MonadStatistics m,
MonadFresh ProblemId m, MonadFresh Int m) =>
QName -> Elims -> Elims -> m Bool
compareUnglueApp QName
q Elims
es Elims
es'
()
_ | QName -> Maybe QName
forall a. a -> Maybe a
Just QName
q Maybe QName -> Maybe QName -> Bool
forall a. Eq a => a -> a -> Bool
== Maybe QName
munglueU -> MonadConversion m => QName -> Elims -> Elims -> m Bool
QName -> Elims -> Elims -> m Bool
compareUnglueUApp QName
q Elims
es Elims
es'
()
_ | QName -> Maybe QName
forall a. a -> Maybe a
Just QName
q Maybe QName -> Maybe QName -> Bool
forall a. Eq a => a -> a -> Bool
== Maybe QName
msubout -> QName -> Elims -> Elims -> m Bool
forall (m :: * -> *).
(MonadMetaSolver m, MonadWarning m, MonadStatistics m,
MonadFresh ProblemId m, MonadFresh Int m) =>
QName -> Elims -> Elims -> m Bool
compareSubApp QName
q Elims
es Elims
es'
()
_ -> Bool -> m Bool
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) = Int -> Elims -> (Elims, Elims)
forall a. Int -> [a] -> ([a], [a])
splitAt Int
5 Elims
es; (Elims
as',Elims
bs') = Int -> Elims -> (Elims, Elims)
forall a. Int -> [a] -> ([a], [a])
splitAt Int
5 Elims
es'
case (Elims -> Maybe [Arg Term]
forall a. [Elim' a] -> Maybe [Arg a]
allApplyElims Elims
as, Elims -> Maybe [Arg Term]
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 <- m Term
forall (m :: * -> *).
(HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) =>
m Term
primSub
Type -> Type -> m ()
forall (m :: * -> *). MonadConversion m => Type -> Type -> m ()
equalType (Sort -> Term -> Type
forall t a. Sort' t -> a -> Type'' t a
El Sort
forall t. Sort' t
Inf (Term -> Type) -> Term -> Type
forall a b. (a -> b) -> a -> b
$ Term -> [Arg Term] -> Term
forall t. Apply t => t -> [Arg Term] -> t
apply Term
tSub ([Arg Term] -> Term) -> [Arg Term] -> Term
forall a b. (a -> b) -> a -> b
$ [Arg Term
a] [Arg Term] -> [Arg Term] -> [Arg Term]
forall a. [a] -> [a] -> [a]
++ (Arg Term -> Arg Term) -> [Arg Term] -> [Arg Term]
forall a b. (a -> b) -> [a] -> [b]
map (Hiding -> Arg Term -> Arg Term
forall a. LensHiding a => Hiding -> a -> a
setHiding Hiding
NotHidden) [Arg Term
bA,Arg Term
phi,Arg Term
u])
(Sort -> Term -> Type
forall t a. Sort' t -> a -> Type'' t a
El Sort
forall t. Sort' t
Inf (Term -> Type) -> Term -> Type
forall a b. (a -> b) -> a -> b
$ Term -> [Arg Term] -> Term
forall t. Apply t => t -> [Arg Term] -> t
apply Term
tSub ([Arg Term] -> Term) -> [Arg Term] -> Term
forall a b. (a -> b) -> a -> b
$ [Arg Term
a] [Arg Term] -> [Arg Term] -> [Arg Term]
forall a. [a] -> [a] -> [a]
++ (Arg Term -> Arg Term) -> [Arg Term] -> [Arg Term]
forall a b. (a -> b) -> [a] -> [b]
map (Hiding -> Arg Term -> Arg Term
forall a. LensHiding a => Hiding -> a -> a
setHiding Hiding
NotHidden) [Arg Term
bA',Arg Term
phi',Arg Term
u'])
Comparison -> CompareAs -> Term -> Term -> m ()
forall (m :: * -> *).
MonadConversion m =>
Comparison -> CompareAs -> Term -> Term -> m ()
compareAtom Comparison
cmp (Type -> CompareAs
AsTermsOf (Type -> CompareAs) -> Type -> CompareAs
forall a b. (a -> b) -> a -> b
$ Sort -> Term -> Type
forall t a. Sort' t -> a -> Type'' t a
El Sort
forall t. Sort' t
Inf (Term -> Type) -> Term -> Type
forall a b. (a -> b) -> a -> b
$ Term -> [Arg Term] -> Term
forall t. Apply t => t -> [Arg Term] -> t
apply Term
tSub ([Arg Term] -> Term) -> [Arg Term] -> Term
forall a b. (a -> b) -> a -> b
$ [Arg Term
a] [Arg Term] -> [Arg Term] -> [Arg Term]
forall a. [a] -> [a] -> [a]
++ (Arg Term -> Arg Term) -> [Arg Term] -> [Arg Term]
forall a b. (a -> b) -> [a] -> [b]
map (Hiding -> Arg Term -> Arg Term
forall a. LensHiding a => Hiding -> a -> a
setHiding Hiding
NotHidden) [Arg Term
bA,Arg Term
phi,Arg Term
u])
(Arg Term -> Term
forall e. Arg e -> e
unArg Arg Term
x) (Arg Term -> Term
forall e. Arg e -> e
unArg Arg Term
x')
[Polarity] -> [IsForced] -> Type -> Term -> Elims -> Elims -> m ()
forall (m :: * -> *).
MonadConversion m =>
[Polarity] -> [IsForced] -> Type -> Term -> Elims -> Elims -> m ()
compareElims [] [] (Sort -> Term -> Type
forall t a. Sort' t -> a -> Type'' t a
El (Term -> Sort
tmSort (Arg Term -> Term
forall e. Arg e -> e
unArg Arg Term
a)) (Arg Term -> Term
forall e. Arg e -> e
unArg Arg Term
bA)) (QName -> Elims -> Term
Def QName
q Elims
as) Elims
bs Elims
bs'
Bool -> m Bool
forall (m :: * -> *) a. Monad m => a -> m a
return Bool
True
(Maybe [Arg Term], Maybe [Arg Term])
_ -> Bool -> m Bool
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) = Int -> Elims -> (Elims, Elims)
forall a. Int -> [a] -> ([a], [a])
splitAt Int
7 Elims
es; (Elims
as',Elims
bs') = Int -> Elims -> (Elims, Elims)
forall a. Int -> [a] -> ([a], [a])
splitAt Int
7 Elims
es'
case (Elims -> Maybe [Arg Term]
forall a. [Elim' a] -> Maybe [Arg a]
allApplyElims Elims
as, Elims -> Maybe [Arg Term]
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 <- VerboseKey -> m Term
forall (m :: * -> *).
(HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) =>
VerboseKey -> m Term
getPrimitiveTerm VerboseKey
builtinGlue
Comparison -> CompareAs -> Term -> Term -> m ()
forall (m :: * -> *).
MonadConversion m =>
Comparison -> CompareAs -> Term -> Term -> m ()
compareAtom Comparison
cmp (Type -> CompareAs
AsTermsOf (Type -> CompareAs) -> Type -> CompareAs
forall a b. (a -> b) -> a -> b
$ Sort -> Term -> Type
forall t a. Sort' t -> a -> Type'' t a
El (Term -> Sort
tmSort (Arg Term -> Term
forall e. Arg e -> e
unArg Arg Term
lb)) (Term -> Type) -> Term -> Type
forall a b. (a -> b) -> a -> b
$ Term -> [Arg Term] -> Term
forall t. Apply t => t -> [Arg Term] -> t
apply Term
tGlue ([Arg Term] -> Term) -> [Arg Term] -> Term
forall a b. (a -> b) -> a -> b
$ [Arg Term
la,Arg Term
lb] [Arg Term] -> [Arg Term] -> [Arg Term]
forall a. [a] -> [a] -> [a]
++ (Arg Term -> Arg Term) -> [Arg Term] -> [Arg Term]
forall a b. (a -> b) -> [a] -> [b]
map (Hiding -> Arg Term -> Arg Term
forall a. LensHiding a => Hiding -> a -> a
setHiding Hiding
NotHidden) [Arg Term
bA,Arg Term
phi,Arg Term
bT,Arg Term
e])
(Arg Term -> Term
forall e. Arg e -> e
unArg Arg Term
b) (Arg Term -> Term
forall e. Arg e -> e
unArg Arg Term
b')
[Polarity] -> [IsForced] -> Type -> Term -> Elims -> Elims -> m ()
forall (m :: * -> *).
MonadConversion m =>
[Polarity] -> [IsForced] -> Type -> Term -> Elims -> Elims -> m ()
compareElims [] [] (Sort -> Term -> Type
forall t a. Sort' t -> a -> Type'' t a
El (Term -> Sort
tmSort (Arg Term -> Term
forall e. Arg e -> e
unArg Arg Term
la)) (Arg Term -> Term
forall e. Arg e -> e
unArg Arg Term
bA)) (QName -> Elims -> Term
Def QName
q Elims
as) Elims
bs Elims
bs'
Bool -> m Bool
forall (m :: * -> *) a. Monad m => a -> m a
return Bool
True
(Maybe [Arg Term], Maybe [Arg Term])
_ -> Bool -> m Bool
forall (m :: * -> *) a. Monad m => a -> m a
return Bool
False
compareUnglueUApp :: MonadConversion m => QName -> Elims -> Elims -> m Bool
compareUnglueUApp :: QName -> Elims -> Elims -> m Bool
compareUnglueUApp QName
q Elims
es Elims
es' = do
let (Elims
as,Elims
bs) = Int -> Elims -> (Elims, Elims)
forall a. Int -> [a] -> ([a], [a])
splitAt Int
5 Elims
es; (Elims
as',Elims
bs') = Int -> Elims -> (Elims, Elims)
forall a. Int -> [a] -> ([a], [a])
splitAt Int
5 Elims
es'
case (Elims -> Maybe [Arg Term]
forall a. [Elim' a] -> Maybe [Arg a]
allApplyElims Elims
as, Elims -> Maybe [Arg Term]
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 <- m Term
forall (m :: * -> *).
(HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) =>
m Term
primHComp
Term
tLSuc <- m Term
forall (m :: * -> *).
(HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) =>
m Term
primLevelSuc
Term
tSubOut <- m Term
forall (m :: * -> *).
(HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) =>
m Term
primSubOut
Term
iz <- m Term
forall (m :: * -> *).
(HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) =>
m Term
primIZero
let lsuc :: Term -> Term
lsuc Term
t = Term
tLSuc Term -> [Arg Term] -> Term
forall t. Apply t => t -> [Arg Term] -> t
`apply` [Term -> Arg Term
forall e. e -> Arg e
argN Term
t]
s :: Sort
s = Term -> Sort
tmSort (Term -> Sort) -> Term -> Sort
forall a b. (a -> b) -> a -> b
$ Arg Term -> Term
forall e. Arg e -> e
unArg Arg Term
la
sucla :: Arg Term
sucla = Term -> Term
lsuc (Term -> Term) -> Arg Term -> Arg Term
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> Arg Term
la
Term
bA <- Names -> NamesT m Term -> m Term
forall (m :: * -> *) a. Names -> NamesT m a -> m a
runNamesT [] (NamesT m Term -> m Term) -> NamesT m Term -> m Term
forall a b. (a -> b) -> a -> b
$ do
[NamesT m Term
la,NamesT m Term
phi,NamesT m Term
bT,NamesT m Term
bAS] <- (Arg Term -> NamesT m (NamesT m Term))
-> [Arg Term] -> NamesT m [NamesT m Term]
forall (t :: * -> *) (m :: * -> *) a b.
(Traversable t, Monad m) =>
(a -> m b) -> t a -> m (t b)
mapM (Term -> NamesT m (NamesT m Term)
forall (m :: * -> *) t a.
(MonadFail m, Subst t a) =>
a -> NamesT m (NamesT m a)
open (Term -> NamesT m (NamesT m Term))
-> (Arg Term -> Term) -> Arg Term -> NamesT m (NamesT m Term)
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Arg Term -> Term
forall e. Arg e -> e
unArg) [Arg Term
la,Arg Term
phi,Arg Term
bT,Arg Term
bAS]
(Term -> NamesT m Term
forall (f :: * -> *) a. Applicative f => a -> f a
pure Term
tSubOut NamesT m Term -> NamesT m Term -> NamesT m Term
forall (tcm :: * -> *).
Monad tcm =>
tcm Term -> tcm Term -> tcm Term
<#> (Term -> NamesT m Term
forall (f :: * -> *) a. Applicative f => a -> f a
pure Term
tLSuc NamesT m Term -> NamesT m Term -> NamesT m Term
forall (tcm :: * -> *).
Monad tcm =>
tcm Term -> tcm Term -> tcm Term
<@> NamesT m Term
la) NamesT m Term -> NamesT m Term -> NamesT m Term
forall (tcm :: * -> *).
Monad tcm =>
tcm Term -> tcm Term -> tcm Term
<#> (Sort -> Term
Sort (Sort -> Term) -> (Term -> Sort) -> Term -> Term
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Term -> Sort
tmSort (Term -> Term) -> NamesT m Term -> NamesT m Term
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> NamesT m Term
la) NamesT m Term -> NamesT m Term -> NamesT m Term
forall (tcm :: * -> *).
Monad tcm =>
tcm Term -> tcm Term -> tcm Term
<#> NamesT m Term
phi NamesT m Term -> NamesT m Term -> NamesT m Term
forall (tcm :: * -> *).
Monad tcm =>
tcm Term -> tcm Term -> tcm Term
<#> (NamesT m Term
bT NamesT m Term -> NamesT m Term -> NamesT m Term
forall (tcm :: * -> *).
Monad tcm =>
tcm Term -> tcm Term -> tcm Term
<@> NamesT m Term
forall (m :: * -> *).
(HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) =>
m Term
primIZero) NamesT m Term -> NamesT m Term -> NamesT m Term
forall (tcm :: * -> *).
Monad tcm =>
tcm Term -> tcm Term -> tcm Term
<@> NamesT m Term
bAS)
Comparison -> CompareAs -> Term -> Term -> m ()
forall (m :: * -> *).
MonadConversion m =>
Comparison -> CompareAs -> Term -> Term -> m ()
compareAtom Comparison
cmp (Type -> CompareAs
AsTermsOf (Type -> CompareAs) -> Type -> CompareAs
forall a b. (a -> b) -> a -> b
$ Sort -> Term -> Type
forall t a. Sort' t -> a -> Type'' t a
El (Term -> Sort
tmSort (Term -> Sort) -> (Arg Term -> Term) -> Arg Term -> Sort
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Arg Term -> Term
forall e. Arg e -> e
unArg (Arg Term -> Sort) -> Arg Term -> Sort
forall a b. (a -> b) -> a -> b
$ Arg Term
sucla) (Term -> Type) -> Term -> Type
forall a b. (a -> b) -> a -> b
$ Term -> [Arg Term] -> Term
forall t. Apply t => t -> [Arg Term] -> t
apply Term
tHComp ([Arg Term] -> Term) -> [Arg Term] -> Term
forall a b. (a -> b) -> a -> b
$ [Arg Term
sucla, Term -> Arg Term
forall e. e -> Arg e
argH (Sort -> Term
Sort Sort
s), Arg Term
phi] [Arg Term] -> [Arg Term] -> [Arg Term]
forall a. [a] -> [a] -> [a]
++ [Term -> Arg Term
forall e. e -> Arg e
argH (Arg Term -> Term
forall e. Arg e -> e
unArg Arg Term
bT), Term -> Arg Term
forall e. e -> Arg e
argH Term
bA])
(Arg Term -> Term
forall e. Arg e -> e
unArg Arg Term
b) (Arg Term -> Term
forall e. Arg e -> e
unArg Arg Term
b')
[Polarity] -> [IsForced] -> Type -> Term -> Elims -> Elims -> m ()
forall (m :: * -> *).
MonadConversion m =>
[Polarity] -> [IsForced] -> Type -> Term -> Elims -> Elims -> m ()
compareElims [] [] (Sort -> Term -> Type
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'
Bool -> m Bool
forall (m :: * -> *) a. Monad m => a -> m a
return Bool
True
(Maybe [Arg Term], Maybe [Arg Term])
_ -> Bool -> m Bool
forall (m :: * -> *) a. Monad m => a -> m a
return Bool
False
conType :: ConHead -> Type -> m Type
conType ConHead
c Type
t = Type
-> (MetaId -> Type -> m Type)
-> (NotBlocked -> Type -> m Type)
-> m Type
forall t (m :: * -> *) a.
(Reduce t, IsMeta t, MonadReduce m, HasBuiltins m) =>
t -> (MetaId -> t -> m a) -> (NotBlocked -> t -> m a) -> m a
ifBlocked Type
t (\ MetaId
_ Type
_ -> m Type
forall (m :: * -> *) a. MonadError TCErr m => m a
patternViolation) ((NotBlocked -> Type -> m Type) -> m Type)
-> (NotBlocked -> Type -> m Type) -> m Type
forall a b. (a -> b) -> a -> b
$ \ NotBlocked
_ Type
t -> do
let impossible :: m b
impossible = do
VerboseKey -> Int -> TCM Doc -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCM Doc -> m ()
reportSDoc VerboseKey
"impossible" Int
10 (TCM Doc -> m ()) -> TCM Doc -> m ()
forall a b. (a -> b) -> a -> b
$
TCM Doc
"expected data/record type, found " TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Type -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Type
t
VerboseKey -> Int -> TCM Doc -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCM Doc -> m ()
reportSDoc VerboseKey
"impossible" Int
70 (TCM Doc -> m ()) -> TCM Doc -> m ()
forall a b. (a -> b) -> a -> b
$ Int -> TCM Doc -> TCM Doc
forall (m :: * -> *). Functor m => Int -> m Doc -> m Doc
nest Int
2 (TCM Doc -> TCM Doc) -> TCM Doc -> TCM Doc
forall a b. (a -> b) -> a -> b
$ TCM Doc
"raw =" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Type -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty Type
t
m b
forall (m :: * -> *) a. MonadError TCErr m => m a
patternViolation
m Type
-> (((QName, Type, [Arg Term]), Type) -> m Type)
-> Maybe ((QName, Type, [Arg Term]), Type)
-> m Type
forall b a. b -> (a -> b) -> Maybe a -> b
maybe m Type
forall b. m b
impossible (Type -> m Type
forall (m :: * -> *) a. Monad m => a -> m a
return (Type -> m Type)
-> (((QName, Type, [Arg Term]), Type) -> Type)
-> ((QName, Type, [Arg Term]), Type)
-> m Type
forall b c a. (b -> c) -> (a -> b) -> a -> c
. ((QName, Type, [Arg Term]), Type) -> Type
forall a b. (a, b) -> b
snd) (Maybe ((QName, Type, [Arg Term]), Type) -> m Type)
-> m (Maybe ((QName, Type, [Arg Term]), Type)) -> m Type
forall (m :: * -> *) a b. Monad m => (a -> m b) -> m a -> m b
=<< ConHead -> Type -> m (Maybe ((QName, Type, [Arg Term]), Type))
forall (m :: * -> *).
(HasConstInfo m, MonadReduce m, MonadDebug 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
VerboseKey -> Int -> VerboseKey -> m () -> m ()
forall (m :: * -> *) a.
MonadDebug m =>
VerboseKey -> Int -> VerboseKey -> m a -> m a
verboseBracket VerboseKey
"tc.conv.fun" Int
15 VerboseKey
"compare function types" (m () -> m ()) -> m () -> m ()
forall a b. (a -> b) -> a -> b
$ do
VerboseKey -> Int -> TCM Doc -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCM Doc -> m ()
reportSDoc VerboseKey
"tc.conv.fun" Int
20 (TCM Doc -> m ()) -> TCM Doc -> m ()
forall a b. (a -> b) -> a -> b
$ Int -> TCM Doc -> TCM Doc
forall (m :: * -> *). Functor m => Int -> m Doc -> m Doc
nest Int
2 (TCM Doc -> TCM Doc) -> TCM Doc -> TCM Doc
forall a b. (a -> b) -> a -> b
$ [TCM Doc] -> TCM Doc
forall (m :: * -> *). Monad m => [m Doc] -> m Doc
vcat
[ TCM Doc
"t1 =" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Term -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Term
t1
, TCM Doc
"t2 =" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Term -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Term
t2
]
Comparison
-> Dom Type
-> Dom Type
-> Abs Type
-> Abs Type
-> m ()
-> m ()
-> m ()
-> m ()
-> m ()
-> m ()
forall (m :: * -> *) c b.
(MonadConversion m, Free c) =>
Comparison
-> Dom Type
-> Dom Type
-> Abs b
-> Abs c
-> m ()
-> m ()
-> m ()
-> m ()
-> m ()
-> m ()
compareDom Comparison
cmp Dom Type
dom2 Dom Type
dom1 Abs Type
b1 Abs Type
b2 m ()
forall a. m a
errH m ()
forall a. m a
errR m ()
forall a. m a
errQ m ()
forall a. m a
errC (m () -> m ()) -> m () -> m ()
forall a b. (a -> b) -> a -> b
$
Comparison -> Type -> Type -> m ()
forall (m :: * -> *).
MonadConversion m =>
Comparison -> Type -> Type -> m ()
compareType Comparison
cmp (Abs Type -> Type
forall t a. Subst t a => Abs a -> a
absBody Abs Type
b1) (Abs Type -> Type
forall t a. Subst t a => Abs a -> a
absBody Abs Type
b2)
where
errH :: m a
errH = TypeError -> m a
forall (m :: * -> *) a.
(MonadTCEnv m, ReadTCState m, MonadError TCErr m) =>
TypeError -> m a
typeError (TypeError -> m a) -> TypeError -> m a
forall a b. (a -> b) -> a -> b
$ Term -> Term -> TypeError
UnequalHiding Term
t1 Term
t2
errR :: m a
errR = TypeError -> m a
forall (m :: * -> *) a.
(MonadTCEnv m, ReadTCState m, MonadError TCErr m) =>
TypeError -> m a
typeError (TypeError -> m a) -> TypeError -> m a
forall a b. (a -> b) -> a -> b
$ Comparison -> Term -> Term -> TypeError
UnequalRelevance Comparison
cmp Term
t1 Term
t2
errQ :: m a
errQ = TypeError -> m a
forall (m :: * -> *) a.
(MonadTCEnv m, ReadTCState m, MonadError TCErr m) =>
TypeError -> m a
typeError (TypeError -> m a) -> TypeError -> m a
forall a b. (a -> b) -> a -> b
$ Comparison -> Term -> Term -> TypeError
UnequalQuantity Comparison
cmp Term
t1 Term
t2
errC :: m a
errC = TypeError -> m a
forall (m :: * -> *) a.
(MonadTCEnv m, ReadTCState m, MonadError TCErr m) =>
TypeError -> m a
typeError (TypeError -> m a) -> TypeError -> m a
forall a b. (a -> b) -> a -> b
$ Comparison -> Term -> Term -> TypeError
UnequalCohesion Comparison
cmp Term
t1 Term
t2
(Term, Term)
_ -> m ()
forall a. HasCallStack => a
__IMPOSSIBLE__
compareDom :: (MonadConversion m , Free c)
=> Comparison
-> Dom Type
-> Dom Type
-> Abs b
-> Abs c
-> m ()
-> m ()
-> m ()
-> m ()
-> m ()
-> m ()
compareDom :: Comparison
-> Dom Type
-> Dom Type
-> Abs b
-> Abs c
-> 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 ()
cont = do
Bool
hasSubtyping <- WithDefault 'False -> Bool
forall (b :: Bool). KnownBool b => WithDefault b -> Bool
collapseDefault (WithDefault 'False -> Bool)
-> (PragmaOptions -> WithDefault 'False) -> PragmaOptions -> Bool
forall b c a. (b -> c) -> (a -> b) -> a -> c
. PragmaOptions -> WithDefault 'False
optSubtyping (PragmaOptions -> Bool) -> m PragmaOptions -> m Bool
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> m PragmaOptions
forall (m :: * -> *). HasOptions m => m PragmaOptions
pragmaOptions
let cmp :: Comparison
cmp = if Bool
hasSubtyping then Comparison
cmp0 else Comparison
CmpEq
if | Bool -> Bool
not (Bool -> Bool) -> Bool -> Bool
forall a b. (a -> b) -> a -> b
$ Dom Type -> Dom Type -> Bool
forall a b. (LensHiding a, LensHiding b) => a -> b -> Bool
sameHiding Dom Type
dom1 Dom Type
dom2 -> m ()
errH
| Bool -> Bool
not (Bool -> Bool) -> Bool -> Bool
forall a b. (a -> b) -> a -> b
$ Comparison -> Relevance -> Relevance -> Bool
compareRelevance Comparison
cmp (Dom Type -> Relevance
forall a. LensRelevance a => a -> Relevance
getRelevance Dom Type
dom1) (Dom Type -> Relevance
forall a. LensRelevance a => a -> Relevance
getRelevance Dom Type
dom2) -> m ()
errR
| Bool -> Bool
not (Bool -> Bool) -> Bool -> Bool
forall a b. (a -> b) -> a -> b
$ Comparison -> Quantity -> Quantity -> Bool
compareQuantity Comparison
cmp (Dom Type -> Quantity
forall a. LensQuantity a => a -> Quantity
getQuantity Dom Type
dom1) (Dom Type -> Quantity
forall a. LensQuantity a => a -> Quantity
getQuantity Dom Type
dom2) -> m ()
errQ
| Bool -> Bool
not (Bool -> Bool) -> Bool -> Bool
forall a b. (a -> b) -> a -> b
$ Comparison -> Cohesion -> Cohesion -> Bool
compareCohesion Comparison
cmp (Dom Type -> Cohesion
forall a. LensCohesion a => a -> Cohesion
getCohesion Dom Type
dom1) (Dom Type -> Cohesion
forall a. LensCohesion a => a -> Cohesion
getCohesion Dom Type
dom2) -> m ()
errC
| Bool
otherwise -> do
let r :: Relevance
r = Relevance -> Relevance -> Relevance
forall a. Ord a => a -> a -> a
max (Dom Type -> Relevance
forall a. LensRelevance a => a -> Relevance
getRelevance Dom Type
dom1) (Dom Type -> Relevance
forall a. LensRelevance a => a -> Relevance
getRelevance Dom Type
dom2)
dependent :: Bool
dependent = (Relevance
r Relevance -> Relevance -> Bool
forall a. Eq a => a -> a -> Bool
/= Relevance
Irrelevant) Bool -> Bool -> Bool
&& Abs c -> Bool
forall a. Free a => Abs a -> Bool
isBinderUsed Abs c
b2
ProblemId
pid <- m () -> m ProblemId
forall (m :: * -> *) a.
(MonadFresh ProblemId m, MonadConstraint m) =>
m a -> m ProblemId
newProblem_ (m () -> m ProblemId) -> m () -> m ProblemId
forall a b. (a -> b) -> a -> b
$ Comparison -> Type -> Type -> m ()
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}) (Type -> Dom Type) -> m Type -> m (Dom Type)
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> Type -> ProblemId -> m Type
forall (m :: * -> *).
(MonadMetaSolver m, MonadFresh Int m) =>
Type -> ProblemId -> m Type
blockTypeOnProblem Type
a1 ProblemId
pid
else Dom Type -> m (Dom Type)
forall (m :: * -> *) a. Monad m => a -> m a
return Dom Type
dom1
let name :: VerboseKey
name = [Suggestion] -> VerboseKey
suggests [ Abs b -> Suggestion
forall a. Suggest a => a -> Suggestion
Suggestion Abs b
b1 , Abs c -> Suggestion
forall a. Suggest a => a -> Suggestion
Suggestion Abs c
b2 ]
(VerboseKey, Dom Type) -> m () -> m ()
forall b (m :: * -> *) a.
(AddContext b, MonadAddContext m) =>
b -> m a -> m a
addContext (VerboseKey
name, Dom Type
dom) (m () -> m ()) -> m () -> m ()
forall a b. (a -> b) -> a -> b
$ m ()
cont
ProblemId -> m ()
forall (m :: * -> *). MonadConstraint m => ProblemId -> m ()
stealConstraints ProblemId
pid
compareRelevance :: Comparison -> Relevance -> Relevance -> Bool
compareRelevance :: Comparison -> Relevance -> Relevance -> Bool
compareRelevance Comparison
CmpEq = Relevance -> Relevance -> Bool
forall a. Eq a => a -> a -> Bool
(==)
compareRelevance Comparison
CmpLeq = Relevance -> Relevance -> Bool
forall a. Ord a => a -> a -> Bool
(<=)
compareQuantity :: Comparison -> Quantity -> Quantity -> Bool
compareQuantity :: Comparison -> Quantity -> Quantity -> Bool
compareQuantity Comparison
CmpEq = Quantity -> Quantity -> Bool
sameQuantity
compareQuantity Comparison
CmpLeq = Quantity -> Quantity -> Bool
moreQuantity
compareCohesion :: Comparison -> Cohesion -> Cohesion -> Bool
compareCohesion :: Comparison -> Cohesion -> Cohesion -> Bool
compareCohesion Comparison
CmpEq = Cohesion -> Cohesion -> Bool
sameCohesion
compareCohesion Comparison
CmpLeq = Cohesion -> Cohesion -> Bool
moreCohesion
antiUnify :: MonadConversion m => ProblemId -> Type -> Term -> Term -> m Term
antiUnify :: ProblemId -> Type -> Term -> Term -> m Term
antiUnify ProblemId
pid Type
a Term
u Term
v = do
((Term
u, Term
v), Bool
eq) <- Term -> Term -> m ((Term, Term), Bool)
forall a (m :: * -> *).
(Instantiate a, SynEq a, MonadReduce m) =>
a -> a -> m ((a, a), Bool)
SynEq.checkSyntacticEquality Term
u Term
v
if Bool
eq then Term -> m Term
forall (m :: * -> *) a. Monad m => a -> m a
return Term
u else do
(Term
u, Term
v) <- (Term, Term) -> m (Term, Term)
forall a (m :: * -> *). (Reduce a, MonadReduce m) => a -> m a
reduce (Term
u, Term
v)
VerboseKey -> Int -> TCM Doc -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCM Doc -> m ()
reportSDoc VerboseKey
"tc.conv.antiUnify" Int
30 (TCM Doc -> m ()) -> TCM Doc -> m ()
forall a b. (a -> b) -> a -> b
$ [TCM Doc] -> TCM Doc
forall (m :: * -> *). Monad m => [m Doc] -> m Doc
vcat
[ TCM Doc
"antiUnify"
, TCM Doc
"a =" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Type -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Type
a
, TCM Doc
"u =" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Term -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Term
u
, TCM Doc
"v =" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Term -> TCM 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 <- ProblemId -> Type -> Type -> m Type
forall (m :: * -> *).
MonadConversion m =>
ProblemId -> Type -> Type -> m Type
antiUnifyType ProblemId
pid (Dom Type -> Type
forall t e. Dom' t e -> e
unDom Dom Type
ua) (Dom Type -> Type
forall t e. Dom' t e -> e
unDom Dom Type
va)
let wa :: Dom Type
wa = Type
wa0 Type -> Dom Type -> Dom Type
forall (f :: * -> *) a b. Functor f => a -> f b -> f a
<$ Dom Type
ua
Type
wb <- Dom Type -> m Type -> m Type
forall b (m :: * -> *) a.
(AddContext b, MonadAddContext m) =>
b -> m a -> m a
addContext Dom Type
wa (m Type -> m Type) -> m Type -> m Type
forall a b. (a -> b) -> a -> b
$ ProblemId -> Type -> Type -> m Type
forall (m :: * -> *).
MonadConversion m =>
ProblemId -> Type -> Type -> m Type
antiUnifyType ProblemId
pid (Abs Type -> Type
forall t a. Subst t a => Abs a -> a
absBody Abs Type
ub) (Abs Type -> Type
forall t a. Subst t a => Abs a -> a
absBody Abs Type
vb)
Term -> m Term
forall (m :: * -> *) a. Monad m => a -> m a
return (Term -> m Term) -> Term -> m Term
forall a b. (a -> b) -> a -> b
$ Dom Type -> Abs Type -> Term
Pi Dom Type
wa (VerboseKey -> Type -> Abs Type
forall t a. (Subst t a, Free a) => VerboseKey -> a -> Abs a
mkAbs (Abs Type -> VerboseKey
forall a. Abs a -> VerboseKey
absName Abs Type
ub) Type
wb)
(Lam ArgInfo
i Abs Term
u, Lam ArgInfo
_ Abs Term
v) ->
Term -> m Term
forall a (m :: * -> *). (Reduce a, MonadReduce m) => a -> m a
reduce (Type -> Term
forall t a. Type'' t a -> a
unEl Type
a) m Term -> (Term -> m Term) -> m Term
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 (Abs Term -> Term) -> (Term -> Abs Term) -> Term -> Term
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (VerboseKey -> Term -> Abs Term
forall t a. (Subst t a, Free a) => VerboseKey -> a -> Abs a
mkAbs (Abs Term -> VerboseKey
forall a. Abs a -> VerboseKey
absName Abs Term
u)) (Term -> Term) -> m Term -> m Term
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> Dom Type -> m Term -> m Term
forall b (m :: * -> *) a.
(AddContext b, MonadAddContext m) =>
b -> m a -> m a
addContext Dom Type
a (ProblemId -> Type -> Term -> Term -> m Term
forall (m :: * -> *).
MonadConversion m =>
ProblemId -> Type -> Term -> Term -> m Term
antiUnify ProblemId
pid (Abs Type -> Type
forall t a. Subst t a => Abs a -> a
absBody Abs Type
b) (Abs Term -> Term
forall t a. Subst t a => Abs a -> a
absBody Abs Term
u) (Abs Term -> Term
forall t a. Subst t a => Abs a -> a
absBody Abs Term
v))
Term
_ -> m Term
fallback
(Var Int
i Elims
us, Var Int
j Elims
vs) | Int
i Int -> Int -> Bool
forall a. Eq a => a -> a -> Bool
== Int
j -> m Term -> m Term
maybeGiveUp (m Term -> m Term) -> m Term -> m Term
forall a b. (a -> b) -> a -> b
$ do
Type
a <- Int -> m Type
forall (m :: * -> *).
(Applicative m, MonadFail m, MonadTCEnv m) =>
Int -> m Type
typeOfBV Int
i
ProblemId -> Type -> Term -> Elims -> Elims -> m Term
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
(Con ConHead
x ConInfo
ci Elims
us, Con ConHead
y ConInfo
_ Elims
vs) | ConHead
x ConHead -> ConHead -> Bool
forall a. Eq a => a -> a -> Bool
== ConHead
y -> m Term -> m Term
maybeGiveUp (m Term -> m Term) -> m Term -> m Term
forall a b. (a -> b) -> a -> b
$ do
Type
a <- m Type
-> (((QName, Type, [Arg Term]), Type) -> m Type)
-> Maybe ((QName, Type, [Arg Term]), Type)
-> m Type
forall b a. b -> (a -> b) -> Maybe a -> b
maybe m Type
forall (m :: * -> *) a. MonadError TCErr m => m a
patternViolation (Type -> m Type
forall (m :: * -> *) a. Monad m => a -> m a
return (Type -> m Type)
-> (((QName, Type, [Arg Term]), Type) -> Type)
-> ((QName, Type, [Arg Term]), Type)
-> m Type
forall b c a. (b -> c) -> (a -> b) -> a -> c
. ((QName, Type, [Arg Term]), Type) -> Type
forall a b. (a, b) -> b
snd) (Maybe ((QName, Type, [Arg Term]), Type) -> m Type)
-> m (Maybe ((QName, Type, [Arg Term]), Type)) -> m Type
forall (m :: * -> *) a b. Monad m => (a -> m b) -> m a -> m b
=<< ConHead -> Type -> m (Maybe ((QName, Type, [Arg Term]), Type))
forall (m :: * -> *).
(MonadReduce m, MonadAddContext m, HasConstInfo m, MonadDebug m) =>
ConHead -> Type -> m (Maybe ((QName, Type, [Arg Term]), Type))
getConType ConHead
x Type
a
ProblemId -> Type -> Term -> Elims -> Elims -> m Term
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 Elims
us, Def QName
g Elims
vs) | QName
f QName -> QName -> Bool
forall a. Eq a => a -> a -> Bool
== QName
g, Elims -> Int
forall (t :: * -> *) a. Foldable t => t a -> Int
length Elims
us Int -> Int -> Bool
forall a. Eq a => a -> a -> Bool
== Elims -> Int
forall (t :: * -> *) a. Foldable t => t a -> Int
length Elims
vs -> m Term -> m Term
maybeGiveUp (m Term -> m Term) -> m Term -> m Term
forall a b. (a -> b) -> a -> b
$ do
Type
a <- QName -> Elims -> Elims -> m Type
forall (m :: * -> *).
MonadConversion m =>
QName -> Elims -> Elims -> m Type
computeElimHeadType QName
f Elims
us Elims
vs
ProblemId -> Type -> Term -> Elims -> Elims -> m Term
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 = m Term -> m Term -> m Term
forall (m :: * -> *) a. MonadConstraint m => m a -> m a -> m a
catchPatternErr m Term
fallback
fallback :: m Term
fallback = Type -> Term -> ProblemId -> m Term
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 :: ProblemId -> Dom Type -> Arg Term -> Arg Term -> m (Arg Term)
antiUnifyArgs ProblemId
pid Dom Type
dom Arg Term
u Arg Term
v
| Arg Term -> Modality
forall a. LensModality a => a -> Modality
getModality Arg Term
u Modality -> Modality -> Bool
forall a. Eq a => a -> a -> Bool
/= Arg Term -> Modality
forall a. LensModality a => a -> Modality
getModality Arg Term
v = m (Arg Term)
forall (m :: * -> *) a. MonadError TCErr m => m a
patternViolation
| Bool
otherwise = Arg Term -> m (Arg Term) -> m (Arg Term)
forall (tcm :: * -> *) m a.
(MonadTCEnv tcm, LensModality m) =>
m -> tcm a -> tcm a
applyModalityToContext Arg Term
u (m (Arg Term) -> m (Arg Term)) -> m (Arg Term) -> m (Arg Term)
forall a b. (a -> b) -> a -> b
$
m Bool -> m (Arg Term) -> m (Arg Term) -> m (Arg Term)
forall (m :: * -> *) a. Monad m => m Bool -> m a -> m a -> m a
ifM (Dom Type -> m Bool
forall a (m :: * -> *).
(LensRelevance a, LensSort a, PrettyTCM a, MonadReduce m,
MonadDebug m) =>
a -> m Bool
isIrrelevantOrPropM Dom Type
dom)
(Arg Term -> m (Arg Term)
forall (m :: * -> *) a. Monad m => a -> m a
return Arg Term
u)
((Term -> Arg Term -> Arg Term
forall (f :: * -> *) a b. Functor f => a -> f b -> f a
<$ Arg Term
u) (Term -> Arg Term) -> m Term -> m (Arg Term)
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> ProblemId -> Type -> Term -> Term -> m Term
forall (m :: * -> *).
MonadConversion m =>
ProblemId -> Type -> Term -> Term -> m Term
antiUnify ProblemId
pid (Dom Type -> Type
forall t e. Dom' t e -> e
unDom Dom Type
dom) (Arg Term -> Term
forall e. Arg e -> e
unArg Arg Term
u) (Arg Term -> Term
forall e. Arg e -> e
unArg Arg Term
v))
antiUnifyType :: MonadConversion m => ProblemId -> Type -> Type -> m Type
antiUnifyType :: ProblemId -> Type -> Type -> m Type
antiUnifyType ProblemId
pid (El Sort
s Term
a) (El Sort
_ Term
b) = m Type -> m Type
forall (m :: * -> *) a.
(MonadTCEnv m, HasOptions m, MonadDebug m) =>
m a -> m a
workOnTypes (m Type -> m Type) -> m Type -> m Type
forall a b. (a -> b) -> a -> b
$ Sort -> Term -> Type
forall t a. Sort' t -> a -> Type'' t a
El Sort
s (Term -> Type) -> m Term -> m Type
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> ProblemId -> Type -> Term -> Term -> m Term
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 :: ProblemId -> Type -> Term -> Elims -> Elims -> m Term
antiUnifyElims ProblemId
pid Type
a Term
self [] [] = Term -> m Term
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 QName -> QName -> Bool
forall a. Eq a => a -> a -> Bool
== QName
g = do
Maybe (Dom Type, Term, Type)
res <- Term
-> Type -> ProjOrigin -> QName -> m (Maybe (Dom Type, Term, Type))
forall (m :: * -> *).
(HasConstInfo m, MonadReduce m, MonadDebug 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) -> ProblemId -> Type -> Term -> Elims -> Elims -> m Term
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 -> m Term
forall (m :: * -> *) a. MonadError TCErr m => m a
patternViolation
antiUnifyElims ProblemId
pid Type
a Term
self (Apply Arg Term
u : Elims
es1) (Apply Arg Term
v : Elims
es2) = do
Term -> m Term
forall a (m :: * -> *). (Reduce a, MonadReduce m) => a -> m a
reduce (Type -> Term
forall t a. Type'' t a -> a
unEl Type
a) m Term -> (Term -> m Term) -> m Term
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 <- ProblemId -> Dom Type -> Arg Term -> Arg Term -> m (Arg Term)
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
ProblemId -> Type -> Term -> Elims -> Elims -> m Term
forall (m :: * -> *).
MonadConversion m =>
ProblemId -> Type -> Term -> Elims -> Elims -> m Term
antiUnifyElims ProblemId
pid (Abs Type
b Abs Type -> Term -> Type
forall t a. Subst t a => Abs a -> t -> a
`lazyAbsApp` Arg Term -> Term
forall e. Arg e -> e
unArg Arg Term
w) (Term -> [Arg Term] -> Term
forall t. Apply t => t -> [Arg Term] -> t
apply Term
self [Arg Term
w]) Elims
es1 Elims
es2
Term
_ -> m Term
forall (m :: * -> *) a. MonadError TCErr m => m a
patternViolation
antiUnifyElims ProblemId
_ Type
_ Term
_ Elims
_ Elims
_ = m Term
forall (m :: * -> *) a. MonadError TCErr m => m a
patternViolation
compareElims :: forall m. MonadConversion m => [Polarity] -> [IsForced] -> Type -> Term -> [Elim] -> [Elim] -> m ()
compareElims :: [Polarity] -> [IsForced] -> Type -> Term -> Elims -> Elims -> m ()
compareElims [Polarity]
pols0 [IsForced]
fors0 Type
a Term
v Elims
els01 Elims
els02 = (Constraint -> m () -> m ()
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 ()) (m () -> m ()) -> m () -> m ()
forall a b. (a -> b) -> a -> b
$ do
let v1 :: Term
v1 = Term -> Elims -> Term
forall t. Apply t => t -> Elims -> t
applyE Term
v Elims
els01
v2 :: Term
v2 = Term -> Elims -> Term
forall t. Apply t => t -> Elims -> t
applyE Term
v Elims
els02
failure :: m a
failure = TypeError -> m a
forall (m :: * -> *) a.
(MonadTCEnv m, ReadTCState m, MonadError TCErr m) =>
TypeError -> m a
typeError (TypeError -> m a) -> TypeError -> m a
forall a b. (a -> b) -> a -> b
$ Comparison -> Term -> Term -> CompareAs -> TypeError
UnequalTerms Comparison
CmpEq Term
v1 Term
v2 (Type -> CompareAs
AsTermsOf Type
a)
Bool -> m () -> m ()
forall (f :: * -> *). Applicative f => Bool -> f () -> f ()
unless (Elims -> Bool
forall (t :: * -> *) a. Foldable t => t a -> Bool
null Elims
els01) (m () -> m ()) -> m () -> m ()
forall a b. (a -> b) -> a -> b
$ do
VerboseKey -> Int -> TCM Doc -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCM Doc -> m ()
reportSDoc VerboseKey
"tc.conv.elim" Int
25 (TCM Doc -> m ()) -> TCM Doc -> m ()
forall a b. (a -> b) -> a -> b
$ TCM Doc
"compareElims" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
$$ do
Int -> TCM Doc -> TCM Doc
forall (m :: * -> *). Functor m => Int -> m Doc -> m Doc
nest Int
2 (TCM Doc -> TCM Doc) -> TCM Doc -> TCM Doc
forall a b. (a -> b) -> a -> b
$ [TCM Doc] -> TCM Doc
forall (m :: * -> *). Monad m => [m Doc] -> m Doc
vcat
[ TCM Doc
"a =" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Type -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Type
a
, TCM Doc
"pols0 (truncated to 10) =" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> [TCM Doc] -> TCM Doc
forall (m :: * -> *). Monad m => [m Doc] -> m Doc
hsep ((Polarity -> TCM Doc) -> [Polarity] -> [TCM Doc]
forall a b. (a -> b) -> [a] -> [b]
map Polarity -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM ([Polarity] -> [TCM Doc]) -> [Polarity] -> [TCM Doc]
forall a b. (a -> b) -> a -> b
$ Int -> [Polarity] -> [Polarity]
forall a. Int -> [a] -> [a]
take Int
10 [Polarity]
pols0)
, TCM Doc
"fors0 (truncated to 10) =" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> [TCM Doc] -> TCM Doc
forall (m :: * -> *). Monad m => [m Doc] -> m Doc
hsep ((IsForced -> TCM Doc) -> [IsForced] -> [TCM Doc]
forall a b. (a -> b) -> [a] -> [b]
map IsForced -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM ([IsForced] -> [TCM Doc]) -> [IsForced] -> [TCM Doc]
forall a b. (a -> b) -> a -> b
$ Int -> [IsForced] -> [IsForced]
forall a. Int -> [a] -> [a]
take Int
10 [IsForced]
fors0)
, TCM Doc
"v =" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Term -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Term
v
, TCM Doc
"els01 =" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Elims -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Elims
els01
, TCM Doc
"els02 =" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Elims -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Elims
els02
]
case (Elims
els01, Elims
els02) of
([] , [] ) -> () -> m ()
forall (m :: * -> *) a. Monad m => a -> m a
return ()
([] , Proj{}:Elims
_ ) -> m ()
forall a. m a
failure
(Proj{} : Elims
_, [] ) -> m ()
forall a. m a
failure
([] , Apply{} : Elims
_) -> m ()
forall a. m a
failure
(Apply{} : Elims
_, [] ) -> m ()
forall a. m a
failure
([] , IApply{} : Elims
_) -> m ()
forall a. m a
failure
(IApply{} : Elims
_, [] ) -> m ()
forall a. m a
failure
(Apply{} : Elims
_, Proj{} : Elims
_) -> ()
forall a. HasCallStack => a
__IMPOSSIBLE__ () -> m () -> m ()
forall (f :: * -> *) a b. Functor f => a -> f b -> f a
<$ Bool -> m ()
forall (m :: * -> *). MonadConstraint m => Bool -> m ()
solveAwakeConstraints' Bool
True
(Proj{} : Elims
_, Apply{} : Elims
_) -> ()
forall a. HasCallStack => a
__IMPOSSIBLE__ () -> m () -> m ()
forall (f :: * -> *) a b. Functor f => a -> f b -> f a
<$ Bool -> m ()
forall (m :: * -> *). MonadConstraint m => Bool -> m ()
solveAwakeConstraints' Bool
True
(IApply{} : Elims
_, Proj{} : Elims
_) -> ()
forall a. HasCallStack => a
__IMPOSSIBLE__ () -> m () -> m ()
forall (f :: * -> *) a b. Functor f => a -> f b -> f a
<$ Bool -> m ()
forall (m :: * -> *). MonadConstraint m => Bool -> m ()
solveAwakeConstraints' Bool
True
(Proj{} : Elims
_, IApply{} : Elims
_) -> ()
forall a. HasCallStack => a
__IMPOSSIBLE__ () -> m () -> m ()
forall (f :: * -> *) a b. Functor f => a -> f b -> f a
<$ Bool -> m ()
forall (m :: * -> *). MonadConstraint m => Bool -> m ()
solveAwakeConstraints' Bool
True
(IApply{} : Elims
_, Apply{} : Elims
_) -> ()
forall a. HasCallStack => a
__IMPOSSIBLE__ () -> m () -> m ()
forall (f :: * -> *) a b. Functor f => a -> f b -> f a
<$ Bool -> m ()
forall (m :: * -> *). MonadConstraint m => Bool -> m ()
solveAwakeConstraints' Bool
True
(Apply{} : Elims
_, IApply{} : Elims
_) -> ()
forall a. HasCallStack => a
__IMPOSSIBLE__ () -> m () -> m ()
forall (f :: * -> *) a b. Functor f => a -> f b -> f a
<$ Bool -> m ()
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
VerboseKey -> Int -> TCM Doc -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCM Doc -> m ()
reportSDoc VerboseKey
"tc.conv.elim" Int
25 (TCM Doc -> m ()) -> TCM Doc -> m ()
forall a b. (a -> b) -> a -> b
$ TCM Doc
"compareElims IApply"
let (Polarity
pol, [Polarity]
pols) = [Polarity] -> (Polarity, [Polarity])
nextPolarity [Polarity]
pols0
Type
-> (MetaId -> Type -> m ()) -> (NotBlocked -> Type -> m ()) -> m ()
forall t (m :: * -> *) a.
(Reduce t, IsMeta t, MonadReduce m, HasBuiltins m) =>
t -> (MetaId -> t -> m a) -> (NotBlocked -> t -> m a) -> m a
ifBlocked Type
a (\ MetaId
m Type
t -> m ()
forall (m :: * -> *) a. MonadError TCErr m => m a
patternViolation) ((NotBlocked -> Type -> m ()) -> m ())
-> (NotBlocked -> Type -> m ()) -> m ()
forall a b. (a -> b) -> a -> b
$ \ NotBlocked
_ Type
a -> do
PathView
va <- Type -> m PathView
forall (m :: * -> *). HasBuiltins m => Type -> m PathView
pathView Type
a
VerboseKey -> Int -> TCM Doc -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCM Doc -> m ()
reportSDoc VerboseKey
"tc.conv.elim.iapply" Int
60 (TCM Doc -> m ()) -> TCM Doc -> m ()
forall a b. (a -> b) -> a -> b
$ TCM Doc
"compareElims IApply" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
$$ do
Int -> TCM Doc -> TCM Doc
forall (m :: * -> *). Functor m => Int -> m Doc -> m Doc
nest Int
2 (TCM Doc -> TCM Doc) -> TCM Doc -> TCM Doc
forall a b. (a -> b) -> a -> b
$ TCM Doc
"va =" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> VerboseKey -> TCM Doc
forall (m :: * -> *). Monad m => VerboseKey -> m Doc
text (Bool -> VerboseKey
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 <- m Term -> m Type
forall (tcm :: * -> *). Functor tcm => tcm Term -> tcm Type
elInf m Term
forall (m :: * -> *).
(HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) =>
m Term
primInterval
Polarity
-> (Comparison -> Term -> Term -> m ()) -> Term -> Term -> m ()
forall (m :: * -> *) a.
MonadConversion m =>
Polarity -> (Comparison -> a -> a -> m ()) -> a -> a -> m ()
compareWithPol Polarity
pol ((Comparison -> Type -> Term -> Term -> m ())
-> Type -> Comparison -> Term -> Term -> m ()
forall a b c. (a -> b -> c) -> b -> a -> c
flip Comparison -> Type -> Term -> Term -> m ()
forall (m :: * -> *).
MonadConversion m =>
Comparison -> Type -> Term -> Term -> m ()
compareTerm Type
b)
Term
r1 Term
r2
let r :: Term
r = Term
r1
Type
codom <- m Term -> m Term -> m Type
forall (m :: * -> *). Monad m => m Term -> m Term -> m Type
el' (Term -> m Term
forall (f :: * -> *) a. Applicative f => a -> f a
pure (Term -> m Term) -> (Arg Term -> Term) -> Arg Term -> m Term
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Arg Term -> Term
forall e. Arg e -> e
unArg (Arg Term -> m Term) -> Arg Term -> m Term
forall a b. (a -> b) -> a -> b
$ Arg Term
l) ((Term -> m Term
forall (f :: * -> *) a. Applicative f => a -> f a
pure (Term -> m Term) -> (Arg Term -> Term) -> Arg Term -> m Term
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Arg Term -> Term
forall e. Arg e -> e
unArg (Arg Term -> m Term) -> Arg Term -> m Term
forall a b. (a -> b) -> a -> b
$ Arg Term
bA) m Term -> m Term -> m Term
forall (tcm :: * -> *).
Monad tcm =>
tcm Term -> tcm Term -> tcm Term
<@> Term -> m Term
forall (f :: * -> *) a. Applicative f => a -> f a
pure Term
r)
[Polarity] -> [IsForced] -> Type -> Term -> Elims -> Elims -> m ()
forall (m :: * -> *).
MonadConversion m =>
[Polarity] -> [IsForced] -> Type -> Term -> Elims -> Elims -> m ()
compareElims [Polarity]
pols [] Type
codom
(Term -> Elims -> Term
forall t. Apply t => t -> Elims -> t
applyE Term
v [Elim' Term
e]) Elims
els1 Elims
els2
OType t :: Type
t@(El Sort
_ Pi{}) -> [Polarity] -> [IsForced] -> Type -> Term -> Elims -> Elims -> m ()
forall (m :: * -> *).
MonadConversion m =>
[Polarity] -> [IsForced] -> Type -> Term -> Elims -> Elims -> m ()
compareElims [Polarity]
pols0 [IsForced]
fors0 Type
t Term
v (Arg Term -> Elim' Term
forall a. Arg a -> Elim' a
Apply (Term -> Arg Term
forall e. e -> Arg e
defaultArg Term
r1) Elim' Term -> Elims -> Elims
forall a. a -> [a] -> [a]
: Elims
els1) (Arg Term -> Elim' Term
forall a. Arg a -> Elim' a
Apply (Term -> Arg Term
forall e. e -> Arg e
defaultArg Term
r2) Elim' Term -> Elims -> Elims
forall a. a -> [a] -> [a]
: Elims
els2)
OType{} -> m ()
forall (m :: * -> *) a. MonadError TCErr m => m a
patternViolation
(Apply Arg Term
arg1 : Elims
els1, Apply Arg Term
arg2 : Elims
els2) ->
(VerboseKey -> Int -> VerboseKey -> m () -> m ()
forall (m :: * -> *) a.
MonadDebug m =>
VerboseKey -> Int -> VerboseKey -> m a -> m a
verboseBracket VerboseKey
"tc.conv.elim" Int
20 VerboseKey
"compare Apply" :: m () -> m ()) (m () -> m ()) -> m () -> m ()
forall a b. (a -> b) -> a -> b
$ do
VerboseKey -> Int -> TCM Doc -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCM Doc -> m ()
reportSDoc VerboseKey
"tc.conv.elim" Int
10 (TCM Doc -> m ()) -> TCM Doc -> m ()
forall a b. (a -> b) -> a -> b
$ Int -> TCM Doc -> TCM Doc
forall (m :: * -> *). Functor m => Int -> m Doc -> m Doc
nest Int
2 (TCM Doc -> TCM Doc) -> TCM Doc -> TCM Doc
forall a b. (a -> b) -> a -> b
$ [TCM Doc] -> TCM Doc
forall (m :: * -> *). Monad m => [m Doc] -> m Doc
vcat
[ TCM Doc
"a =" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Type -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Type
a
, TCM Doc
"v =" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Term -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Term
v
, TCM Doc
"arg1 =" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Arg Term -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Arg Term
arg1
, TCM Doc
"arg2 =" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Arg Term -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Arg Term
arg2
]
VerboseKey -> Int -> TCM Doc -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCM Doc -> m ()
reportSDoc VerboseKey
"tc.conv.elim" Int
50 (TCM Doc -> m ()) -> TCM Doc -> m ()
forall a b. (a -> b) -> a -> b
$ Int -> TCM Doc -> TCM Doc
forall (m :: * -> *). Functor m => Int -> m Doc -> m Doc
nest Int
2 (TCM Doc -> TCM Doc) -> TCM Doc -> TCM Doc
forall a b. (a -> b) -> a -> b
$ [TCM Doc] -> TCM Doc
forall (m :: * -> *). Monad m => [m Doc] -> m Doc
vcat
[ TCM Doc
"raw:"
, TCM Doc
"a =" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Type -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty Type
a
, TCM Doc
"v =" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Term -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty Term
v
, TCM Doc
"arg1 =" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Arg Term -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty Arg Term
arg1
, TCM Doc
"arg2 =" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Arg Term -> TCM Doc
forall (m :: * -> *) a. (Monad 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
-> (MetaId -> Type -> m ()) -> (NotBlocked -> Type -> m ()) -> m ()
forall t (m :: * -> *) a.
(Reduce t, IsMeta t, MonadReduce m, HasBuiltins m) =>
t -> (MetaId -> t -> m a) -> (NotBlocked -> t -> m a) -> m a
ifBlocked Type
a (\ MetaId
m Type
t -> m ()
forall (m :: * -> *) a. MonadError TCErr m => m a
patternViolation) ((NotBlocked -> Type -> m ()) -> m ())
-> (NotBlocked -> Type -> m ()) -> m ()
forall a b. (a -> b) -> a -> b
$ \ NotBlocked
_ Type
a -> do
VerboseKey -> Int -> VerboseKey -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> VerboseKey -> m ()
reportSLn VerboseKey
"tc.conv.elim" Int
90 (VerboseKey -> m ()) -> VerboseKey -> m ()
forall a b. (a -> b) -> a -> b
$ VerboseKey
"type is not blocked"
case Type -> Term
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
VerboseKey -> Int -> VerboseKey -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> VerboseKey -> m ()
reportSLn VerboseKey
"tc.conv.elim" Int
90 (VerboseKey -> m ()) -> VerboseKey -> m ()
forall a b. (a -> b) -> a -> b
$ VerboseKey
"type is a function type"
Maybe Term
mlvl <- m Term -> m (Maybe Term)
forall e (m :: * -> *) a.
(MonadError e m, Functor m) =>
m a -> m (Maybe a)
tryMaybe m Term
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 Int -> a -> Bool
forall a. Free a => Int -> a -> Bool
`freeInIgnoringSorts` a
c
freeInCoDom Abs a
_ = Bool
False
dependent :: Bool
dependent = (Term -> Maybe Term
forall a. a -> Maybe a
Just (Type -> Term
forall t a. Type'' t a -> a
unEl Type
b) Maybe Term -> Maybe Term -> Bool
forall a. Eq a => a -> a -> Bool
/= Maybe Term
mlvl) Bool -> Bool -> Bool
&& Abs Type -> Bool
forall a. Free a => Abs a -> Bool
freeInCoDom Abs Type
codom
ProblemId
pid <- m () -> m ProblemId
forall (m :: * -> *) a.
(MonadFresh ProblemId m, MonadConstraint m) =>
m a -> m ProblemId
newProblem_ (m () -> m ProblemId) -> m () -> m ProblemId
forall a b. (a -> b) -> a -> b
$ ArgInfo -> m () -> m ()
forall (tcm :: * -> *) m a.
(MonadTCEnv tcm, LensModality m) =>
m -> tcm a -> tcm a
applyModalityToContext ArgInfo
info (m () -> m ()) -> m () -> m ()
forall a b. (a -> b) -> a -> b
$
if IsForced -> Bool
isForced IsForced
for then
VerboseKey -> Int -> VerboseKey -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> VerboseKey -> m ()
reportSLn VerboseKey
"tc.conv.elim" Int
90 (VerboseKey -> m ()) -> VerboseKey -> m ()
forall a b. (a -> b) -> a -> b
$ VerboseKey
"argument is forced"
else if ArgInfo -> Bool
forall a. LensRelevance a => a -> Bool
isIrrelevant ArgInfo
info then do
VerboseKey -> Int -> VerboseKey -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> VerboseKey -> m ()
reportSLn VerboseKey
"tc.conv.elim" Int
90 (VerboseKey -> m ()) -> VerboseKey -> m ()
forall a b. (a -> b) -> a -> b
$ VerboseKey
"argument is irrelevant"
Type -> Term -> Term -> m ()
forall (m :: * -> *).
MonadConversion m =>
Type -> Term -> Term -> m ()
compareIrrelevant Type
b (Arg Term -> Term
forall e. Arg e -> e
unArg Arg Term
arg1) (Arg Term -> Term
forall e. Arg e -> e
unArg Arg Term
arg2)
else do
VerboseKey -> Int -> VerboseKey -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> VerboseKey -> m ()
reportSLn VerboseKey
"tc.conv.elim" Int
90 (VerboseKey -> m ()) -> VerboseKey -> m ()
forall a b. (a -> b) -> a -> b
$ VerboseKey
"argument has polarity " VerboseKey -> VerboseKey -> VerboseKey
forall a. [a] -> [a] -> [a]
++ Polarity -> VerboseKey
forall a. Show a => a -> VerboseKey
show Polarity
pol
Polarity
-> (Comparison -> Term -> Term -> m ()) -> Term -> Term -> m ()
forall (m :: * -> *) a.
MonadConversion m =>
Polarity -> (Comparison -> a -> a -> m ()) -> a -> a -> m ()
compareWithPol Polarity
pol ((Comparison -> Type -> Term -> Term -> m ())
-> Type -> Comparison -> Term -> Term -> m ()
forall a b c. (a -> b -> c) -> b -> a -> c
flip Comparison -> Type -> Term -> Term -> m ()
forall (m :: * -> *).
MonadConversion m =>
Comparison -> Type -> Term -> Term -> m ()
compareTerm Type
b)
(Arg Term -> Term
forall e. Arg e -> e
unArg Arg Term
arg1) (Arg Term -> Term
forall e. Arg e -> e
unArg Arg Term
arg2)
Bool
solved <- ProblemId -> m Bool
forall (m :: * -> *).
(MonadTCEnv m, ReadTCState m) =>
ProblemId -> m Bool
isProblemSolved ProblemId
pid
VerboseKey -> Int -> VerboseKey -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> VerboseKey -> m ()
reportSLn VerboseKey
"tc.conv.elim" Int
90 (VerboseKey -> m ()) -> VerboseKey -> m ()
forall a b. (a -> b) -> a -> b
$ VerboseKey
"solved = " VerboseKey -> VerboseKey -> VerboseKey
forall a. [a] -> [a] -> [a]
++ Bool -> VerboseKey
forall a. Show a => a -> VerboseKey
show Bool
solved
Arg Term
arg <- if Bool
dependent Bool -> Bool -> Bool
&& Bool -> Bool
not Bool
solved
then ArgInfo -> m (Arg Term) -> m (Arg Term)
forall (tcm :: * -> *) m a.
(MonadTCEnv tcm, LensModality m) =>
m -> tcm a -> tcm a
applyModalityToContext ArgInfo
info (m (Arg Term) -> m (Arg Term)) -> m (Arg Term) -> m (Arg Term)
forall a b. (a -> b) -> a -> b
$ do
VerboseKey -> Int -> TCM Doc -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCM Doc -> m ()
reportSDoc VerboseKey
"tc.conv.elims" Int
30 (TCM Doc -> m ()) -> TCM Doc -> m ()
forall a b. (a -> b) -> a -> b
$ [TCM Doc] -> TCM Doc
forall (m :: * -> *). Monad m => [m Doc] -> m Doc
vcat ([TCM Doc] -> TCM Doc) -> [TCM Doc] -> TCM Doc
forall a b. (a -> b) -> a -> b
$
[ TCM Doc
"Trying antiUnify:"
, Int -> TCM Doc -> TCM Doc
forall (m :: * -> *). Functor m => Int -> m Doc -> m Doc
nest Int
2 (TCM Doc -> TCM Doc) -> TCM Doc -> TCM Doc
forall a b. (a -> b) -> a -> b
$ TCM Doc
"b =" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Type -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Type
b
, Int -> TCM Doc -> TCM Doc
forall (m :: * -> *). Functor m => Int -> m Doc -> m Doc
nest Int
2 (TCM Doc -> TCM Doc) -> TCM Doc -> TCM Doc
forall a b. (a -> b) -> a -> b
$ TCM Doc
"arg1 =" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Arg Term -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Arg Term
arg1
, Int -> TCM Doc -> TCM Doc
forall (m :: * -> *). Functor m => Int -> m Doc -> m Doc
nest Int
2 (TCM Doc -> TCM Doc) -> TCM Doc -> TCM Doc
forall a b. (a -> b) -> a -> b
$ TCM Doc
"arg2 =" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Arg Term -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Arg Term
arg2
]
Arg Term
arg <- (Arg Term
arg1 Arg Term -> Term -> Arg Term
forall (f :: * -> *) a b. Functor f => f a -> b -> f b
$>) (Term -> Arg Term) -> m Term -> m (Arg Term)
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> ProblemId -> Type -> Term -> Term -> m Term
forall (m :: * -> *).
MonadConversion m =>
ProblemId -> Type -> Term -> Term -> m Term
antiUnify ProblemId
pid Type
b (Arg Term -> Term
forall e. Arg e -> e
unArg Arg Term
arg1) (Arg Term -> Term
forall e. Arg e -> e
unArg Arg Term
arg2)
VerboseKey -> Int -> TCM Doc -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCM Doc -> m ()
reportSDoc VerboseKey
"tc.conv.elims" Int
30 (TCM Doc -> m ()) -> TCM Doc -> m ()
forall a b. (a -> b) -> a -> b
$ TCM Doc -> Int -> TCM Doc -> TCM Doc
forall (m :: * -> *).
Applicative m =>
m Doc -> Int -> m Doc -> m Doc
hang TCM Doc
"Anti-unification:" Int
2 (Arg Term -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Arg Term
arg)
VerboseKey -> Int -> TCM Doc -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCM Doc -> m ()
reportSDoc VerboseKey
"tc.conv.elims" Int
70 (TCM Doc -> m ()) -> TCM Doc -> m ()
forall a b. (a -> b) -> a -> b
$ Int -> TCM Doc -> TCM Doc
forall (m :: * -> *). Functor m => Int -> m Doc -> m Doc
nest Int
2 (TCM Doc -> TCM Doc) -> TCM Doc -> TCM Doc
forall a b. (a -> b) -> a -> b
$ TCM Doc
"raw:" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Arg Term -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty Arg Term
arg
Arg Term -> m (Arg Term)
forall (m :: * -> *) a. Monad m => a -> m a
return Arg Term
arg
else Arg Term -> m (Arg Term)
forall (m :: * -> *) a. Monad m => a -> m a
return Arg Term
arg1
[Polarity] -> [IsForced] -> Type -> Term -> Elims -> Elims -> m ()
forall (m :: * -> *).
MonadConversion m =>
[Polarity] -> [IsForced] -> Type -> Term -> Elims -> Elims -> m ()
compareElims [Polarity]
pols [IsForced]
fors (Abs Type
codom Abs Type -> Term -> Type
forall t a. Subst t a => Abs a -> t -> a
`lazyAbsApp` Arg Term -> Term
forall e. Arg e -> e
unArg Arg Term
arg) (Term -> [Arg Term] -> Term
forall t. Apply t => t -> [Arg Term] -> t
apply Term
v [Arg Term
arg]) Elims
els1 Elims
els2
VerboseKey -> Int -> VerboseKey -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> VerboseKey -> m ()
reportSLn VerboseKey
"tc.conv.elim" Int
90 (VerboseKey -> m ()) -> VerboseKey -> m ()
forall a b. (a -> b) -> a -> b
$ VerboseKey
"stealing constraints from problem " VerboseKey -> VerboseKey -> VerboseKey
forall a. [a] -> [a] -> [a]
++ ProblemId -> VerboseKey
forall a. Show a => a -> VerboseKey
show ProblemId
pid
ProblemId -> m ()
forall (m :: * -> *). MonadConstraint m => ProblemId -> m ()
stealConstraints ProblemId
pid
Term
a -> do
VerboseKey -> Int -> TCM Doc -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCM Doc -> m ()
reportSDoc VerboseKey
"impossible" Int
10 (TCM Doc -> m ()) -> TCM Doc -> m ()
forall a b. (a -> b) -> a -> b
$
TCM Doc
"unexpected type when comparing apply eliminations " TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Term -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Term
a
VerboseKey -> Int -> TCM Doc -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCM Doc -> m ()
reportSDoc VerboseKey
"impossible" Int
50 (TCM Doc -> m ()) -> TCM Doc -> m ()
forall a b. (a -> b) -> a -> b
$ TCM Doc
"raw type:" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Term -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty Term
a
m ()
forall (m :: * -> *) a. MonadError TCErr m => m a
patternViolation
(Proj ProjOrigin
o QName
f : Elims
els1, Proj ProjOrigin
_ QName
f' : Elims
els2)
| QName
f QName -> QName -> Bool
forall a. Eq a => a -> a -> Bool
/= QName
f' -> TypeError -> m ()
forall (m :: * -> *) a.
(MonadTCEnv m, ReadTCState m, MonadError TCErr m) =>
TypeError -> m a
typeError (TypeError -> m ()) -> (Doc -> TypeError) -> Doc -> m ()
forall b c a. (b -> c) -> (a -> b) -> a -> c
. VerboseKey -> TypeError
GenericError (VerboseKey -> TypeError)
-> (Doc -> VerboseKey) -> Doc -> TypeError
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Doc -> VerboseKey
forall a. Show a => a -> VerboseKey
show (Doc -> m ()) -> m Doc -> m ()
forall (m :: * -> *) a b. Monad m => (a -> m b) -> m a -> m b
=<< QName -> m Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM QName
f m Doc -> m Doc -> m Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> m Doc
"/=" m Doc -> m Doc -> m Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> QName -> m Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM QName
f'
| Bool
otherwise -> Type
-> (MetaId -> Type -> m ()) -> (NotBlocked -> Type -> m ()) -> m ()
forall t (m :: * -> *) a.
(Reduce t, IsMeta t, MonadReduce m, HasBuiltins m) =>
t -> (MetaId -> t -> m a) -> (NotBlocked -> t -> m a) -> m a
ifBlocked Type
a (\ MetaId
m Type
t -> m ()
forall (m :: * -> *) a. MonadError TCErr m => m a
patternViolation) ((NotBlocked -> Type -> m ()) -> m ())
-> (NotBlocked -> Type -> m ()) -> m ()
forall a b. (a -> b) -> a -> b
$ \ NotBlocked
_ Type
a -> do
Maybe (Dom Type, Term, Type)
res <- Term
-> Type -> ProjOrigin -> QName -> m (Maybe (Dom Type, Term, Type))
forall (m :: * -> *).
(HasConstInfo m, MonadReduce m, MonadDebug m) =>
Term
-> Type -> ProjOrigin -> QName -> m (Maybe (Dom Type, Term, Type))
projectTyped Term
v Type
a ProjOrigin
o QName
f
case Maybe (Dom Type, Term, Type)
res of
Just (Dom Type
_, Term
u, Type
t) -> do
[Polarity] -> [IsForced] -> Type -> Term -> Elims -> Elims -> m ()
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
VerboseKey -> Int -> TCM Doc -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCM Doc -> m ()
reportSDoc VerboseKey
"tc.conv.elims" Int
30 (TCM Doc -> m ()) -> TCM Doc -> m ()
forall a b. (a -> b) -> a -> b
$ [TCM Doc] -> TCM Doc
forall (m :: * -> *). Monad m => [m Doc] -> m Doc
sep
[ VerboseKey -> TCM Doc
forall (m :: * -> *). Monad m => VerboseKey -> m Doc
text (VerboseKey -> TCM Doc) -> VerboseKey -> TCM Doc
forall a b. (a -> b) -> a -> b
$ VerboseKey
"projection " VerboseKey -> VerboseKey -> VerboseKey
forall a. [a] -> [a] -> [a]
++ QName -> VerboseKey
forall a. Show a => a -> VerboseKey
show QName
f
, VerboseKey -> TCM Doc
forall (m :: * -> *). Monad m => VerboseKey -> m Doc
text VerboseKey
"applied to value " TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Term -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Term
v
, VerboseKey -> TCM Doc
forall (m :: * -> *). Monad m => VerboseKey -> m Doc
text VerboseKey
"of unexpected type " TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Type -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Type
a
]
m ()
forall (m :: * -> *) a. MonadError TCErr m => m a
patternViolation
compareIrrelevant :: MonadConversion m => Type -> Term -> Term -> m ()
compareIrrelevant :: 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
VerboseKey -> Int -> TCM Doc -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCM Doc -> m ()
reportSDoc VerboseKey
"tc.conv.irr" Int
20 (TCM Doc -> m ()) -> TCM Doc -> m ()
forall a b. (a -> b) -> a -> b
$ [TCM Doc] -> TCM Doc
forall (m :: * -> *). Monad m => [m Doc] -> m Doc
vcat
[ TCM Doc
"compareIrrelevant"
, Int -> TCM Doc -> TCM Doc
forall (m :: * -> *). Functor m => Int -> m Doc -> m Doc
nest Int
2 (TCM Doc -> TCM Doc) -> TCM Doc -> TCM Doc
forall a b. (a -> b) -> a -> b
$ TCM Doc
"v =" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Term -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Term
v
, Int -> TCM Doc -> TCM Doc
forall (m :: * -> *). Functor m => Int -> m Doc -> m Doc
nest Int
2 (TCM Doc -> TCM Doc) -> TCM Doc -> TCM Doc
forall a b. (a -> b) -> a -> b
$ TCM Doc
"w =" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Term -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Term
w
]
VerboseKey -> Int -> TCM Doc -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCM Doc -> m ()
reportSDoc VerboseKey
"tc.conv.irr" Int
50 (TCM Doc -> m ()) -> TCM Doc -> m ()
forall a b. (a -> b) -> a -> b
$ [TCM Doc] -> TCM Doc
forall (m :: * -> *). Monad m => [m Doc] -> m Doc
vcat
[ Int -> TCM Doc -> TCM Doc
forall (m :: * -> *). Functor m => Int -> m Doc -> m Doc
nest Int
2 (TCM Doc -> TCM Doc) -> TCM Doc -> TCM Doc
forall a b. (a -> b) -> a -> b
$ TCM Doc
"v =" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Term -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty Term
v
, Int -> TCM Doc -> TCM Doc
forall (m :: * -> *). Functor m => Int -> m Doc -> m Doc
nest Int
2 (TCM Doc -> TCM Doc) -> TCM Doc -> TCM Doc
forall a b. (a -> b) -> a -> b
$ TCM Doc
"w =" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Term -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty Term
w
]
Term -> Term -> m () -> m ()
forall (m :: * -> *).
(MonadMetaSolver m, MonadWarning m, MonadStatistics m,
MonadFresh ProblemId m, MonadFresh Int m) =>
Term -> Term -> m () -> m ()
try Term
v Term
w (m () -> m ()) -> m () -> m ()
forall a b. (a -> b) -> a -> b
$ Term -> Term -> m () -> m ()
forall (m :: * -> *).
(MonadMetaSolver m, MonadWarning m, MonadStatistics m,
MonadFresh ProblemId m, MonadFresh Int m) =>
Term -> Term -> m () -> m ()
try Term
w Term
v (m () -> m ()) -> m () -> m ()
forall a b. (a -> b) -> a -> b
$ () -> m ()
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
MetaVariable
mv <- MetaId -> m MetaVariable
forall (m :: * -> *).
(MonadFail m, ReadTCState m) =>
MetaId -> m MetaVariable
lookupMeta MetaId
x
let rel :: Relevance
rel = MetaVariable -> Relevance
getMetaRelevance MetaVariable
mv
inst :: Bool
inst = case MetaVariable -> MetaInstantiation
mvInstantiation MetaVariable
mv of
InstV{} -> Bool
True
MetaInstantiation
_ -> Bool
False
VerboseKey -> Int -> TCM Doc -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCM Doc -> m ()
reportSDoc VerboseKey
"tc.conv.irr" Int
20 (TCM Doc -> m ()) -> TCM Doc -> m ()
forall a b. (a -> b) -> a -> b
$ [TCM Doc] -> TCM Doc
forall (m :: * -> *). Monad m => [m Doc] -> m Doc
vcat
[ Int -> TCM Doc -> TCM Doc
forall (m :: * -> *). Functor m => Int -> m Doc -> m Doc
nest Int
2 (TCM Doc -> TCM Doc) -> TCM Doc -> TCM Doc
forall a b. (a -> b) -> a -> b
$ VerboseKey -> TCM Doc
forall (m :: * -> *). Monad m => VerboseKey -> m Doc
text (VerboseKey -> TCM Doc) -> VerboseKey -> TCM Doc
forall a b. (a -> b) -> a -> b
$ VerboseKey
"rel = " VerboseKey -> VerboseKey -> VerboseKey
forall a. [a] -> [a] -> [a]
++ Relevance -> VerboseKey
forall a. Show a => a -> VerboseKey
show Relevance
rel
, Int -> TCM Doc -> TCM Doc
forall (m :: * -> *). Functor m => Int -> m Doc -> m Doc
nest Int
2 (TCM Doc -> TCM Doc) -> TCM Doc -> TCM Doc
forall a b. (a -> b) -> a -> b
$ TCM Doc
"inst =" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Bool -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty Bool
inst
]
if Bool -> Bool
not (Relevance -> Bool
forall a. LensRelevance a => a -> Bool
isIrrelevant Relevance
rel) Bool -> Bool -> Bool
|| Bool
inst
then m ()
fallback
else (CompareDirection
-> MetaId
-> Elims
-> Term
-> CompareAs
-> (Term -> Term -> m ())
-> m ()
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) ((Term -> Term -> m ()) -> m ()) -> (Term -> Term -> m ()) -> m ()
forall a b. (a -> b) -> a -> b
$ Type -> Term -> Term -> m ()
forall (m :: * -> *).
MonadConversion m =>
Type -> Term -> Term -> m ()
compareIrrelevant Type
t) m () -> (TCErr -> m ()) -> m ()
forall e (m :: * -> *) a.
MonadError e m =>
m a -> (e -> m a) -> m a
`catchError` \ TCErr
_ -> m ()
fallback
try Term
v Term
w m ()
fallback = m ()
fallback
compareWithPol :: MonadConversion m => Polarity -> (Comparison -> a -> a -> m ()) -> a -> a -> m ()
compareWithPol :: 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 = () -> m ()
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
compareArgs :: MonadConversion m => [Polarity] -> [IsForced] -> Type -> Term -> Args -> Args -> m ()
compareArgs :: [Polarity]
-> [IsForced] -> Type -> Term -> [Arg Term] -> [Arg Term] -> m ()
compareArgs [Polarity]
pol [IsForced]
for Type
a Term
v [Arg Term]
args1 [Arg Term]
args2 =
[Polarity] -> [IsForced] -> Type -> Term -> Elims -> Elims -> m ()
forall (m :: * -> *).
MonadConversion m =>
[Polarity] -> [IsForced] -> Type -> Term -> Elims -> Elims -> m ()
compareElims [Polarity]
pol [IsForced]
for Type
a Term
v ((Arg Term -> Elim' Term) -> [Arg Term] -> Elims
forall a b. (a -> b) -> [a] -> [b]
map Arg Term -> Elim' Term
forall a. Arg a -> Elim' a
Apply [Arg Term]
args1) ((Arg Term -> Elim' Term) -> [Arg Term] -> Elims
forall a b. (a -> b) -> [a] -> [b]
map Arg Term -> Elim' Term
forall a. Arg a -> Elim' a
Apply [Arg Term]
args2)
compareType :: MonadConversion m => Comparison -> Type -> Type -> m ()
compareType :: Comparison -> Type -> Type -> m ()
compareType Comparison
cmp ty1 :: Type
ty1@(El Sort
s1 Term
a1) ty2 :: Type
ty2@(El Sort
s2 Term
a2) =
m () -> m ()
forall (m :: * -> *) a.
(MonadTCEnv m, HasOptions m, MonadDebug m) =>
m a -> m a
workOnTypes (m () -> m ()) -> m () -> m ()
forall a b. (a -> b) -> a -> b
$
VerboseKey -> Int -> VerboseKey -> m () -> m ()
forall (m :: * -> *) a.
MonadDebug m =>
VerboseKey -> Int -> VerboseKey -> m a -> m a
verboseBracket VerboseKey
"tc.conv.type" Int
20 VerboseKey
"compareType" (m () -> m ()) -> m () -> m ()
forall a b. (a -> b) -> a -> b
$ do
VerboseKey -> Int -> TCM Doc -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCM Doc -> m ()
reportSDoc VerboseKey
"tc.conv.type" Int
50 (TCM Doc -> m ()) -> TCM Doc -> m ()
forall a b. (a -> b) -> a -> b
$ [TCM Doc] -> TCM Doc
forall (m :: * -> *). Monad m => [m Doc] -> m Doc
vcat
[ TCM Doc
"compareType" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> [TCM Doc] -> TCM Doc
forall (m :: * -> *). Monad m => [m Doc] -> m Doc
sep [ Type -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Type
ty1 TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Comparison -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Comparison
cmp
, Type -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Type
ty2 ]
, [TCM Doc] -> TCM Doc
forall (m :: * -> *). Monad m => [m Doc] -> m Doc
hsep [ TCM Doc
" sorts:", Sort -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Sort
s1, TCM Doc
" and ", Sort -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Sort
s2 ]
]
Comparison -> CompareAs -> Term -> Term -> m ()
forall (m :: * -> *).
MonadConversion m =>
Comparison -> CompareAs -> Term -> Term -> m ()
compareAs Comparison
cmp CompareAs
AsTypes Term
a1 Term
a2
m Bool -> m () -> m ()
forall (m :: * -> *). Monad m => m Bool -> m () -> m ()
unlessM ((PragmaOptions -> Bool
optCumulativity (PragmaOptions -> Bool) -> m PragmaOptions -> m Bool
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> m PragmaOptions
forall (m :: * -> *). HasOptions m => m PragmaOptions
pragmaOptions) m Bool -> m Bool -> m Bool
forall (m :: * -> *). Monad m => m Bool -> m Bool -> m Bool
`or2M`
(Bool -> Bool
not (Bool -> Bool) -> (PragmaOptions -> Bool) -> PragmaOptions -> Bool
forall b c a. (b -> c) -> (a -> b) -> a -> c
. PragmaOptions -> Bool
optCompareSorts (PragmaOptions -> Bool) -> m PragmaOptions -> m Bool
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> m PragmaOptions
forall (m :: * -> *). HasOptions m => m PragmaOptions
pragmaOptions)) (m () -> m ()) -> m () -> m ()
forall a b. (a -> b) -> a -> b
$
Comparison -> Sort -> Sort -> m ()
forall (m :: * -> *).
MonadConversion m =>
Comparison -> Sort -> Sort -> m ()
compareSort Comparison
CmpEq Sort
s1 Sort
s2
() -> m ()
forall (m :: * -> *) a. Monad m => a -> m a
return ()
leqType :: MonadConversion m => Type -> Type -> m ()
leqType :: Type -> Type -> m ()
leqType = Comparison -> Type -> Type -> m ()
forall (m :: * -> *).
MonadConversion m =>
Comparison -> Type -> Type -> m ()
compareType Comparison
CmpLeq
coerce :: (MonadConversion m, MonadTCM m) => Comparison -> Term -> Type -> Type -> m Term
coerce :: Comparison -> Term -> Type -> Type -> m Term
coerce Comparison
cmp Term
v Type
t1 Type
t2 = Type -> m Term -> m Term
forall (m :: * -> *).
(MonadMetaSolver m, MonadConstraint m, MonadFresh Int m,
MonadFresh ProblemId m) =>
Type -> m Term -> m Term
blockTerm Type
t2 (m Term -> m Term) -> m Term -> m Term
forall a b. (a -> b) -> a -> b
$ do
VerboseKey -> Int -> m () -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> m () -> m ()
verboseS VerboseKey
"tc.conv.coerce" Int
10 (m () -> m ()) -> m () -> m ()
forall a b. (a -> b) -> a -> b
$ do
(Expr
a1,Expr
a2) <- (Type, Type) -> m (Expr, Expr)
forall i a (m :: * -> *). (Reify i a, MonadReify m) => i -> m a
reify (Type
t1,Type
t2)
let dbglvl :: Int
dbglvl = if Expr -> Bool
isSet Expr
a1 Bool -> Bool -> Bool
&& Expr -> Bool
isSet Expr
a2 then Int
50 else Int
10
VerboseKey -> Int -> TCM Doc -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCM Doc -> m ()
reportSDoc VerboseKey
"tc.conv.coerce" Int
dbglvl (TCM Doc -> m ()) -> TCM Doc -> m ()
forall a b. (a -> b) -> a -> b
$
TCM Doc
"coerce" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> [TCM Doc] -> TCM Doc
forall (m :: * -> *). Monad m => [m Doc] -> m Doc
vcat
[ TCM Doc
"term v =" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Term -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Term
v
, TCM Doc
"from type t1 =" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Expr -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Expr
a1
, TCM Doc
"to type t2 =" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Expr -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Expr
a2
, TCM Doc
"comparison =" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Comparison -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Comparison
cmp
]
VerboseKey -> Int -> TCM Doc -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCM Doc -> m ()
reportSDoc VerboseKey
"tc.conv.coerce" Int
70 (TCM Doc -> m ()) -> TCM Doc -> m ()
forall a b. (a -> b) -> a -> b
$
TCM Doc
"coerce" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> [TCM Doc] -> TCM Doc
forall (m :: * -> *). Monad m => [m Doc] -> m Doc
vcat
[ TCM Doc
"term v =" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Term -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty Term
v
, TCM Doc
"from type t1 =" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Type -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty Type
t1
, TCM Doc
"to type t2 =" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Type -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty Type
t2
, TCM Doc
"comparison =" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Comparison -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty Comparison
cmp
]
TelV Telescope
tel1 Type
b1 <- Int -> (Dom Type -> Bool) -> Type -> m (TelV Type)
forall (m :: * -> *).
(MonadReduce m, MonadAddContext m) =>
Int -> (Dom Type -> Bool) -> Type -> m (TelV Type)
telViewUpTo' (-Int
1) Dom Type -> Bool
forall a. LensHiding a => a -> Bool
notVisible Type
t1
TelV Telescope
tel2 Type
b2 <- Int -> (Dom Type -> Bool) -> Type -> m (TelV Type)
forall (m :: * -> *).
(MonadReduce m, MonadAddContext m) =>
Int -> (Dom Type -> Bool) -> Type -> m (TelV Type)
telViewUpTo' (-Int
1) Dom Type -> Bool
forall a. LensHiding a => a -> Bool
notVisible Type
t2
let n :: Int
n = Telescope -> Int
forall a. Sized a => a -> Int
size Telescope
tel1 Int -> Int -> Int
forall a. Num a => a -> a -> a
- Telescope -> Int
forall a. Sized a => a -> Int
size Telescope
tel2
if Int
n Int -> Int -> Bool
forall a. Ord a => a -> a -> Bool
<= Int
0 then m Term
fallback else do
Type
-> (MetaId -> Type -> m Term)
-> (NotBlocked -> Type -> m Term)
-> m Term
forall t (m :: * -> *) a.
(Reduce t, IsMeta t, MonadReduce m, HasBuiltins m) =>
t -> (MetaId -> t -> m a) -> (NotBlocked -> t -> m a) -> m a
ifBlocked Type
b2 (\ MetaId
_ Type
_ -> m Term
fallback) ((NotBlocked -> Type -> m Term) -> m Term)
-> (NotBlocked -> Type -> m Term) -> m Term
forall a b. (a -> b) -> a -> b
$ \ NotBlocked
_ Type
_ -> do
([Arg Term]
args, Type
t1') <- Int -> (Hiding -> Bool) -> Type -> m ([Arg Term], Type)
forall (m :: * -> *).
(MonadReduce m, MonadMetaSolver m, MonadDebug m, MonadTCM m) =>
Int -> (Hiding -> Bool) -> Type -> m ([Arg Term], Type)
implicitArgs Int
n Hiding -> Bool
forall a. LensHiding a => a -> Bool
notVisible Type
t1
let v' :: Term
v' = Term
v Term -> [Arg Term] -> Term
forall t. Apply t => t -> [Arg Term] -> t
`apply` [Arg Term]
args
Term
v' Term -> m () -> m Term
forall (f :: * -> *) a b. Functor f => a -> f b -> f a
<$ (Type -> Type -> m ()) -> Term -> Type -> Type -> m ()
forall (m :: * -> *).
MonadConversion m =>
(Type -> Type -> m ()) -> Term -> Type -> Type -> m ()
coerceSize (Comparison -> Type -> Type -> m ()
forall (m :: * -> *).
MonadConversion m =>
Comparison -> Type -> Type -> m ()
compareType Comparison
cmp) Term
v' Type
t1' Type
t2
where
fallback :: m Term
fallback = Term
v Term -> m () -> m Term
forall (f :: * -> *) a b. Functor f => a -> f b -> f a
<$ (Type -> Type -> m ()) -> Term -> Type -> Type -> m ()
forall (m :: * -> *).
MonadConversion m =>
(Type -> Type -> m ()) -> Term -> Type -> Type -> m ()
coerceSize (Comparison -> Type -> Type -> m ()
forall (m :: * -> *).
MonadConversion m =>
Comparison -> Type -> Type -> m ()
compareType Comparison
cmp) Term
v Type
t1 Type
t2
coerceSize :: MonadConversion m => (Type -> Type -> m ()) -> Term -> Type -> Type -> m ()
coerceSize :: (Type -> Type -> m ()) -> Term -> Type -> Type -> m ()
coerceSize Type -> Type -> m ()
leqType Term
v Type
t1 Type
t2 = VerboseKey -> Int -> VerboseKey -> m () -> m ()
forall (m :: * -> *) a.
MonadDebug m =>
VerboseKey -> Int -> VerboseKey -> m a -> m a
verboseBracket VerboseKey
"tc.conv.size.coerce" Int
45 VerboseKey
"coerceSize" (m () -> m ()) -> m () -> m ()
forall a b. (a -> b) -> a -> b
$
m () -> m ()
forall (m :: * -> *) a.
(MonadTCEnv m, HasOptions m, MonadDebug m) =>
m a -> m a
workOnTypes (m () -> m ()) -> m () -> m ()
forall a b. (a -> b) -> a -> b
$ do
VerboseKey -> Int -> TCM Doc -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCM Doc -> m ()
reportSDoc VerboseKey
"tc.conv.size.coerce" Int
70 (TCM Doc -> m ()) -> TCM Doc -> m ()
forall a b. (a -> b) -> a -> b
$
TCM Doc
"coerceSize" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> [TCM Doc] -> TCM Doc
forall (m :: * -> *). Monad m => [m Doc] -> m Doc
vcat
[ TCM Doc
"term v =" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Term -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty Term
v
, TCM Doc
"from type t1 =" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Type -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty Type
t1
, TCM Doc
"to type t2 =" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Type -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty Type
t2
]
let fallback :: m ()
fallback = Type -> Type -> m ()
leqType Type
t1 Type
t2
done :: m ()
done = m (Maybe BoundedSize) -> m () -> (BoundedSize -> m ()) -> m ()
forall (m :: * -> *) a b.
Monad m =>
m (Maybe a) -> m b -> (a -> m b) -> m b
caseMaybeM (Type -> m (Maybe BoundedSize)
forall a (m :: * -> *).
(IsSizeType a, HasOptions m, HasBuiltins m) =>
a -> m (Maybe BoundedSize)
isSizeType (Type -> m (Maybe BoundedSize)) -> m Type -> m (Maybe BoundedSize)
forall (m :: * -> *) a b. Monad m => (a -> m b) -> m a -> m b
=<< Type -> m Type
forall a (m :: * -> *). (Reduce a, MonadReduce m) => a -> m a
reduce Type
t1) m ()
fallback ((BoundedSize -> m ()) -> m ()) -> (BoundedSize -> m ()) -> m ()
forall a b. (a -> b) -> a -> b
$ \ BoundedSize
_ -> () -> m ()
forall (m :: * -> *) a. Monad m => a -> m a
return ()
m (Maybe BoundedSize) -> m () -> (BoundedSize -> m ()) -> m ()
forall (m :: * -> *) a b.
Monad m =>
m (Maybe a) -> m b -> (a -> m b) -> m b
caseMaybeM (Type -> m (Maybe BoundedSize)
forall a (m :: * -> *).
(IsSizeType a, HasOptions m, HasBuiltins m) =>
a -> m (Maybe BoundedSize)
isSizeType (Type -> m (Maybe BoundedSize)) -> m Type -> m (Maybe BoundedSize)
forall (m :: * -> *) a b. Monad m => (a -> m b) -> m a -> m b
=<< Type -> m Type
forall a (m :: * -> *). (Reduce a, MonadReduce m) => a -> m a
reduce Type
t2) m ()
fallback ((BoundedSize -> m ()) -> m ()) -> (BoundedSize -> m ()) -> m ()
forall a b. (a -> b) -> a -> b
$ \ BoundedSize
b2 -> do
SizeMaxView
mv <- Term -> m SizeMaxView
forall (m :: * -> *).
(MonadReduce m, HasBuiltins m) =>
Term -> m SizeMaxView
sizeMaxView Term
v
if (DeepSizeView -> Bool) -> SizeMaxView -> Bool
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
m Bool -> m () -> m ()
forall (m :: * -> *). Monad m => m Bool -> m () -> m ()
unlessM (m () -> m Bool
forall (m :: * -> *).
(MonadConstraint m, MonadWarning m, MonadError TCErr m,
MonadFresh ProblemId m) =>
m () -> m Bool
tryConversion (m () -> m Bool) -> m () -> m Bool
forall a b. (a -> b) -> a -> b
$ m () -> m ()
forall (m :: * -> *) a.
(MonadTCEnv m, HasOptions m, MonadDebug m) =>
m a -> m a
dontAssignMetas (m () -> m ()) -> m () -> m ()
forall a b. (a -> b) -> a -> b
$ Type -> Type -> m ()
leqType Type
t1 Type
t2) (m () -> m ()) -> m () -> m ()
forall a b. (a -> b) -> a -> b
$ do
VerboseKey -> Int -> TCM Doc -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCM Doc -> m ()
reportSDoc VerboseKey
"tc.conv.size.coerce" Int
20 (TCM Doc -> m ()) -> TCM Doc -> m ()
forall a b. (a -> b) -> a -> b
$ TCM Doc
"coercing to a size type"
case BoundedSize
b2 of
BoundedSize
BoundedNo -> m ()
done
BoundedLt Term
v2 -> do
SizeView
sv2 <- Term -> m SizeView
forall (m :: * -> *).
(HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) =>
Term -> m SizeView
sizeView Term
v2
case SizeView
sv2 of
SizeView
SizeInf -> m ()
done
OtherSize{} -> do
Term
vinc <- Int -> Term -> m Term
forall (m :: * -> *). HasBuiltins m => Int -> Term -> m Term
sizeSuc Int
1 Term
v
Comparison -> Term -> Term -> m ()
forall (m :: * -> *).
MonadConversion m =>
Comparison -> Term -> Term -> m ()
compareSizes Comparison
CmpLeq Term
vinc Term
v2
m ()
done
SizeSuc Term
a2 -> do
Comparison -> Term -> Term -> m ()
forall (m :: * -> *).
MonadConversion m =>
Comparison -> Term -> Term -> m ()
compareSizes Comparison
CmpLeq Term
v Term
a2
m ()
done
compareLevel :: MonadConversion m => Comparison -> Level -> Level -> m ()
compareLevel :: Comparison -> Level -> Level -> m ()
compareLevel Comparison
CmpLeq Level
u Level
v = Level -> Level -> m ()
forall (m :: * -> *). MonadConversion m => Level -> Level -> m ()
leqLevel Level
u Level
v
compareLevel Comparison
CmpEq Level
u Level
v = Level -> Level -> m ()
forall (m :: * -> *). MonadConversion m => Level -> Level -> m ()
equalLevel Level
u Level
v
compareSort :: MonadConversion m => Comparison -> Sort -> Sort -> m ()
compareSort :: Comparison -> Sort -> Sort -> m ()
compareSort Comparison
CmpEq = Sort -> Sort -> m ()
forall (m :: * -> *). MonadConversion m => Sort -> Sort -> m ()
equalSort
compareSort Comparison
CmpLeq = Sort -> Sort -> m ()
forall (m :: * -> *). MonadConversion m => Sort -> Sort -> m ()
leqSort
leqSort :: forall m. MonadConversion m => Sort -> Sort -> m ()
leqSort :: Sort -> Sort -> m ()
leqSort Sort
s1 Sort
s2 = (Constraint -> m () -> m ()
forall (m :: * -> *).
MonadConstraint m =>
Constraint -> m () -> m ()
catchConstraint (Comparison -> Sort -> Sort -> Constraint
SortCmp Comparison
CmpLeq Sort
s1 Sort
s2) :: m () -> m ()) (m () -> m ()) -> m () -> m ()
forall a b. (a -> b) -> a -> b
$ do
(Sort
s1,Sort
s2) <- (Sort, Sort) -> m (Sort, Sort)
forall a (m :: * -> *). (Reduce a, MonadReduce m) => a -> m a
reduce (Sort
s1,Sort
s2)
let postpone :: m ()
postpone = Constraint -> m ()
forall (m :: * -> *). MonadConstraint m => Constraint -> m ()
addConstraint (Comparison -> Sort -> Sort -> Constraint
SortCmp Comparison
CmpLeq Sort
s1 Sort
s2)
no :: m a
no = TypeError -> m a
forall (m :: * -> *) a.
(MonadTCEnv m, ReadTCState m, MonadError TCErr m) =>
TypeError -> m a
typeError (TypeError -> m a) -> TypeError -> m a
forall a b. (a -> b) -> a -> b
$ Sort -> Sort -> TypeError
NotLeqSort Sort
s1 Sort
s2
yes :: m ()
yes = () -> m ()
forall (m :: * -> *) a. Monad m => a -> m a
return ()
synEq :: m ()
synEq = m Bool -> m () -> m () -> m ()
forall (m :: * -> *) a. Monad m => m Bool -> m a -> m a -> m a
ifNotM (PragmaOptions -> Bool
optSyntacticEquality (PragmaOptions -> Bool) -> m PragmaOptions -> m Bool
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> m PragmaOptions
forall (m :: * -> *). HasOptions m => m PragmaOptions
pragmaOptions) m ()
postpone (m () -> m ()) -> m () -> m ()
forall a b. (a -> b) -> a -> b
$ do
((Sort
s1,Sort
s2) , Bool
equal) <- Sort -> Sort -> m ((Sort, Sort), Bool)
forall a (m :: * -> *).
(Instantiate a, SynEq a, MonadReduce m) =>
a -> a -> m ((a, a), Bool)
SynEq.checkSyntacticEquality Sort
s1 Sort
s2
if | Bool
equal -> m ()
yes
| Bool
otherwise -> m ()
postpone
VerboseKey -> Int -> TCM Doc -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCM Doc -> m ()
reportSDoc VerboseKey
"tc.conv.sort" Int
30 (TCM Doc -> m ()) -> TCM Doc -> m ()
forall a b. (a -> b) -> a -> b
$
[TCM Doc] -> TCM Doc
forall (m :: * -> *). Monad m => [m Doc] -> m Doc
sep [ TCM Doc
"leqSort"
, Int -> TCM Doc -> TCM Doc
forall (m :: * -> *). Functor m => Int -> m Doc -> m Doc
nest Int
2 (TCM Doc -> TCM Doc) -> TCM Doc -> TCM Doc
forall a b. (a -> b) -> a -> b
$ [TCM Doc] -> TCM Doc
forall (m :: * -> *). Monad m => [m Doc] -> m Doc
fsep [ Sort -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Sort
s1 TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> TCM Doc
"=<"
, Sort -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Sort
s2 ]
]
Bool
propEnabled <- m Bool
forall (m :: * -> *). HasOptions m => m Bool
isPropEnabled
let fvsRHS :: Int -> Bool
fvsRHS = (Int -> IntSet -> Bool
`IntSet.member` Sort -> IntSet
forall t. Free t => t -> IntSet
allFreeVars Sort
s2)
Bool
badRigid <- Sort
s1 Sort -> (Int -> Bool) -> m Bool
forall (m :: * -> *) a.
(MonadReduce m, MonadAddContext m, MonadTCEnv m, MonadDebug m,
AnyRigid a) =>
a -> (Int -> Bool) -> m Bool
`rigidVarsNotContainedIn` Int -> Bool
fvsRHS
case (Sort
s1, Sort
s2) of
(DummyS VerboseKey
s, Sort
_) -> VerboseKey -> m ()
forall (m :: * -> *) a b.
(ReportS [a], MonadDebug m, IsString a) =>
a -> m b
impossibleSort VerboseKey
s
(Sort
_, DummyS VerboseKey
s) -> VerboseKey -> m ()
forall (m :: * -> *) a b.
(ReportS [a], MonadDebug m, IsString a) =>
a -> m b
impossibleSort VerboseKey
s
(Type Level
a , Type Level
b ) -> Level -> Level -> m ()
forall (m :: * -> *). MonadConversion m => Level -> Level -> m ()
leqLevel Level
a Level
b
(Prop Level
a , Prop Level
b ) -> Level -> Level -> m ()
forall (m :: * -> *). MonadConversion m => Level -> Level -> m ()
leqLevel Level
a Level
b
(Prop Level
a , Type Level
b ) -> Level -> Level -> m ()
forall (m :: * -> *). MonadConversion m => Level -> Level -> m ()
leqLevel Level
a Level
b
(Type Level
a , Prop Level
b ) -> m ()
forall a. m a
no
(Sort
_ , Sort
Inf ) -> m ()
yes
(Sort
Inf , Sort
_ ) -> Sort -> Sort -> m ()
forall (m :: * -> *). MonadConversion m => Sort -> Sort -> m ()
equalSort Sort
s1 Sort
s2
(Sort
_ , Sort
SizeUniv) -> Sort -> Sort -> m ()
forall (m :: * -> *). MonadConversion m => Sort -> Sort -> m ()
equalSort Sort
s1 Sort
s2
(Sort
_ , Prop (Max Integer
0 [])) -> Sort -> Sort -> m ()
forall (m :: * -> *). MonadConversion m => Sort -> Sort -> m ()
equalSort Sort
s1 Sort
s2
(Sort
_ , Type (Max Integer
0 []))
| Bool -> Bool
not Bool
propEnabled -> Sort -> Sort -> m ()
forall (m :: * -> *). MonadConversion m => Sort -> Sort -> m ()
equalSort Sort
s1 Sort
s2
(Sort
SizeUniv, Type{} ) -> m ()
forall a. m a
no
(Sort
SizeUniv, Prop{} ) -> m ()
forall a. m a
no
(Sort
_ , Sort
_ ) | Bool
badRigid -> Sort -> Sort -> m ()
forall (m :: * -> *). MonadConversion m => Sort -> Sort -> m ()
equalSort Sort
s2 Sort
forall t. Sort' t
Inf
(UnivSort Sort
Inf , UnivSort Sort
Inf) -> m ()
yes
(PiSort{}, Sort
_ ) -> m ()
synEq
(Sort
_ , PiSort{}) -> m ()
synEq
(FunSort{}, Sort
_ ) -> m ()
synEq
(Sort
_ , FunSort{}) -> m ()
synEq
(UnivSort{}, Sort
_ ) -> m ()
synEq
(Sort
_ , UnivSort{}) -> m ()
synEq
(MetaS{} , Sort
_ ) -> m ()
synEq
(Sort
_ , MetaS{} ) -> m ()
synEq
(DefS{} , Sort
_ ) -> m ()
synEq
(Sort
_ , DefS{}) -> m ()
synEq
where
impossibleSort :: a -> m b
impossibleSort a
s = do
VerboseKey -> Int -> [a] -> m ()
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
]
m b
forall a. HasCallStack => a
__IMPOSSIBLE__
leqLevel :: MonadConversion m => Level -> Level -> m ()
leqLevel :: Level -> Level -> m ()
leqLevel Level
a Level
b = do
VerboseKey -> Int -> TCM Doc -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCM Doc -> m ()
reportSDoc VerboseKey
"tc.conv.nat" Int
30 (TCM Doc -> m ()) -> TCM Doc -> m ()
forall a b. (a -> b) -> a -> b
$
TCM Doc
"compareLevel" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+>
[TCM Doc] -> TCM Doc
forall (m :: * -> *). Monad m => [m Doc] -> m Doc
sep [ Level -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Level
a TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> TCM Doc
"=<"
, Level -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Level
b ]
Level
a <- Level -> m Level
forall a (m :: * -> *). (Normalise a, MonadReduce m) => a -> m a
normalise Level
a
Level
b <- Level -> m Level
forall a (m :: * -> *). (Normalise a, MonadReduce m) => a -> m a
normalise Level
b
Level -> Level -> m ()
forall (m :: * -> *). MonadConversion m => Level -> Level -> m ()
leqView Level
a Level
b
where
leqView :: MonadConversion m => Level -> Level -> m ()
leqView :: Level -> Level -> m ()
leqView Level
a Level
b = Constraint -> m () -> m ()
forall (m :: * -> *).
MonadConstraint m =>
Constraint -> m () -> m ()
catchConstraint (Comparison -> Level -> Level -> Constraint
LevelCmp Comparison
CmpLeq Level
a Level
b) (m () -> m ()) -> m () -> m ()
forall a b. (a -> b) -> a -> b
$ do
VerboseKey -> Int -> TCM Doc -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCM Doc -> m ()
reportSDoc VerboseKey
"tc.conv.level" Int
30 (TCM Doc -> m ()) -> TCM Doc -> m ()
forall a b. (a -> b) -> a -> b
$
TCM Doc
"compareLevelView" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+>
[TCM Doc] -> TCM Doc
forall (m :: * -> *). Monad m => [m Doc] -> m Doc
sep [ Level -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty Level
a TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> TCM Doc
"=<"
, Level -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty Level
b ]
Bool
cumulativity <- PragmaOptions -> Bool
optCumulativity (PragmaOptions -> Bool) -> m PragmaOptions -> m Bool
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> m PragmaOptions
forall (m :: * -> *). HasOptions m => m PragmaOptions
pragmaOptions
VerboseKey -> Int -> TCM Doc -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCM Doc -> m ()
reportSDoc VerboseKey
"tc.conv.level" Int
40 (TCM Doc -> m ()) -> TCM Doc -> m ()
forall a b. (a -> b) -> a -> b
$
TCM Doc
"compareLevelView" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+>
[TCM Doc] -> TCM Doc
forall (m :: * -> *). Monad m => [m Doc] -> m Doc
sep [ [TCM Doc] -> TCM Doc
forall (m :: * -> *).
(Monad m, Semigroup (m Doc)) =>
[m Doc] -> m Doc
prettyList_ ((SingleLevel -> TCM Doc) -> [SingleLevel] -> [TCM Doc]
forall a b. (a -> b) -> [a] -> [b]
map (Level -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty (Level -> TCM Doc)
-> (SingleLevel -> Level) -> SingleLevel -> TCM Doc
forall b c a. (b -> c) -> (a -> b) -> a -> c
. SingleLevel -> Level
unSingleLevel) ([SingleLevel] -> [TCM Doc]) -> [SingleLevel] -> [TCM Doc]
forall a b. (a -> b) -> a -> b
$ NonEmpty SingleLevel -> [SingleLevel]
forall a. NonEmpty a -> [a]
NonEmpty.toList (NonEmpty SingleLevel -> [SingleLevel])
-> NonEmpty SingleLevel -> [SingleLevel]
forall a b. (a -> b) -> a -> b
$ Level -> NonEmpty SingleLevel
levelMaxView Level
a)
, TCM Doc
"=<"
, [TCM Doc] -> TCM Doc
forall (m :: * -> *).
(Monad m, Semigroup (m Doc)) =>
[m Doc] -> m Doc
prettyList_ ((SingleLevel -> TCM Doc) -> [SingleLevel] -> [TCM Doc]
forall a b. (a -> b) -> [a] -> [b]
map (Level -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty (Level -> TCM Doc)
-> (SingleLevel -> Level) -> SingleLevel -> TCM Doc
forall b c a. (b -> c) -> (a -> b) -> a -> c
. SingleLevel -> Level
unSingleLevel) ([SingleLevel] -> [TCM Doc]) -> [SingleLevel] -> [TCM Doc]
forall a b. (a -> b) -> a -> b
$ NonEmpty SingleLevel -> [SingleLevel]
forall a. NonEmpty a -> [a]
NonEmpty.toList (NonEmpty SingleLevel -> [SingleLevel])
-> NonEmpty SingleLevel -> [SingleLevel]
forall a b. (a -> b) -> a -> b
$ Level -> NonEmpty SingleLevel
levelMaxView Level
b)
]
m () -> m ()
wrap (m () -> m ()) -> m () -> m ()
forall a b. (a -> b) -> a -> b
$ case (Level -> NonEmpty SingleLevel
levelMaxView Level
a, Level -> NonEmpty SingleLevel
levelMaxView Level
b) of
(NonEmpty SingleLevel, NonEmpty SingleLevel)
_ | Level
a Level -> Level -> Bool
forall a. Eq a => a -> a -> Bool
== Level
b -> m ()
ok
(SingleClosed Integer
0 :| [] , NonEmpty SingleLevel
_) -> m ()
ok
(NonEmpty SingleLevel
as , SingleClosed Integer
0 :| []) ->
[m ()] -> m ()
forall (t :: * -> *) (m :: * -> *) a.
(Foldable t, Monad m) =>
t (m a) -> m ()
sequence_ [ Level -> Level -> m ()
forall (m :: * -> *). MonadConversion m => Level -> Level -> m ()
equalLevel (SingleLevel -> Level
unSingleLevel SingleLevel
a') (Integer -> Level
ClosedLevel Integer
0) | SingleLevel
a' <- NonEmpty SingleLevel -> [SingleLevel]
forall a. NonEmpty a -> [a]
NonEmpty.toList NonEmpty SingleLevel
as ]
(SingleClosed Integer
m :| [], SingleClosed Integer
n :| []) -> if Integer
m Integer -> Integer -> Bool
forall a. Ord a => a -> a -> Bool
<= Integer
n then m ()
ok else m ()
notok
(SingleClosed Integer
m :| [] , NonEmpty SingleLevel
_)
| Integer
m Integer -> Integer -> Bool
forall a. Ord a => a -> a -> Bool
<= Level -> Integer
levelLowerBound Level
b -> m ()
ok
(NonEmpty SingleLevel
as, NonEmpty SingleLevel
bs)
| (SingleLevel -> Bool) -> NonEmpty SingleLevel -> Bool
forall (t :: * -> *) a. Foldable t => (a -> Bool) -> t a -> Bool
all SingleLevel -> Bool
neutralOrClosed NonEmpty SingleLevel
bs , Level -> Integer
levelLowerBound Level
a Integer -> Integer -> Bool
forall a. Ord a => a -> a -> Bool
> Level -> Integer
levelLowerBound Level
b -> m ()
notok
(as :: NonEmpty SingleLevel
as@(SingleLevel
_:|SingleLevel
_:[SingleLevel]
_), SingleLevel
b :| []) ->
[m ()] -> m ()
forall (t :: * -> *) (m :: * -> *) a.
(Foldable t, Monad m) =>
t (m a) -> m ()
sequence_ [ Level -> Level -> m ()
forall (m :: * -> *). MonadConversion m => Level -> Level -> m ()
leqView (SingleLevel -> Level
unSingleLevel SingleLevel
a') (SingleLevel -> Level
unSingleLevel SingleLevel
b) | SingleLevel
a' <- NonEmpty SingleLevel -> [SingleLevel]
forall a. NonEmpty a -> [a]
NonEmpty.toList NonEmpty SingleLevel
as ]
(NonEmpty SingleLevel
as, NonEmpty SingleLevel
bs)
| let minN :: Integer
minN = Integer -> Integer -> Integer
forall a. Ord a => a -> a -> a
min ((Integer, Level) -> Integer
forall a b. (a, b) -> a
fst ((Integer, Level) -> Integer) -> (Integer, Level) -> Integer
forall a b. (a -> b) -> a -> b
$ Level -> (Integer, Level)
levelPlusView Level
a) ((Integer, Level) -> Integer
forall a b. (a, b) -> a
fst ((Integer, Level) -> Integer) -> (Integer, Level) -> Integer
forall a b. (a -> b) -> a -> b
$ Level -> (Integer, Level)
levelPlusView Level
b)
a' :: Level
a' = Level -> Maybe Level -> Level
forall a. a -> Maybe a -> a
fromMaybe Level
forall a. HasCallStack => a
__IMPOSSIBLE__ (Maybe Level -> Level) -> Maybe Level -> Level
forall a b. (a -> b) -> a -> b
$ Integer -> Level -> Maybe Level
subLevel Integer
minN Level
a
b' :: Level
b' = Level -> Maybe Level -> Level
forall a. a -> Maybe a -> a
fromMaybe Level
forall a. HasCallStack => a
__IMPOSSIBLE__ (Maybe Level -> Level) -> Maybe Level -> Level
forall a b. (a -> b) -> a -> b
$ Integer -> Level -> Maybe Level
subLevel Integer
minN Level
b
, Integer
minN Integer -> Integer -> Bool
forall a. Ord a => a -> a -> Bool
> Integer
0 -> Level -> Level -> m ()
forall (m :: * -> *). MonadConversion m => Level -> Level -> m ()
leqView Level
a' Level
b'
(NonEmpty SingleLevel
as, NonEmpty SingleLevel
bs)
| (subsumed :: [SingleLevel]
subsumed@(SingleLevel
_:[SingleLevel]
_) , [SingleLevel]
as') <- (SingleLevel -> Bool)
-> [SingleLevel] -> ([SingleLevel], [SingleLevel])
forall a. (a -> Bool) -> [a] -> ([a], [a])
List.partition SingleLevel -> Bool
isSubsumed (NonEmpty SingleLevel -> [SingleLevel]
forall a. NonEmpty a -> [a]
NonEmpty.toList NonEmpty SingleLevel
as)
-> Level -> Level -> m ()
forall (m :: * -> *). MonadConversion m => Level -> Level -> m ()
leqView ([SingleLevel] -> Level
unSingleLevels [SingleLevel]
as') Level
b
where
isSubsumed :: SingleLevel -> Bool
isSubsumed SingleLevel
a = (SingleLevel -> Bool) -> [SingleLevel] -> Bool
forall (t :: * -> *) a. Foldable t => (a -> Bool) -> t a -> Bool
any (SingleLevel -> SingleLevel -> Bool
`subsumes` SingleLevel
a) (NonEmpty SingleLevel -> [SingleLevel]
forall a. NonEmpty a -> [a]
NonEmpty.toList NonEmpty SingleLevel
bs)
subsumes :: SingleLevel -> SingleLevel -> Bool
subsumes :: SingleLevel -> SingleLevel -> Bool
subsumes (SingleClosed Integer
m) (SingleClosed Integer
n) = Integer
m Integer -> Integer -> Bool
forall a. Ord a => a -> a -> Bool
>= Integer
n
subsumes (SinglePlus (Plus Integer
m LevelAtom' Term
_)) (SingleClosed Integer
n) = Integer
m Integer -> Integer -> Bool
forall a. Ord a => a -> a -> Bool
>= Integer
n
subsumes (SinglePlus (Plus Integer
m LevelAtom' Term
a)) (SinglePlus (Plus Integer
n LevelAtom' Term
b)) = LevelAtom' Term
a LevelAtom' Term -> LevelAtom' Term -> Bool
forall a. Eq a => a -> a -> Bool
== LevelAtom' Term
b Bool -> Bool -> Bool
&& Integer
m Integer -> Integer -> Bool
forall a. Ord a => a -> a -> Bool
>= Integer
n
subsumes SingleLevel
_ SingleLevel
_ = Bool
False
(NonEmpty SingleLevel
as , NonEmpty SingleLevel
bs)
| Bool
cumulativity
, Just (mb :: LevelAtom' Term
mb@(MetaLevel MetaId
x Elims
es) , [SingleLevel]
bs') <- [SingleLevel] -> Maybe (LevelAtom' Term, [SingleLevel])
singleMetaView (NonEmpty SingleLevel -> [SingleLevel]
forall a. NonEmpty a -> [a]
NonEmpty.toList NonEmpty SingleLevel
bs)
, [SingleLevel] -> Bool
forall (t :: * -> *) a. Foldable t => t a -> Bool
null [SingleLevel]
bs' Bool -> Bool -> Bool
|| (Term, Level) -> Bool
forall a. TermLike a => a -> Bool
noMetas (Level -> Term
Level Level
a , [SingleLevel] -> Level
unSingleLevels [SingleLevel]
bs') -> do
MetaVariable
mv <- MetaId -> m MetaVariable
forall (m :: * -> *).
(MonadFail m, ReadTCState m) =>
MetaId -> m MetaVariable
lookupMeta MetaId
x
Bool
abort <- (Maybe InteractionId -> Bool
forall a. Maybe a -> Bool
isJust (Maybe InteractionId -> Bool) -> m (Maybe InteractionId) -> m Bool
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> MetaId -> m (Maybe InteractionId)
forall (m :: * -> *).
ReadTCState m =>
MetaId -> m (Maybe InteractionId)
isInteractionMeta MetaId
x) m Bool -> m Bool -> m Bool
forall (m :: * -> *). Monad m => m Bool -> m Bool -> m Bool
`or2M`
((DoGeneralize -> DoGeneralize -> Bool
forall a. Eq a => a -> a -> Bool
== DoGeneralize
YesGeneralize) (DoGeneralize -> Bool) -> m DoGeneralize -> m Bool
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> MetaId -> m DoGeneralize
forall (m :: * -> *).
(ReadTCState m, MonadFail m) =>
MetaId -> m DoGeneralize
isGeneralizableMeta MetaId
x)
if | Bool
abort -> m ()
forall a. m a
postpone
| Bool
otherwise -> do
MetaId
x' <- case MetaVariable -> Judgement MetaId
mvJudgement MetaVariable
mv of
IsSort{} -> m MetaId
forall a. HasCallStack => a
__IMPOSSIBLE__
HasType MetaId
_ Comparison
cmp Type
t -> do
TelV Telescope
tel Type
t' <- Type -> m (TelV Type)
forall (m :: * -> *).
(MonadReduce m, MonadAddContext m) =>
Type -> m (TelV Type)
telView Type
t
Frozen
-> MetaInfo
-> MetaPriority
-> Permutation
-> Judgement ()
-> m MetaId
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 (Int -> Permutation) -> Int -> Permutation
forall a b. (a -> b) -> a -> b
$ Telescope -> Int
forall a. Sized a => a -> Int
size Telescope
tel) (Judgement () -> m MetaId) -> Judgement () -> m MetaId
forall a b. (a -> b) -> a -> b
$ () -> Comparison -> Type -> Judgement ()
forall a. a -> Comparison -> Type -> Judgement a
HasType () Comparison
cmp Type
t
VerboseKey -> Int -> TCM Doc -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCM Doc -> m ()
reportSDoc VerboseKey
"tc.conv.level" Int
20 (TCM Doc -> m ()) -> TCM Doc -> m ()
forall a b. (a -> b) -> a -> b
$ [TCM Doc] -> TCM Doc
forall (m :: * -> *). Monad m => [m Doc] -> m Doc
fsep
[ TCM Doc
"attempting to solve" , Term -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM (MetaId -> Elims -> Term
MetaV MetaId
x Elims
es) , TCM Doc
"to the maximum of"
, Term -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM (Level -> Term
Level Level
a) , TCM Doc
"and the fresh meta" , Term -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM (MetaId -> Elims -> Term
MetaV MetaId
x' Elims
es)
]
Level -> Level -> m ()
forall (m :: * -> *). MonadConversion m => Level -> Level -> m ()
equalLevel (LevelAtom' Term -> Level
atomicLevel LevelAtom' Term
mb) (Level -> m ()) -> Level -> m ()
forall a b. (a -> b) -> a -> b
$ Level -> Level -> Level
levelLub Level
a (LevelAtom' Term -> Level
atomicLevel (LevelAtom' Term -> Level) -> LevelAtom' Term -> Level
forall a b. (a -> b) -> a -> b
$ MetaId -> Elims -> LevelAtom' Term
forall t. MetaId -> [Elim' t] -> LevelAtom' t
MetaLevel MetaId
x' Elims
es)
(NonEmpty SingleLevel, NonEmpty SingleLevel)
_ | (Term, Term) -> Bool
forall a. TermLike a => a -> Bool
noMetas (Level -> Term
Level Level
a , Level -> Term
Level Level
b) -> m ()
notok
| Bool
otherwise -> m ()
forall a. m a
postpone
where
ok :: m ()
ok = () -> m ()
forall (m :: * -> *) a. Monad m => a -> m a
return ()
notok :: m ()
notok = m Bool -> m () -> m ()
forall (m :: * -> *). Monad m => m Bool -> m () -> m ()
unlessM m Bool
forall (m :: * -> *). HasOptions m => m Bool
typeInType (m () -> m ()) -> m () -> m ()
forall a b. (a -> b) -> a -> b
$ TypeError -> m ()
forall (m :: * -> *) a.
(MonadTCEnv m, ReadTCState m, MonadError TCErr m) =>
TypeError -> m a
typeError (TypeError -> m ()) -> TypeError -> m ()
forall a b. (a -> b) -> a -> b
$ Sort -> Sort -> TypeError
NotLeqSort (Level -> Sort
forall t. Level' t -> Sort' t
Type Level
a) (Level -> Sort
forall t. Level' t -> Sort' t
Type Level
b)
postpone :: m a
postpone = m a
forall (m :: * -> *) a. MonadError TCErr m => m a
patternViolation
wrap :: m () -> m ()
wrap m ()
m = m ()
m m () -> (TCErr -> m ()) -> m ()
forall e (m :: * -> *) a.
MonadError e m =>
m a -> (e -> m a) -> m a
`catchError` \case
TypeError{} -> m ()
notok
TCErr
err -> TCErr -> m ()
forall e (m :: * -> *) a. MonadError e m => e -> m a
throwError TCErr
err
neutralOrClosed :: SingleLevel -> Bool
neutralOrClosed (SingleClosed Integer
_) = Bool
True
neutralOrClosed (SinglePlus (Plus Integer
_ NeutralLevel{})) = Bool
True
neutralOrClosed SingleLevel
_ = Bool
False
singleMetaView :: [SingleLevel] -> Maybe (LevelAtom, [SingleLevel])
singleMetaView :: [SingleLevel] -> Maybe (LevelAtom' Term, [SingleLevel])
singleMetaView (SinglePlus (Plus Integer
0 l :: LevelAtom' Term
l@(MetaLevel MetaId
m Elims
es)) : [SingleLevel]
ls)
| (SingleLevel -> Bool) -> [SingleLevel] -> Bool
forall (t :: * -> *) a. Foldable t => (a -> Bool) -> t a -> Bool
all (Bool -> Bool
not (Bool -> Bool) -> (SingleLevel -> Bool) -> SingleLevel -> Bool
forall b c a. (b -> c) -> (a -> b) -> a -> c
. SingleLevel -> Bool
isMetaLevel) [SingleLevel]
ls = (LevelAtom' Term, [SingleLevel])
-> Maybe (LevelAtom' Term, [SingleLevel])
forall a. a -> Maybe a
Just (LevelAtom' Term
l,[SingleLevel]
ls)
singleMetaView (SingleLevel
l : [SingleLevel]
ls)
| Bool -> Bool
not (Bool -> Bool) -> Bool -> Bool
forall a b. (a -> b) -> a -> b
$ SingleLevel -> Bool
isMetaLevel SingleLevel
l = ([SingleLevel] -> [SingleLevel])
-> (LevelAtom' Term, [SingleLevel])
-> (LevelAtom' Term, [SingleLevel])
forall (a :: * -> * -> *) b c d.
Arrow a =>
a b c -> a (d, b) (d, c)
second (SingleLevel
lSingleLevel -> [SingleLevel] -> [SingleLevel]
forall a. a -> [a] -> [a]
:) ((LevelAtom' Term, [SingleLevel])
-> (LevelAtom' Term, [SingleLevel]))
-> Maybe (LevelAtom' Term, [SingleLevel])
-> Maybe (LevelAtom' Term, [SingleLevel])
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> [SingleLevel] -> Maybe (LevelAtom' Term, [SingleLevel])
singleMetaView [SingleLevel]
ls
singleMetaView [SingleLevel]
_ = Maybe (LevelAtom' Term, [SingleLevel])
forall a. Maybe a
Nothing
isMetaLevel :: SingleLevel -> Bool
isMetaLevel :: SingleLevel -> Bool
isMetaLevel (SinglePlus (Plus Integer
_ MetaLevel{})) = Bool
True
isMetaLevel (SinglePlus (Plus Integer
_ UnreducedLevel{})) = Bool
forall a. HasCallStack => a
__IMPOSSIBLE__
isMetaLevel SingleLevel
_ = Bool
False
equalLevel :: MonadConversion m => Level -> Level -> m ()
equalLevel :: Level -> Level -> m ()
equalLevel Level
a Level
b = do
(Level
a, Level
b) <- (Level, Level) -> m (Level, Level)
forall a (m :: * -> *). (Normalise a, MonadReduce m) => a -> m a
normalise (Level
a, Level
b)
Level -> Level -> m ()
forall (m :: * -> *). MonadConversion m => Level -> Level -> m ()
equalLevel' Level
a Level
b
equalLevel' :: forall m. MonadConversion m => Level -> Level -> m ()
equalLevel' :: Level -> Level -> m ()
equalLevel' Level
a Level
b = do
VerboseKey -> Int -> TCM Doc -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCM Doc -> m ()
reportSDoc VerboseKey
"tc.conv.level" Int
50 (TCM Doc -> m ()) -> TCM Doc -> m ()
forall a b. (a -> b) -> a -> b
$ [TCM Doc] -> TCM Doc
forall (m :: * -> *). Monad m => [m Doc] -> m Doc
sep [ TCM Doc
"equalLevel", Int -> TCM Doc -> TCM Doc
forall (m :: * -> *). Functor m => Int -> m Doc -> m Doc
nest Int
2 (TCM Doc -> TCM Doc) -> TCM Doc -> TCM Doc
forall a b. (a -> b) -> a -> b
$ TCM Doc -> TCM Doc
forall (m :: * -> *). Functor m => m Doc -> m Doc
parens (TCM Doc -> TCM Doc) -> TCM Doc -> TCM Doc
forall a b. (a -> b) -> a -> b
$ Level -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty Level
a, Int -> TCM Doc -> TCM Doc
forall (m :: * -> *). Functor m => Int -> m Doc -> m Doc
nest Int
2 (TCM Doc -> TCM Doc) -> TCM Doc -> TCM Doc
forall a b. (a -> b) -> a -> b
$ TCM Doc -> TCM Doc
forall (m :: * -> *). Functor m => m Doc -> m Doc
parens (TCM Doc -> TCM Doc) -> TCM Doc -> TCM Doc
forall a b. (a -> b) -> a -> b
$ Level -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty Level
b ]
VerboseKey -> Int -> TCM Doc -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCM Doc -> m ()
reportSDoc VerboseKey
"tc.conv.level" Int
40 (TCM Doc -> m ()) -> TCM Doc -> m ()
forall a b. (a -> b) -> a -> b
$
[TCM Doc] -> TCM Doc
forall (m :: * -> *). Monad m => [m Doc] -> m Doc
sep [ TCM Doc
"equalLevel"
, [TCM Doc] -> TCM Doc
forall (m :: * -> *). Monad m => [m Doc] -> m Doc
vcat [ Int -> TCM Doc -> TCM Doc
forall (m :: * -> *). Functor m => Int -> m Doc -> m Doc
nest Int
2 (TCM Doc -> TCM Doc) -> TCM Doc -> TCM Doc
forall a b. (a -> b) -> a -> b
$ [TCM Doc] -> TCM Doc
forall (m :: * -> *). Monad m => [m Doc] -> m Doc
sep [ Level -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Level
a TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> TCM Doc
"=="
, Level -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Level
b
]
]
]
let (Level
a',Level
b') = Level -> Level -> (Level, Level)
removeSubsumed Level
a Level
b
VerboseKey -> Int -> TCM Doc -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCM Doc -> m ()
reportSDoc VerboseKey
"tc.conv.level" Int
50 (TCM Doc -> m ()) -> TCM Doc -> m ()
forall a b. (a -> b) -> a -> b
$
[TCM Doc] -> TCM Doc
forall (m :: * -> *). Monad m => [m Doc] -> m Doc
sep [ TCM Doc
"equalLevel (w/o subsumed)"
, [TCM Doc] -> TCM Doc
forall (m :: * -> *). Monad m => [m Doc] -> m Doc
vcat [ Int -> TCM Doc -> TCM Doc
forall (m :: * -> *). Functor m => Int -> m Doc -> m Doc
nest Int
2 (TCM Doc -> TCM Doc) -> TCM Doc -> TCM Doc
forall a b. (a -> b) -> a -> b
$ [TCM Doc] -> TCM Doc
forall (m :: * -> *). Monad m => [m Doc] -> m Doc
sep [ Level -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Level
a' TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> TCM Doc
"=="
, Level -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Level
b'
]
]
]
let as :: NonEmpty SingleLevel
as = Level -> NonEmpty SingleLevel
levelMaxView Level
a'
bs :: NonEmpty SingleLevel
bs = Level -> NonEmpty SingleLevel
levelMaxView Level
b'
VerboseKey -> Int -> TCM Doc -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCM Doc -> m ()
reportSDoc VerboseKey
"tc.conv.level" Int
50 (TCM Doc -> m ()) -> TCM Doc -> m ()
forall a b. (a -> b) -> a -> b
$
[TCM Doc] -> TCM Doc
forall (m :: * -> *). Monad m => [m Doc] -> m Doc
sep [ VerboseKey -> TCM Doc
forall (m :: * -> *). Monad m => VerboseKey -> m Doc
text VerboseKey
"equalLevel"
, [TCM Doc] -> TCM Doc
forall (m :: * -> *). Monad m => [m Doc] -> m Doc
vcat [ Int -> TCM Doc -> TCM Doc
forall (m :: * -> *). Functor m => Int -> m Doc -> m Doc
nest Int
2 (TCM Doc -> TCM Doc) -> TCM Doc -> TCM Doc
forall a b. (a -> b) -> a -> b
$ [TCM Doc] -> TCM Doc
forall (m :: * -> *). Monad m => [m Doc] -> m Doc
sep [ [TCM Doc] -> TCM Doc
forall (m :: * -> *).
(Monad m, Semigroup (m Doc)) =>
[m Doc] -> m Doc
prettyList_ ((SingleLevel -> TCM Doc) -> [SingleLevel] -> [TCM Doc]
forall a b. (a -> b) -> [a] -> [b]
map (Level -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty (Level -> TCM Doc)
-> (SingleLevel -> Level) -> SingleLevel -> TCM Doc
forall b c a. (b -> c) -> (a -> b) -> a -> c
. SingleLevel -> Level
unSingleLevel) ([SingleLevel] -> [TCM Doc]) -> [SingleLevel] -> [TCM Doc]
forall a b. (a -> b) -> a -> b
$ NonEmpty SingleLevel -> [SingleLevel]
forall a. NonEmpty a -> [a]
NonEmpty.toList (NonEmpty SingleLevel -> [SingleLevel])
-> NonEmpty SingleLevel -> [SingleLevel]
forall a b. (a -> b) -> a -> b
$ NonEmpty SingleLevel
as)
, TCM Doc
"=="
, [TCM Doc] -> TCM Doc
forall (m :: * -> *).
(Monad m, Semigroup (m Doc)) =>
[m Doc] -> m Doc
prettyList_ ((SingleLevel -> TCM Doc) -> [SingleLevel] -> [TCM Doc]
forall a b. (a -> b) -> [a] -> [b]
map (Level -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty (Level -> TCM Doc)
-> (SingleLevel -> Level) -> SingleLevel -> TCM Doc
forall b c a. (b -> c) -> (a -> b) -> a -> c
. SingleLevel -> Level
unSingleLevel) ([SingleLevel] -> [TCM Doc]) -> [SingleLevel] -> [TCM Doc]
forall a b. (a -> b) -> a -> b
$ NonEmpty SingleLevel -> [SingleLevel]
forall a. NonEmpty a -> [a]
NonEmpty.toList (NonEmpty SingleLevel -> [SingleLevel])
-> NonEmpty SingleLevel -> [SingleLevel]
forall a b. (a -> b) -> a -> b
$ NonEmpty SingleLevel
bs)
]
]
]
VerboseKey -> Int -> TCM Doc -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCM Doc -> m ()
reportSDoc VerboseKey
"tc.conv.level" Int
80 (TCM Doc -> m ()) -> TCM Doc -> m ()
forall a b. (a -> b) -> a -> b
$
[TCM Doc] -> TCM Doc
forall (m :: * -> *). Monad m => [m Doc] -> m Doc
sep [ VerboseKey -> TCM Doc
forall (m :: * -> *). Monad m => VerboseKey -> m Doc
text VerboseKey
"equalLevel"
, [TCM Doc] -> TCM Doc
forall (m :: * -> *). Monad m => [m Doc] -> m Doc
vcat [ Int -> TCM Doc -> TCM Doc
forall (m :: * -> *). Functor m => Int -> m Doc -> m Doc
nest Int
2 (TCM Doc -> TCM Doc) -> TCM Doc -> TCM Doc
forall a b. (a -> b) -> a -> b
$ [TCM Doc] -> TCM Doc
forall (m :: * -> *). Monad m => [m Doc] -> m Doc
sep [ [TCM Doc] -> TCM Doc
forall (m :: * -> *).
(Monad m, Semigroup (m Doc)) =>
[m Doc] -> m Doc
prettyList_ ((SingleLevel -> TCM Doc) -> [SingleLevel] -> [TCM Doc]
forall a b. (a -> b) -> [a] -> [b]
map (VerboseKey -> TCM Doc
forall (m :: * -> *). Monad m => VerboseKey -> m Doc
text (VerboseKey -> TCM Doc)
-> (SingleLevel -> VerboseKey) -> SingleLevel -> TCM Doc
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Level -> VerboseKey
forall a. Show a => a -> VerboseKey
show (Level -> VerboseKey)
-> (SingleLevel -> Level) -> SingleLevel -> VerboseKey
forall b c a. (b -> c) -> (a -> b) -> a -> c
. SingleLevel -> Level
unSingleLevel) ([SingleLevel] -> [TCM Doc]) -> [SingleLevel] -> [TCM Doc]
forall a b. (a -> b) -> a -> b
$ NonEmpty SingleLevel -> [SingleLevel]
forall a. NonEmpty a -> [a]
NonEmpty.toList (NonEmpty SingleLevel -> [SingleLevel])
-> NonEmpty SingleLevel -> [SingleLevel]
forall a b. (a -> b) -> a -> b
$ NonEmpty SingleLevel
as)
, TCM Doc
"=="
, [TCM Doc] -> TCM Doc
forall (m :: * -> *).
(Monad m, Semigroup (m Doc)) =>
[m Doc] -> m Doc
prettyList_ ((SingleLevel -> TCM Doc) -> [SingleLevel] -> [TCM Doc]
forall a b. (a -> b) -> [a] -> [b]
map (VerboseKey -> TCM Doc
forall (m :: * -> *). Monad m => VerboseKey -> m Doc
text (VerboseKey -> TCM Doc)
-> (SingleLevel -> VerboseKey) -> SingleLevel -> TCM Doc
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Level -> VerboseKey
forall a. Show a => a -> VerboseKey
show (Level -> VerboseKey)
-> (SingleLevel -> Level) -> SingleLevel -> VerboseKey
forall b c a. (b -> c) -> (a -> b) -> a -> c
. SingleLevel -> Level
unSingleLevel) ([SingleLevel] -> [TCM Doc]) -> [SingleLevel] -> [TCM Doc]
forall a b. (a -> b) -> a -> b
$ NonEmpty SingleLevel -> [SingleLevel]
forall a. NonEmpty a -> [a]
NonEmpty.toList (NonEmpty SingleLevel -> [SingleLevel])
-> NonEmpty SingleLevel -> [SingleLevel]
forall a b. (a -> b) -> a -> b
$ NonEmpty SingleLevel
bs)
]
]
]
Constraint -> m () -> m ()
forall (m :: * -> *).
MonadConstraint m =>
Constraint -> m () -> m ()
catchConstraint (Comparison -> Level -> Level -> Constraint
LevelCmp Comparison
CmpEq Level
a Level
b) (m () -> m ()) -> m () -> m ()
forall a b. (a -> b) -> a -> b
$ case (NonEmpty SingleLevel
as, NonEmpty SingleLevel
bs) of
(NonEmpty SingleLevel, NonEmpty SingleLevel)
_ | Level
a Level -> Level -> Bool
forall a. Eq a => a -> a -> Bool
== Level
b -> m ()
ok
(SingleClosed Integer
m :| [], SingleClosed Integer
n :| [])
| Integer
m Integer -> Integer -> Bool
forall a. Eq a => a -> a -> Bool
== Integer
n -> m ()
ok
| Bool
otherwise -> m ()
notok
(SingleClosed Integer
m :| [] , NonEmpty SingleLevel
bs) | (SingleLevel -> Bool) -> NonEmpty SingleLevel -> Bool
forall (t :: * -> *) a. Foldable t => (a -> Bool) -> t a -> Bool
any SingleLevel -> Bool
isNeutral NonEmpty SingleLevel
bs -> m ()
notok
(NonEmpty SingleLevel
as , SingleClosed Integer
n :| []) | (SingleLevel -> Bool) -> NonEmpty SingleLevel -> Bool
forall (t :: * -> *) a. Foldable t => (a -> Bool) -> t a -> Bool
any SingleLevel -> Bool
isNeutral NonEmpty SingleLevel
as -> m ()
notok
(SingleClosed Integer
m :| [] , NonEmpty SingleLevel
_) | Integer
m Integer -> Integer -> Bool
forall a. Ord a => a -> a -> Bool
< Level -> Integer
levelLowerBound Level
b -> m ()
notok
(NonEmpty SingleLevel
_ , SingleClosed Integer
n :| []) | Integer
n Integer -> Integer -> Bool
forall a. Ord a => a -> a -> Bool
< Level -> Integer
levelLowerBound Level
a -> m ()
notok
(SingleClosed Integer
0 :| [] , bs :: NonEmpty SingleLevel
bs@(SingleLevel
_:|SingleLevel
_:[SingleLevel]
_)) ->
[m ()] -> m ()
forall (t :: * -> *) (m :: * -> *) a.
(Foldable t, Monad m) =>
t (m a) -> m ()
sequence_ [ Level -> Level -> m ()
forall (m :: * -> *). MonadConversion m => Level -> Level -> m ()
equalLevel' (Integer -> Level
ClosedLevel Integer
0) (SingleLevel -> Level
unSingleLevel SingleLevel
b') | SingleLevel
b' <- NonEmpty SingleLevel -> [SingleLevel]
forall a. NonEmpty a -> [a]
NonEmpty.toList NonEmpty SingleLevel
bs ]
(as :: NonEmpty SingleLevel
as@(SingleLevel
_:|SingleLevel
_:[SingleLevel]
_) , SingleClosed Integer
0 :| []) ->
[m ()] -> m ()
forall (t :: * -> *) (m :: * -> *) a.
(Foldable t, Monad m) =>
t (m a) -> m ()
sequence_ [ Level -> Level -> m ()
forall (m :: * -> *). MonadConversion m => Level -> Level -> m ()
equalLevel' (SingleLevel -> Level
unSingleLevel SingleLevel
a') (Integer -> Level
ClosedLevel Integer
0) | SingleLevel
a' <- NonEmpty SingleLevel -> [SingleLevel]
forall a. NonEmpty a -> [a]
NonEmpty.toList NonEmpty SingleLevel
as ]
(SinglePlus (Plus Integer
k (MetaLevel MetaId
x Elims
as)) :| [] , NonEmpty SingleLevel
bs)
| (SingleLevel -> Bool) -> NonEmpty SingleLevel -> Bool
forall (t :: * -> *) a. Foldable t => (a -> Bool) -> t a -> Bool
any (MetaId -> SingleLevel -> Bool
isThisMeta MetaId
x) NonEmpty SingleLevel
bs -> m ()
forall b. m b
postpone
(NonEmpty SingleLevel
as , SinglePlus (Plus Integer
k (MetaLevel MetaId
x Elims
bs)) :| [])
| (SingleLevel -> Bool) -> NonEmpty SingleLevel -> Bool
forall (t :: * -> *) a. Foldable t => (a -> Bool) -> t a -> Bool
any (MetaId -> SingleLevel -> Bool
isThisMeta MetaId
x) NonEmpty SingleLevel
as -> m ()
forall b. m b
postpone
(SinglePlus (Plus Integer
k (MetaLevel MetaId
x Elims
as')) :| [] , SinglePlus (Plus Integer
l (MetaLevel MetaId
y Elims
bs')) :| [])
| Integer
k Integer -> Integer -> Bool
forall a. Eq a => a -> a -> Bool
== Integer
l -> if
| MetaId
y MetaId -> MetaId -> Bool
forall a. Ord a => a -> a -> Bool
< MetaId
x -> MetaId -> Elims -> Level -> m ()
forall (m :: * -> *).
(MonadMetaSolver m, MonadWarning m, MonadStatistics m,
MonadFresh ProblemId m, MonadFresh Int m) =>
MetaId -> Elims -> Level -> m ()
meta MetaId
x Elims
as' (Level -> m ()) -> Level -> m ()
forall a b. (a -> b) -> a -> b
$ LevelAtom' Term -> Level
atomicLevel (LevelAtom' Term -> Level) -> LevelAtom' Term -> Level
forall a b. (a -> b) -> a -> b
$ MetaId -> Elims -> LevelAtom' Term
forall t. MetaId -> [Elim' t] -> LevelAtom' t
MetaLevel MetaId
y Elims
bs'
| Bool
otherwise -> MetaId -> Elims -> Level -> m ()
forall (m :: * -> *).
(MonadMetaSolver m, MonadWarning m, MonadStatistics m,
MonadFresh ProblemId m, MonadFresh Int m) =>
MetaId -> Elims -> Level -> m ()
meta MetaId
y Elims
bs' (Level -> m ()) -> Level -> m ()
forall a b. (a -> b) -> a -> b
$ LevelAtom' Term -> Level
atomicLevel (LevelAtom' Term -> Level) -> LevelAtom' Term -> Level
forall a b. (a -> b) -> a -> b
$ MetaId -> Elims -> LevelAtom' Term
forall t. MetaId -> [Elim' t] -> LevelAtom' t
MetaLevel MetaId
x Elims
as'
(SinglePlus (Plus Integer
k (MetaLevel MetaId
x Elims
as')) :| [] , NonEmpty SingleLevel
_)
| Just Level
b' <- Integer -> Level -> Maybe Level
subLevel Integer
k Level
b -> MetaId -> Elims -> Level -> m ()
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'
(NonEmpty SingleLevel
_ , SinglePlus (Plus Integer
l (MetaLevel MetaId
y Elims
bs')) :| [])
| Just Level
a' <- Integer -> Level -> Maybe Level
subLevel Integer
l Level
a -> MetaId -> Elims -> Level -> m ()
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'
(NonEmpty SingleLevel, NonEmpty SingleLevel)
_ | Just Level
a' <- Level -> Level -> Maybe Level
levelMaxDiff Level
a Level
b
, Level
b Level -> Level -> Bool
forall a. Eq a => a -> a -> Bool
/= Integer -> Level
ClosedLevel Integer
0 -> Level -> Level -> m ()
forall (m :: * -> *). MonadConversion m => Level -> Level -> m ()
leqLevel Level
a' Level
b
(NonEmpty SingleLevel, NonEmpty SingleLevel)
_ | Just Level
b' <- Level -> Level -> Maybe Level
levelMaxDiff Level
b Level
a
, Level
a Level -> Level -> Bool
forall a. Eq a => a -> a -> Bool
/= Integer -> Level
ClosedLevel Integer
0 -> Level -> Level -> m ()
forall (m :: * -> *). MonadConversion m => Level -> Level -> m ()
leqLevel Level
b' Level
a
(NonEmpty SingleLevel
as , NonEmpty SingleLevel
bs)
| (SingleLevel -> Bool) -> [SingleLevel] -> Bool
forall (t :: * -> *) a. Foldable t => (a -> Bool) -> t a -> Bool
all SingleLevel -> Bool
isNeutralOrClosed (NonEmpty SingleLevel -> [SingleLevel]
forall a. NonEmpty a -> [a]
NonEmpty.toList NonEmpty SingleLevel
as [SingleLevel] -> [SingleLevel] -> [SingleLevel]
forall a. [a] -> [a] -> [a]
++ NonEmpty SingleLevel -> [SingleLevel]
forall a. NonEmpty a -> [a]
NonEmpty.toList NonEmpty SingleLevel
bs)
, Bool -> Bool
not ((SingleLevel -> Bool) -> [SingleLevel] -> Bool
forall (t :: * -> *) a. Foldable t => (a -> Bool) -> t a -> Bool
any SingleLevel -> Bool
hasMeta (NonEmpty SingleLevel -> [SingleLevel]
forall a. NonEmpty a -> [a]
NonEmpty.toList NonEmpty SingleLevel
as [SingleLevel] -> [SingleLevel] -> [SingleLevel]
forall a. [a] -> [a] -> [a]
++ NonEmpty SingleLevel -> [SingleLevel]
forall a. NonEmpty a -> [a]
NonEmpty.toList NonEmpty SingleLevel
bs))
, NonEmpty SingleLevel -> Int
forall (t :: * -> *) a. Foldable t => t a -> Int
length NonEmpty SingleLevel
as Int -> Int -> Bool
forall a. Eq a => a -> a -> Bool
== NonEmpty SingleLevel -> Int
forall (t :: * -> *) a. Foldable t => t a -> Int
length NonEmpty SingleLevel
bs -> do
VerboseKey -> Int -> VerboseKey -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> VerboseKey -> m ()
reportSLn VerboseKey
"tc.conv.level" Int
60 (VerboseKey -> m ()) -> VerboseKey -> m ()
forall a b. (a -> b) -> a -> b
$ VerboseKey
"equalLevel: all are neutral or closed"
(SingleLevel -> SingleLevel -> m ())
-> [SingleLevel] -> [SingleLevel] -> m ()
forall (m :: * -> *) a b c.
Applicative m =>
(a -> b -> m c) -> [a] -> [b] -> m ()
zipWithM_ (Term -> Term -> m ()
forall (m :: * -> *).
(MonadMetaSolver m, MonadWarning m, MonadStatistics m,
MonadFresh ProblemId m, MonadFresh Int m) =>
Term -> Term -> m ()
(===) (Term -> Term -> m ())
-> (SingleLevel -> Term) -> SingleLevel -> SingleLevel -> m ()
forall b c a. (b -> b -> c) -> (a -> b) -> a -> a -> c
`on` Level -> Term
levelTm (Level -> Term) -> (SingleLevel -> Level) -> SingleLevel -> Term
forall b c a. (b -> c) -> (a -> b) -> a -> c
. SingleLevel -> Level
unSingleLevel) (NonEmpty SingleLevel -> [SingleLevel]
forall a. NonEmpty a -> [a]
NonEmpty.toList NonEmpty SingleLevel
as) (NonEmpty SingleLevel -> [SingleLevel]
forall a. NonEmpty a -> [a]
NonEmpty.toList NonEmpty SingleLevel
bs)
(NonEmpty SingleLevel, NonEmpty SingleLevel)
_ | (Term, Term) -> Bool
forall a. TermLike a => a -> Bool
noMetas (Level -> Term
Level Level
a , Level -> Term
Level Level
b) -> m ()
notok
| Bool
otherwise -> m ()
forall b. m b
postpone
where
Term
a === :: Term -> Term -> m ()
=== Term
b = m Bool -> m () -> m ()
forall (m :: * -> *). Monad m => m Bool -> m () -> m ()
unlessM m Bool
forall (m :: * -> *). HasOptions m => m Bool
typeInType (m () -> m ()) -> m () -> m ()
forall a b. (a -> b) -> a -> b
$ do
Type
lvl <- m Type
forall (m :: * -> *). HasBuiltins m => m Type
levelType
CompareAs -> Term -> Term -> m ()
forall (m :: * -> *).
MonadConversion m =>
CompareAs -> Term -> Term -> m ()
equalAtom (Type -> CompareAs
AsTermsOf Type
lvl) Term
a Term
b
ok :: m ()
ok = () -> m ()
forall (m :: * -> *) a. Monad m => a -> m a
return ()
notok :: m ()
notok = m Bool -> m () -> m ()
forall (m :: * -> *). Monad m => m Bool -> m () -> m ()
unlessM m Bool
forall (m :: * -> *). HasOptions m => m Bool
typeInType m ()
forall b. m b
notOk
notOk :: m a
notOk = TypeError -> m a
forall (m :: * -> *) a.
(MonadTCEnv m, ReadTCState m, MonadError TCErr m) =>
TypeError -> m a
typeError (TypeError -> m a) -> TypeError -> m a
forall a b. (a -> b) -> a -> b
$ Comparison -> Level -> Level -> TypeError
UnequalLevel Comparison
CmpEq Level
a Level
b
postpone :: m b
postpone = do
VerboseKey -> Int -> TCM Doc -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCM Doc -> m ()
reportSDoc VerboseKey
"tc.conv.level" Int
30 (TCM Doc -> m ()) -> TCM Doc -> m ()
forall a b. (a -> b) -> a -> b
$ TCM Doc -> Int -> TCM Doc -> TCM Doc
forall (m :: * -> *).
Applicative m =>
m Doc -> Int -> m Doc -> m Doc
hang TCM Doc
"postponing:" Int
2 (TCM Doc -> TCM Doc) -> TCM Doc -> TCM Doc
forall a b. (a -> b) -> a -> b
$ TCM Doc -> Int -> TCM Doc -> TCM Doc
forall (m :: * -> *).
Applicative m =>
m Doc -> Int -> m Doc -> m Doc
hang (Level -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty Level
a TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> TCM Doc
"==") Int
0 (Level -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty Level
b)
m b
forall (m :: * -> *) a. MonadError TCErr m => m a
patternViolation
meta :: MetaId -> Elims -> Level -> m ()
meta MetaId
x Elims
as Level
b = do
VerboseKey -> Int -> VerboseKey -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> VerboseKey -> m ()
reportSLn VerboseKey
"tc.meta.level" Int
30 (VerboseKey -> m ()) -> VerboseKey -> m ()
forall a b. (a -> b) -> a -> b
$ VerboseKey
"Assigning meta level"
VerboseKey -> Int -> TCM Doc -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCM Doc -> m ()
reportSDoc VerboseKey
"tc.meta.level" Int
50 (TCM Doc -> m ()) -> TCM Doc -> m ()
forall a b. (a -> b) -> a -> b
$ TCM Doc
"meta" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> [TCM Doc] -> TCM Doc
forall (m :: * -> *). Monad m => [m Doc] -> m Doc
sep [[TCM Doc] -> TCM Doc
forall (m :: * -> *).
(Monad m, Semigroup (m Doc)) =>
[m Doc] -> m Doc
prettyList ([TCM Doc] -> TCM Doc) -> [TCM Doc] -> TCM Doc
forall a b. (a -> b) -> a -> b
$ (Elim' Term -> TCM Doc) -> Elims -> [TCM Doc]
forall a b. (a -> b) -> [a] -> [b]
map Elim' Term -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty Elims
as, Level -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty Level
b]
Type
lvl <- m Type
forall (m :: * -> *). HasBuiltins m => m Type
levelType
CompareDirection
-> MetaId
-> Elims
-> Term
-> CompareAs
-> (Term -> Term -> m ())
-> m ()
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) Term -> Term -> m ()
forall (m :: * -> *).
(MonadMetaSolver m, MonadWarning m, MonadStatistics m,
MonadFresh ProblemId m, MonadFresh Int m) =>
Term -> Term -> m ()
(===)
wrap :: m () -> m ()
wrap m ()
m = m ()
m m () -> (TCErr -> m ()) -> m ()
forall e (m :: * -> *) a.
MonadError e m =>
m a -> (e -> m a) -> m a
`catchError` \case
TypeError{} -> m ()
notok
TCErr
err -> TCErr -> m ()
forall e (m :: * -> *) a. MonadError e m => e -> m a
throwError TCErr
err
isNeutral :: SingleLevel -> Bool
isNeutral (SinglePlus (Plus Integer
_ NeutralLevel{})) = Bool
True
isNeutral SingleLevel
_ = Bool
False
isNeutralOrClosed :: SingleLevel -> Bool
isNeutralOrClosed (SingleClosed Integer
_) = Bool
True
isNeutralOrClosed (SinglePlus (Plus Integer
_ NeutralLevel{})) = Bool
True
isNeutralOrClosed SingleLevel
_ = Bool
False
hasMeta :: SingleLevel -> Bool
hasMeta (SinglePlus PlusLevel' Term
a) = case PlusLevel' Term
a of
Plus Integer
_ MetaLevel{} -> Bool
True
Plus Integer
_ (BlockedLevel MetaId
_ Term
v) -> Maybe MetaId -> Bool
forall a. Maybe a -> Bool
isJust (Maybe MetaId -> Bool) -> Maybe MetaId -> Bool
forall a b. (a -> b) -> a -> b
$ Term -> Maybe MetaId
forall a. TermLike a => a -> Maybe MetaId
firstMeta Term
v
Plus Integer
_ (NeutralLevel NotBlocked
_ Term
v) -> Maybe MetaId -> Bool
forall a. Maybe a -> Bool
isJust (Maybe MetaId -> Bool) -> Maybe MetaId -> Bool
forall a b. (a -> b) -> a -> b
$ Term -> Maybe MetaId
forall a. TermLike a => a -> Maybe MetaId
firstMeta Term
v
Plus Integer
_ (UnreducedLevel Term
v) -> Maybe MetaId -> Bool
forall a. Maybe a -> Bool
isJust (Maybe MetaId -> Bool) -> Maybe MetaId -> Bool
forall a b. (a -> b) -> a -> b
$ Term -> Maybe MetaId
forall a. TermLike a => a -> Maybe MetaId
firstMeta Term
v
hasMeta (SingleClosed Integer
_) = Bool
False
isThisMeta :: MetaId -> SingleLevel -> Bool
isThisMeta MetaId
x (SinglePlus (Plus Integer
_ (MetaLevel MetaId
y Elims
_))) = MetaId
x MetaId -> MetaId -> Bool
forall a. Eq a => a -> a -> Bool
== MetaId
y
isThisMeta MetaId
_ SingleLevel
_ = Bool
False
removeSubsumed :: Level -> Level -> (Level, Level)
removeSubsumed Level
a Level
b =
let as :: [SingleLevel]
as = NonEmpty SingleLevel -> [SingleLevel]
forall a. NonEmpty a -> [a]
NonEmpty.toList (NonEmpty SingleLevel -> [SingleLevel])
-> NonEmpty SingleLevel -> [SingleLevel]
forall a b. (a -> b) -> a -> b
$ Level -> NonEmpty SingleLevel
levelMaxView Level
a
bs :: [SingleLevel]
bs = NonEmpty SingleLevel -> [SingleLevel]
forall a. NonEmpty a -> [a]
NonEmpty.toList (NonEmpty SingleLevel -> [SingleLevel])
-> NonEmpty SingleLevel -> [SingleLevel]
forall a b. (a -> b) -> a -> b
$ Level -> NonEmpty SingleLevel
levelMaxView Level
b
a' :: Level
a' = [SingleLevel] -> Level
unSingleLevels ([SingleLevel] -> Level) -> [SingleLevel] -> Level
forall a b. (a -> b) -> a -> b
$ (SingleLevel -> Bool) -> [SingleLevel] -> [SingleLevel]
forall a. (a -> Bool) -> [a] -> [a]
filter (Bool -> Bool
not (Bool -> Bool) -> (SingleLevel -> Bool) -> SingleLevel -> Bool
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (SingleLevel -> [SingleLevel] -> Bool
forall (t :: * -> *).
Foldable t =>
SingleLevel -> t SingleLevel -> Bool
`isStrictlySubsumedBy` [SingleLevel]
bs)) [SingleLevel]
as
b' :: Level
b' = [SingleLevel] -> Level
unSingleLevels ([SingleLevel] -> Level) -> [SingleLevel] -> Level
forall a b. (a -> b) -> a -> b
$ (SingleLevel -> Bool) -> [SingleLevel] -> [SingleLevel]
forall a. (a -> Bool) -> [a] -> [a]
filter (Bool -> Bool
not (Bool -> Bool) -> (SingleLevel -> Bool) -> SingleLevel -> Bool
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (SingleLevel -> [SingleLevel] -> Bool
forall (t :: * -> *).
Foldable t =>
SingleLevel -> t SingleLevel -> Bool
`isStrictlySubsumedBy` [SingleLevel]
as)) [SingleLevel]
bs
in (Level
a',Level
b')
SingleLevel
x isStrictlySubsumedBy :: SingleLevel -> t SingleLevel -> Bool
`isStrictlySubsumedBy` t SingleLevel
ys = (SingleLevel -> Bool) -> t SingleLevel -> Bool
forall (t :: * -> *) a. Foldable t => (a -> Bool) -> t a -> Bool
any (SingleLevel -> SingleLevel -> Bool
`strictlySubsumes` SingleLevel
x) t SingleLevel
ys
SingleClosed Integer
m strictlySubsumes :: SingleLevel -> SingleLevel -> Bool
`strictlySubsumes` SingleClosed Integer
n = Integer
m Integer -> Integer -> Bool
forall a. Ord a => a -> a -> Bool
> Integer
n
SinglePlus (Plus Integer
m LevelAtom' Term
a) `strictlySubsumes` SingleClosed Integer
n = Integer
m Integer -> Integer -> Bool
forall a. Ord a => a -> a -> Bool
> Integer
n
SinglePlus (Plus Integer
m LevelAtom' Term
a) `strictlySubsumes` SinglePlus (Plus Integer
n LevelAtom' Term
b) = LevelAtom' Term
a LevelAtom' Term -> LevelAtom' Term -> Bool
forall a. Eq a => a -> a -> Bool
== LevelAtom' Term
b Bool -> Bool -> Bool
&& Integer
m Integer -> Integer -> Bool
forall a. Ord a => a -> a -> Bool
> Integer
n
SingleLevel
_ `strictlySubsumes` SingleLevel
_ = Bool
False
equalSort :: forall m. MonadConversion m => Sort -> Sort -> m ()
equalSort :: Sort -> Sort -> m ()
equalSort Sort
s1 Sort
s2 = do
Constraint -> m () -> m ()
forall (m :: * -> *).
MonadConstraint m =>
Constraint -> m () -> m ()
catchConstraint (Comparison -> Sort -> Sort -> Constraint
SortCmp Comparison
CmpEq Sort
s1 Sort
s2) (m () -> m ()) -> m () -> m ()
forall a b. (a -> b) -> a -> b
$ do
(Sort
s1,Sort
s2) <- (Sort, Sort) -> m (Sort, Sort)
forall a (m :: * -> *). (Reduce a, MonadReduce m) => a -> m a
reduce (Sort
s1,Sort
s2)
let yes :: m ()
yes = () -> m ()
forall (m :: * -> *) a. Monad m => a -> m a
return ()
no :: m a
no = TypeError -> m a
forall (m :: * -> *) a.
(MonadTCEnv m, ReadTCState m, MonadError TCErr m) =>
TypeError -> m a
typeError (TypeError -> m a) -> TypeError -> m a
forall a b. (a -> b) -> a -> b
$ Sort -> Sort -> TypeError
UnequalSorts Sort
s1 Sort
s2
VerboseKey -> Int -> TCM Doc -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCM Doc -> m ()
reportSDoc VerboseKey
"tc.conv.sort" Int
30 (TCM Doc -> m ()) -> TCM Doc -> m ()
forall a b. (a -> b) -> a -> b
$ [TCM Doc] -> TCM Doc
forall (m :: * -> *). Monad m => [m Doc] -> m Doc
sep
[ TCM Doc
"equalSort"
, [TCM Doc] -> TCM Doc
forall (m :: * -> *). Monad m => [m Doc] -> m Doc
vcat [ Int -> TCM Doc -> TCM Doc
forall (m :: * -> *). Functor m => Int -> m Doc -> m Doc
nest Int
2 (TCM Doc -> TCM Doc) -> TCM Doc -> TCM Doc
forall a b. (a -> b) -> a -> b
$ [TCM Doc] -> TCM Doc
forall (m :: * -> *). Monad m => [m Doc] -> m Doc
fsep [ Sort -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Sort
s1 TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> TCM Doc
"=="
, Sort -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Sort
s2 ]
, Int -> TCM Doc -> TCM Doc
forall (m :: * -> *). Functor m => Int -> m Doc -> m Doc
nest Int
2 (TCM Doc -> TCM Doc) -> TCM Doc -> TCM Doc
forall a b. (a -> b) -> a -> b
$ [TCM Doc] -> TCM Doc
forall (m :: * -> *). Monad m => [m Doc] -> m Doc
fsep [ Sort -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty Sort
s1 TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> TCM Doc
"=="
, Sort -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty Sort
s2 ]
]
]
Bool
propEnabled <- m Bool
forall (m :: * -> *). HasOptions m => m Bool
isPropEnabled
Bool
typeInTypeEnabled <- m Bool
forall (m :: * -> *). HasOptions m => m Bool
typeInType
case (Sort
s1, Sort
s2) of
(DummyS VerboseKey
s, Sort
_) -> VerboseKey -> m ()
forall (m :: * -> *) a b.
(ReportS [a], MonadDebug m, IsString a) =>
a -> m b
impossibleSort VerboseKey
s
(Sort
_, DummyS VerboseKey
s) -> VerboseKey -> m ()
forall (m :: * -> *) a b.
(ReportS [a], MonadDebug m, IsString a) =>
a -> m b
impossibleSort VerboseKey
s
(MetaS MetaId
x Elims
es , MetaS MetaId
y Elims
es')
| MetaId
x MetaId -> MetaId -> Bool
forall a. Eq a => a -> a -> Bool
== MetaId
y -> Sort -> Sort -> m ()
synEq Sort
s1 Sort
s2
| MetaId
x MetaId -> MetaId -> Bool
forall a. Ord a => a -> a -> Bool
< MetaId
y -> MetaId -> Elims -> Sort -> m ()
meta MetaId
y Elims
es' Sort
s1
| Bool
otherwise -> MetaId -> Elims -> Sort -> m ()
meta MetaId
x Elims
es Sort
s2
(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
(Type Level
a , Type Level
b ) -> Level -> Level -> m ()
forall (m :: * -> *). MonadConversion m => Level -> Level -> m ()
equalLevel Level
a Level
b m () -> m () -> m ()
forall (m :: * -> *) a. MonadError TCErr m => m a -> m a -> m a
`catchInequalLevel` m ()
forall a. m a
no
(Sort
SizeUniv , Sort
SizeUniv ) -> m ()
yes
(Prop Level
a , Prop Level
b ) -> Level -> Level -> m ()
forall (m :: * -> *). MonadConversion m => Level -> Level -> m ()
equalLevel Level
a Level
b m () -> m () -> m ()
forall (m :: * -> *) a. MonadError TCErr m => m a -> m a -> m a
`catchInequalLevel` m ()
forall a. m a
no
(Sort
Inf , Sort
Inf ) -> m ()
yes
(Type{} , Sort
Inf )
| Bool
typeInTypeEnabled -> m ()
yes
(Sort
Inf , Type{} )
| Bool
typeInTypeEnabled -> m ()
yes
(Sort
s1 , PiSort Dom Type
a Abs Sort
b) -> Sort -> Dom Type -> Abs Sort -> m ()
piSortEquals Sort
s1 Dom Type
a Abs Sort
b
(PiSort Dom Type
a Abs Sort
b , Sort
s2) -> Sort -> Dom Type -> Abs Sort -> m ()
piSortEquals Sort
s2 Dom Type
a Abs Sort
b
(Sort
s1 , FunSort Sort
a Sort
b) -> Sort -> Sort -> Sort -> m ()
funSortEquals Sort
s1 Sort
a Sort
b
(FunSort Sort
a Sort
b , Sort
s2) -> Sort -> Sort -> Sort -> m ()
funSortEquals Sort
s2 Sort
a Sort
b
(Sort
s1 , UnivSort Sort
s2) -> Sort -> Sort -> m ()
univSortEquals Sort
s1 Sort
s2
(UnivSort Sort
s1 , Sort
s2 ) -> Sort -> Sort -> m ()
univSortEquals Sort
s2 Sort
s1
(DefS QName
d Elims
es , DefS QName
d' Elims
es')
| QName
d QName -> QName -> Bool
forall a. Eq a => a -> a -> Bool
== QName
d' -> Sort -> Sort -> m ()
synEq Sort
s1 Sort
s2
| Bool
otherwise -> m ()
forall a. m a
no
(Sort
_ , Sort
_ ) -> m ()
forall a. m a
no
where
meta :: MetaId -> [Elim' Term] -> Sort -> m ()
meta :: MetaId -> Elims -> Sort -> m ()
meta MetaId
x Elims
es Sort
s = do
VerboseKey -> Int -> VerboseKey -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> VerboseKey -> m ()
reportSLn VerboseKey
"tc.meta.sort" Int
30 (VerboseKey -> m ()) -> VerboseKey -> m ()
forall a b. (a -> b) -> a -> b
$ VerboseKey
"Assigning meta sort"
VerboseKey -> Int -> TCM Doc -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCM Doc -> m ()
reportSDoc VerboseKey
"tc.meta.sort" Int
50 (TCM Doc -> m ()) -> TCM Doc -> m ()
forall a b. (a -> b) -> a -> b
$ TCM Doc
"meta" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> [TCM Doc] -> TCM Doc
forall (m :: * -> *). Monad m => [m Doc] -> m Doc
sep [MetaId -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty MetaId
x, [TCM Doc] -> TCM Doc
forall (m :: * -> *).
(Monad m, Semigroup (m Doc)) =>
[m Doc] -> m Doc
prettyList ([TCM Doc] -> TCM Doc) -> [TCM Doc] -> TCM Doc
forall a b. (a -> b) -> a -> b
$ (Elim' Term -> TCM Doc) -> Elims -> [TCM Doc]
forall a b. (a -> b) -> [a] -> [b]
map Elim' Term -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty Elims
es, Sort -> TCM Doc
forall (m :: * -> *) a. (Monad m, Pretty a) => a -> m Doc
pretty Sort
s]
CompareDirection
-> MetaId
-> Elims
-> Term
-> CompareAs
-> (Term -> Term -> m ())
-> m ()
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 Term -> Term -> m ()
forall a. HasCallStack => a
__IMPOSSIBLE__
synEq :: Sort -> Sort -> m ()
synEq :: Sort -> Sort -> m ()
synEq Sort
s1 Sort
s2 = do
let postpone :: m ()
postpone = Constraint -> m ()
forall (m :: * -> *). MonadConstraint m => Constraint -> m ()
addConstraint (Constraint -> m ()) -> Constraint -> m ()
forall a b. (a -> b) -> a -> b
$ Comparison -> Sort -> Sort -> Constraint
SortCmp Comparison
CmpEq Sort
s1 Sort
s2
Bool
doSynEq <- PragmaOptions -> Bool
optSyntacticEquality (PragmaOptions -> Bool) -> m PragmaOptions -> m Bool
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> m PragmaOptions
forall (m :: * -> *). HasOptions m => m PragmaOptions
pragmaOptions
if | Bool
doSynEq -> do
((Sort
s1,Sort
s2) , Bool
equal) <- Sort -> Sort -> m ((Sort, Sort), Bool)
forall a (m :: * -> *).
(Instantiate a, SynEq a, MonadReduce m) =>
a -> a -> m ((a, a), Bool)
SynEq.checkSyntacticEquality Sort
s1 Sort
s2
if | Bool
equal -> () -> m ()
forall (m :: * -> *) a. Monad m => a -> m a
return ()
| Bool
otherwise -> m ()
postpone
| Bool
otherwise -> m ()
postpone
set0 :: Sort
set0 = Integer -> Sort
mkType Integer
0
prop0 :: Sort
prop0 = Integer -> Sort
mkProp Integer
0
univSortEquals :: Sort -> Sort -> m ()
univSortEquals :: Sort -> Sort -> m ()
univSortEquals Sort
s1 Sort
s2 = do
VerboseKey -> Int -> TCM Doc -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCM Doc -> m ()
reportSDoc VerboseKey
"tc.conv.sort" Int
35 (TCM Doc -> m ()) -> TCM Doc -> m ()
forall a b. (a -> b) -> a -> b
$ [TCM Doc] -> TCM Doc
forall (m :: * -> *). Monad m => [m Doc] -> m Doc
vcat
[ TCM Doc
"univSortEquals"
, TCM Doc
" s1 =" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Sort -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Sort
s1
, TCM Doc
" s2 =" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Sort -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Sort
s2
]
let no :: m a
no = TypeError -> m a
forall (m :: * -> *) a.
(MonadTCEnv m, ReadTCState m, MonadError TCErr m) =>
TypeError -> m a
typeError (TypeError -> m a) -> TypeError -> m a
forall a b. (a -> b) -> a -> b
$ Sort -> Sort -> TypeError
UnequalSorts Sort
s1 (Sort -> Sort
forall t. Sort' t -> Sort' t
UnivSort Sort
s2)
case Sort
s1 of
Type Level
l1 -> do
Bool
propEnabled <- m Bool
forall (m :: * -> *). HasOptions m => m Bool
isPropEnabled
if | Sort
Inf <- Sort
s2 -> m ()
forall a. m a
no
| Sort
SizeUniv <- Sort
s2 -> m ()
forall a. m a
no
| Bool -> Bool
not Bool
propEnabled -> do
Level
l2 <- case Integer -> Level -> Maybe Level
subLevel Integer
1 Level
l1 of
Just Level
l2 -> Level -> m Level
forall (m :: * -> *) a. Monad m => a -> m a
return Level
l2
Maybe Level
Nothing -> do
Level
l2 <- m Level
forall (m :: * -> *). MonadMetaSolver m => m Level
newLevelMeta
Level -> Level -> m ()
forall (m :: * -> *). MonadConversion m => Level -> Level -> m ()
equalLevel Level
l1 (Level -> Level
levelSuc Level
l2)
Level -> m Level
forall (m :: * -> *) a. Monad m => a -> m a
return Level
l2
Sort -> Sort -> m ()
forall (m :: * -> *). MonadConversion m => Sort -> Sort -> m ()
equalSort (Level -> Sort
forall t. Level' t -> Sort' t
Type Level
l2) Sort
s2
| Bool
otherwise -> Sort -> Sort -> m ()
synEq (Level -> Sort
forall t. Level' t -> Sort' t
Type Level
l1) (Sort -> Sort
forall t. Sort' t -> Sort' t
UnivSort Sort
s2)
Sort
Inf -> do
Bool
infInInf <- (PragmaOptions -> Bool
optOmegaInOmega (PragmaOptions -> Bool) -> m PragmaOptions -> m Bool
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> m PragmaOptions
forall (m :: * -> *). HasOptions m => m PragmaOptions
pragmaOptions) m Bool -> m Bool -> m Bool
forall (m :: * -> *). Monad m => m Bool -> m Bool -> m Bool
`or2M` m Bool
forall (m :: * -> *). HasOptions m => m Bool
typeInType
if | Bool
infInInf -> Sort -> Sort -> m ()
forall (m :: * -> *). MonadConversion m => Sort -> Sort -> m ()
equalSort Sort
forall t. Sort' t
Inf Sort
s2
| Bool
otherwise -> m ()
forall a. m a
no
Prop{} -> m ()
forall a. m a
no
SizeUniv{} -> m ()
forall a. m a
no
Sort
_ -> Sort -> Sort -> m ()
synEq Sort
s1 (Sort -> Sort
forall t. Sort' t -> Sort' t
UnivSort Sort
s2)
piSortEquals :: Sort -> Dom Type -> Abs Sort -> m ()
piSortEquals :: Sort -> Dom Type -> Abs Sort -> m ()
piSortEquals Sort
s Dom Type
a NoAbs{} = m ()
forall a. HasCallStack => a
__IMPOSSIBLE__
piSortEquals Sort
s Dom Type
a bAbs :: Abs Sort
bAbs@(Abs VerboseKey
x Sort
b) = do
VerboseKey -> Int -> TCM Doc -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCM Doc -> m ()
reportSDoc VerboseKey
"tc.conv.sort" Int
35 (TCM Doc -> m ()) -> TCM Doc -> m ()
forall a b. (a -> b) -> a -> b
$ [TCM Doc] -> TCM Doc
forall (m :: * -> *). Monad m => [m Doc] -> m Doc
vcat
[ TCM Doc
"piSortEquals"
, TCM Doc
" s =" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Sort -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Sort
s
, TCM Doc
" a =" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Dom Type -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Dom Type
a
, TCM Doc
" b =" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> (VerboseKey, Dom Type) -> TCM Doc -> TCM Doc
forall b (m :: * -> *) a.
(AddContext b, MonadAddContext m) =>
b -> m a -> m a
addContext (VerboseKey
x,Dom Type
a) (Sort -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Sort
b)
]
Bool
propEnabled <- m Bool
forall (m :: * -> *). HasOptions m => m Bool
isPropEnabled
if | Sort -> Bool
definitelyNotInf Sort
s -> do
Sort
b' <- m Sort
forall (m :: * -> *). MonadMetaSolver m => m Sort
newSortMeta
(VerboseKey, Dom Type) -> m () -> m ()
forall b (m :: * -> *) a.
(AddContext b, MonadAddContext m) =>
b -> m a -> m a
addContext (VerboseKey
x,Dom Type
a) (m () -> m ()) -> m () -> m ()
forall a b. (a -> b) -> a -> b
$ Sort -> Sort -> m ()
forall (m :: * -> *). MonadConversion m => Sort -> Sort -> m ()
equalSort Sort
b (Int -> Sort -> Sort
forall t a. Subst t a => Int -> a -> a
raise Int
1 Sort
b')
Sort -> Sort -> Sort -> m ()
funSortEquals Sort
s (Dom Type -> Sort
forall a. LensSort a => a -> Sort
getSort Dom Type
a) Sort
b'
| Bool
otherwise -> Sort -> Sort -> m ()
synEq (Dom Type -> Abs Sort -> Sort
forall t. Dom' t (Type'' t t) -> Abs (Sort' t) -> Sort' t
PiSort Dom Type
a Abs Sort
bAbs) Sort
s
funSortEquals :: Sort -> Sort -> Sort -> m ()
funSortEquals :: Sort -> Sort -> Sort -> m ()
funSortEquals Sort
s0 Sort
s1 Sort
s2 = do
VerboseKey -> Int -> TCM Doc -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCM Doc -> m ()
reportSDoc VerboseKey
"tc.conv.sort" Int
35 (TCM Doc -> m ()) -> TCM Doc -> m ()
forall a b. (a -> b) -> a -> b
$ [TCM Doc] -> TCM Doc
forall (m :: * -> *). Monad m => [m Doc] -> m Doc
vcat
[ TCM Doc
"funSortEquals"
, TCM Doc
" s0 =" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Sort -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Sort
s0
, TCM Doc
" s1 =" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Sort -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Sort
s1
, TCM Doc
" s2 =" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Sort -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Sort
s2
]
Bool
propEnabled <- m Bool
forall (m :: * -> *). HasOptions m => m Bool
isPropEnabled
Bool
sizedTypesEnabled <- m Bool
forall (m :: * -> *). HasOptions m => m Bool
sizedTypesOption
case Sort
s0 of
Sort
Inf | Sort -> Bool
definitelyNotInf Sort
s1 Bool -> Bool -> Bool
&& Sort -> Bool
definitelyNotInf Sort
s2 -> do
TypeError -> m ()
forall (m :: * -> *) a.
(MonadTCEnv m, ReadTCState m, MonadError TCErr m) =>
TypeError -> m a
typeError (TypeError -> m ()) -> TypeError -> m ()
forall a b. (a -> b) -> a -> b
$ Sort -> Sort -> TypeError
UnequalSorts Sort
s0 (Sort -> Sort -> Sort
forall t. Sort' t -> Sort' t -> Sort' t
FunSort Sort
s1 Sort
s2)
| Sort -> Bool
definitelyNotInf Sort
s1 -> Sort -> Sort -> m ()
forall (m :: * -> *). MonadConversion m => Sort -> Sort -> m ()
equalSort Sort
forall t. Sort' t
Inf Sort
s2
| Sort -> Bool
definitelyNotInf Sort
s2 -> Sort -> Sort -> m ()
forall (m :: * -> *). MonadConversion m => Sort -> Sort -> m ()
equalSort Sort
forall t. Sort' t
Inf Sort
s1
| Bool
otherwise -> Sort -> Sort -> m ()
synEq Sort
s0 (Sort -> Sort -> Sort
forall t. Sort' t -> Sort' t -> Sort' t
FunSort Sort
s1 Sort
s2)
Type Level
l -> do
Level
l2 <- Sort -> m Level
forceType Sort
s2
Level -> Level -> m ()
forall (m :: * -> *). MonadConversion m => Level -> Level -> m ()
leqLevel Level
l2 Level
l
if | Bool
propEnabled -> case Sort -> Sort -> Maybe Sort
funSort' Sort
s1 (Level -> Sort
forall t. Level' t -> Sort' t
Type Level
l2) of
Just Sort
s -> Sort -> Sort -> m ()
forall (m :: * -> *). MonadConversion m => Sort -> Sort -> m ()
equalSort (Level -> Sort
forall t. Level' t -> Sort' t
Type Level
l) Sort
s
Maybe Sort
Nothing -> Sort -> Sort -> m ()
synEq (Level -> Sort
forall t. Level' t -> Sort' t
Type Level
l) (Sort -> Sort -> Sort
forall t. Sort' t -> Sort' t -> Sort' t
FunSort Sort
s1 (Sort -> Sort) -> Sort -> Sort
forall a b. (a -> b) -> a -> b
$ Level -> Sort
forall t. Level' t -> Sort' t
Type Level
l2)
| Bool
otherwise -> do
Level
l1 <- Sort -> m Level
forceType Sort
s1
Level -> Level -> m ()
forall (m :: * -> *). MonadConversion m => Level -> Level -> m ()
equalLevel Level
l (Level -> Level -> Level
levelLub Level
l1 Level
l2)
Prop Level
l -> do
Level
l2 <- Sort -> m Level
forceProp Sort
s2
Level -> Level -> m ()
forall (m :: * -> *). MonadConversion m => Level -> Level -> m ()
leqLevel Level
l2 Level
l
case Sort -> Sort -> Maybe Sort
funSort' Sort
s1 (Level -> Sort
forall t. Level' t -> Sort' t
Prop Level
l2) of
Just Sort
s -> Sort -> Sort -> m ()
forall (m :: * -> *). MonadConversion m => Sort -> Sort -> m ()
equalSort (Level -> Sort
forall t. Level' t -> Sort' t
Prop Level
l) Sort
s
Maybe Sort
Nothing -> Sort -> Sort -> m ()
synEq (Level -> Sort
forall t. Level' t -> Sort' t
Prop Level
l) (Sort -> Sort -> Sort
forall t. Sort' t -> Sort' t -> Sort' t
FunSort Sort
s1 (Sort -> Sort) -> Sort -> Sort
forall a b. (a -> b) -> a -> b
$ Level -> Sort
forall t. Level' t -> Sort' t
Prop Level
l2)
Sort
SizeUniv -> Sort -> Sort -> m ()
forall (m :: * -> *). MonadConversion m => Sort -> Sort -> m ()
equalSort Sort
forall t. Sort' t
SizeUniv Sort
s2
Sort
_ -> Sort -> Sort -> m ()
synEq Sort
s0 (Sort -> Sort -> Sort
forall t. Sort' t -> Sort' t -> Sort' t
FunSort Sort
s1 Sort
s2)
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
definitelyNotInf :: Sort -> Bool
definitelyNotInf :: Sort -> Bool
definitelyNotInf = \case
Sort
Inf -> Bool
False
Type{} -> Bool
True
Prop{} -> Bool
True
Sort
SizeUniv -> Bool
True
PiSort{} -> Bool
False
FunSort{} -> Bool
False
UnivSort{} -> Bool
False
MetaS{} -> Bool
False
DefS{} -> Bool
False
DummyS{} -> Bool
False
forceType :: Sort -> m Level
forceType :: Sort -> m Level
forceType (Type Level
l) = Level -> m Level
forall (m :: * -> *) a. Monad m => a -> m a
return Level
l
forceType Sort
s = do
Level
l <- m Level
forall (m :: * -> *). MonadMetaSolver m => m Level
newLevelMeta
Sort -> Sort -> m ()
forall (m :: * -> *). MonadConversion m => Sort -> Sort -> m ()
equalSort Sort
s (Level -> Sort
forall t. Level' t -> Sort' t
Type Level
l)
Level -> m Level
forall (m :: * -> *) a. Monad m => a -> m a
return Level
l
forceProp :: Sort -> m Level
forceProp :: Sort -> m Level
forceProp (Prop Level
l) = Level -> m Level
forall (m :: * -> *) a. Monad m => a -> m a
return Level
l
forceProp Sort
s = do
Level
l <- m Level
forall (m :: * -> *). MonadMetaSolver m => m Level
newLevelMeta
Sort -> Sort -> m ()
forall (m :: * -> *). MonadConversion m => Sort -> Sort -> m ()
equalSort Sort
s (Level -> Sort
forall t. Level' t -> Sort' t
Prop Level
l)
Level -> m Level
forall (m :: * -> *) a. Monad m => a -> m a
return Level
l
impossibleSort :: a -> m b
impossibleSort a
s = do
VerboseKey -> Int -> [a] -> m ()
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
]
m b
forall a. HasCallStack => a
__IMPOSSIBLE__
catchInequalLevel :: m a -> m a -> m a
catchInequalLevel m a
m m a
fail = m a
m m a -> (TCErr -> m a) -> m a
forall e (m :: * -> *) a.
MonadError e m =>
m a -> (e -> m a) -> m a
`catchError` \case
TypeError{} -> m a
fail
TCErr
err -> TCErr -> m a
forall e (m :: * -> *) a. MonadError e m => e -> m a
throwError TCErr
err
forallFaceMaps :: MonadConversion m => Term -> (Map.Map Int Bool -> MetaId -> Term -> m a) -> (Substitution -> m a) -> m [a]
forallFaceMaps :: Term
-> (Map Int Bool -> MetaId -> Term -> m a)
-> (Substitution -> m a)
-> m [a]
forallFaceMaps Term
t Map Int Bool -> MetaId -> Term -> m a
kb Substitution -> m a
k = do
VerboseKey -> Int -> TCM Doc -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCM Doc -> m ()
reportSDoc VerboseKey
"conv.forall" Int
20 (TCM Doc -> m ()) -> TCM Doc -> m ()
forall a b. (a -> b) -> a -> b
$
[TCM Doc] -> TCM Doc
forall (m :: * -> *). Monad m => [m Doc] -> m Doc
fsep [TCM Doc
"forallFaceMaps"
, Term -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Term
t
]
[(Map Int Bool, [Term])]
as <- Term -> m [(Map Int Bool, [Term])]
forall (m :: * -> *).
HasBuiltins m =>
Term -> m [(Map Int Bool, [Term])]
decomposeInterval Term
t
Bool -> Term
boolToI <- do
Term
io <- m Term
forall (m :: * -> *).
(HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) =>
m Term
primIOne
Term
iz <- m Term
forall (m :: * -> *).
(HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) =>
m Term
primIZero
(Bool -> Term) -> m (Bool -> Term)
forall (m :: * -> *) a. Monad m => a -> m a
return (\Bool
b -> if Bool
b then Term
io else Term
iz)
[(Map Int Bool, [Term])]
-> ((Map Int Bool, [Term]) -> m a) -> m [a]
forall (t :: * -> *) (m :: * -> *) a b.
(Traversable t, Monad m) =>
t a -> (a -> m b) -> m (t b)
forM [(Map Int Bool, [Term])]
as (((Map Int Bool, [Term]) -> m a) -> m [a])
-> ((Map Int Bool, [Term]) -> m a) -> m [a]
forall a b. (a -> b) -> a -> b
$ \ (Map Int Bool
ms,[Term]
ts) -> do
[Term]
-> (MetaId -> Term -> m a) -> (NotBlocked -> Term -> m a) -> m a
forall (m :: * -> *) (t :: * -> *) b.
(HasBuiltins m, MonadError TCErr m, MonadReduce m, Foldable t) =>
t Term
-> (MetaId -> Term -> m b) -> (NotBlocked -> Term -> m b) -> m b
ifBlockeds [Term]
ts (Map Int Bool -> MetaId -> Term -> m a
kb Map Int Bool
ms) ((NotBlocked -> Term -> m a) -> m a)
-> (NotBlocked -> Term -> m a) -> m a
forall a b. (a -> b) -> a -> b
$ \ NotBlocked
_ Term
_ -> do
let xs :: [(Int, Term)]
xs = ((Int, Bool) -> (Int, Term)) -> [(Int, Bool)] -> [(Int, Term)]
forall a b. (a -> b) -> [a] -> [b]
map (Int -> Int
forall a. a -> a
id (Int -> Int) -> (Bool -> Term) -> (Int, Bool) -> (Int, Term)
forall a c b d. (a -> c) -> (b -> d) -> (a, b) -> (c, d)
-*- Bool -> Term
boolToI) ([(Int, Bool)] -> [(Int, Term)]) -> [(Int, Bool)] -> [(Int, Term)]
forall a b. (a -> b) -> a -> b
$ Map Int Bool -> [(Int, Bool)]
forall k a. Map k a -> [(k, a)]
Map.toAscList Map Int Bool
ms
[Dom (Name, Type)]
cxt <- m [Dom (Name, Type)]
forall (m :: * -> *). MonadTCEnv m => m [Dom (Name, Type)]
getContext
VerboseKey -> Int -> TCM Doc -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCM Doc -> m ()
reportSDoc VerboseKey
"conv.forall" Int
20 (TCM Doc -> m ()) -> TCM Doc -> m ()
forall a b. (a -> b) -> a -> b
$
[TCM Doc] -> TCM Doc
forall (m :: * -> *). Monad m => [m Doc] -> m Doc
fsep [TCM Doc
"substContextN"
, [Dom (Name, Type)] -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM [Dom (Name, Type)]
cxt
, [(Int, Term)] -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM [(Int, Term)]
xs
]
([Dom (Name, Type)]
cxt',Substitution
sigma) <- [Dom (Name, Type)]
-> [(Int, Term)] -> m ([Dom (Name, Type)], Substitution)
forall (m :: * -> *).
MonadConversion m =>
[Dom (Name, Type)]
-> [(Int, Term)] -> m ([Dom (Name, Type)], Substitution)
substContextN [Dom (Name, Type)]
cxt [(Int, Term)]
xs
[(Dom (Name, Type), Term)]
resolved <- [(Int, Term)]
-> ((Int, Term) -> m (Dom (Name, Type), Term))
-> m [(Dom (Name, Type), Term)]
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) -> (,) (Dom (Name, Type) -> Term -> (Dom (Name, Type), Term))
-> m (Dom (Name, Type)) -> m (Term -> (Dom (Name, Type), Term))
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> Int -> m (Dom (Name, Type))
forall (m :: * -> *).
(MonadFail m, MonadTCEnv m) =>
Int -> m (Dom (Name, Type))
lookupBV Int
i m (Term -> (Dom (Name, Type), Term))
-> m Term -> m (Dom (Name, Type), Term)
forall (f :: * -> *) a b. Applicative f => f (a -> b) -> f a -> f b
<*> Term -> m Term
forall (m :: * -> *) a. Monad m => a -> m a
return (Substitution -> Term -> Term
forall t a. Subst t a => Substitution' t -> a -> a
applySubst Substitution
sigma Term
t))
Substitution
-> ([Dom (Name, Type)] -> [Dom (Name, Type)]) -> m a -> m a
forall (m :: * -> *) a.
MonadAddContext m =>
Substitution
-> ([Dom (Name, Type)] -> [Dom (Name, Type)]) -> m a -> m a
updateContext Substitution
sigma ([Dom (Name, Type)] -> [Dom (Name, Type)] -> [Dom (Name, Type)]
forall a b. a -> b -> a
const [Dom (Name, Type)]
cxt') (m a -> m a) -> m a -> m a
forall a b. (a -> b) -> a -> b
$
[(Dom (Name, Type), Term)] -> m a -> m a
forall (m :: * -> *) t a.
MonadAddContext m =>
[(Dom' t (Name, Type), Term)] -> m a -> m a
addBindings [(Dom (Name, Type), Term)]
resolved (m a -> m a) -> m a -> m a
forall a b. (a -> b) -> a -> b
$ do
Closure ()
cl <- () -> m (Closure ())
forall (m :: * -> *) a.
(MonadTCEnv m, ReadTCState m) =>
a -> m (Closure a)
buildClosure ()
Telescope
tel <- m Telescope
forall (m :: * -> *). (Applicative m, MonadTCEnv m) => m Telescope
getContextTelescope
ModuleName
m <- m ModuleName
forall (m :: * -> *). MonadTCEnv m => m ModuleName
currentModule
Maybe Substitution
sub <- ModuleName -> m (Maybe Substitution)
forall (m :: * -> *).
(MonadTCEnv m, ReadTCState m) =>
ModuleName -> m (Maybe Substitution)
getModuleParameterSub ModuleName
m
VerboseKey -> Int -> Names -> m ()
forall a (m :: * -> *).
(ReportS a, MonadDebug m) =>
VerboseKey -> Int -> a -> m ()
reportS VerboseKey
"conv.forall" Int
10
[ Int -> Char -> VerboseKey
forall a. Int -> a -> [a]
replicate Int
10 Char
'-'
, ModuleName -> VerboseKey
forall a. Show a => a -> VerboseKey
show (TCEnv -> ModuleName
envCurrentModule (TCEnv -> ModuleName) -> TCEnv -> ModuleName
forall a b. (a -> b) -> a -> b
$ Closure () -> TCEnv
forall a. Closure a -> TCEnv
clEnv Closure ()
cl)
, LetBindings -> VerboseKey
forall a. Show a => a -> VerboseKey
show (TCEnv -> LetBindings
envLetBindings (TCEnv -> LetBindings) -> TCEnv -> LetBindings
forall a b. (a -> b) -> a -> b
$ Closure () -> TCEnv
forall a. Closure a -> TCEnv
clEnv Closure ()
cl)
, Telescope -> VerboseKey
forall a. Show a => a -> VerboseKey
show Telescope
tel
, Substitution -> VerboseKey
forall a. Show a => a -> VerboseKey
show Substitution
sigma
, ModuleName -> VerboseKey
forall a. Show a => a -> VerboseKey
show ModuleName
m
, Maybe Substitution -> VerboseKey
forall a. Show a => a -> VerboseKey
show Maybe Substitution
sub
]
Substitution -> m a
k Substitution
sigma
where
ifBlockeds :: t Term
-> (MetaId -> Term -> m b) -> (NotBlocked -> Term -> m b) -> m b
ifBlockeds t Term
ts MetaId -> Term -> m b
blocked NotBlocked -> Term -> m b
unblocked = do
Term
and <- VerboseKey -> m Term
forall (m :: * -> *).
(HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) =>
VerboseKey -> m Term
getPrimitiveTerm VerboseKey
"primIMin"
Term
io <- m Term
forall (m :: * -> *).
(HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) =>
m Term
primIOne
let t :: Term
t = (Term -> Term -> Term) -> Term -> t Term -> Term
forall (t :: * -> *) a b.
Foldable t =>
(a -> b -> b) -> b -> t a -> b
foldr (\ Term
x Term
r -> Term
and Term -> [Arg Term] -> Term
forall t. Apply t => t -> [Arg Term] -> t
`apply` [Term -> Arg Term
forall e. e -> Arg e
argN Term
x,Term -> Arg Term
forall e. e -> Arg e
argN Term
r]) Term
io t Term
ts
Term
-> (MetaId -> Term -> m b) -> (NotBlocked -> Term -> m b) -> m b
forall t (m :: * -> *) a.
(Reduce t, IsMeta t, MonadReduce m, HasBuiltins m) =>
t -> (MetaId -> t -> m a) -> (NotBlocked -> t -> m a) -> m a
ifBlocked Term
t MetaId -> 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 = ArgInfo -> Name -> Term -> Type -> m a -> m a
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 :: [Dom (Name, Type)]
-> [(Int, Term)] -> m ([Dom (Name, Type)], Substitution)
substContextN [Dom (Name, Type)]
c [] = ([Dom (Name, Type)], Substitution)
-> m ([Dom (Name, Type)], Substitution)
forall (m :: * -> *) a. Monad m => a -> m a
return ([Dom (Name, Type)]
c, Substitution
forall a. Substitution' a
idS)
substContextN [Dom (Name, Type)]
c ((Int
i,Term
t):[(Int, Term)]
xs) = do
([Dom (Name, Type)]
c', Substitution
sigma) <- Int
-> Term
-> [Dom (Name, Type)]
-> m ([Dom (Name, Type)], Substitution)
forall (m :: * -> *).
MonadConversion m =>
Int
-> Term
-> [Dom (Name, Type)]
-> m ([Dom (Name, Type)], Substitution)
substContext Int
i Term
t [Dom (Name, Type)]
c
([Dom (Name, Type)]
c'', Substitution
sigma') <- [Dom (Name, Type)]
-> [(Int, Term)] -> m ([Dom (Name, Type)], Substitution)
forall (m :: * -> *).
MonadConversion m =>
[Dom (Name, Type)]
-> [(Int, Term)] -> m ([Dom (Name, Type)], Substitution)
substContextN [Dom (Name, Type)]
c' (((Int, Term) -> (Int, Term)) -> [(Int, Term)] -> [(Int, Term)]
forall a b. (a -> b) -> [a] -> [b]
map (Int -> Int -> Int
forall a. Num a => a -> a -> a
subtract Int
1 (Int -> Int) -> (Term -> Term) -> (Int, Term) -> (Int, Term)
forall a c b d. (a -> c) -> (b -> d) -> (a, b) -> (c, d)
-*- Substitution -> Term -> Term
forall t a. Subst t a => Substitution' t -> a -> a
applySubst Substitution
sigma) [(Int, Term)]
xs)
([Dom (Name, Type)], Substitution)
-> m ([Dom (Name, Type)], Substitution)
forall (m :: * -> *) a. Monad m => a -> m a
return ([Dom (Name, Type)]
c'', Substitution -> Substitution -> Substitution
forall t a. Subst t a => Substitution' t -> a -> a
applySubst Substitution
sigma' Substitution
sigma)
substContext :: MonadConversion m => Int -> Term -> Context -> m (Context , Substitution)
substContext :: Int
-> Term
-> [Dom (Name, Type)]
-> m ([Dom (Name, Type)], Substitution)
substContext Int
i Term
t [] = m ([Dom (Name, Type)], Substitution)
forall a. HasCallStack => a
__IMPOSSIBLE__
substContext Int
i Term
t (Dom (Name, Type)
x:[Dom (Name, Type)]
xs) | Int
i Int -> Int -> Bool
forall a. Eq a => a -> a -> Bool
== Int
0 = ([Dom (Name, Type)], Substitution)
-> m ([Dom (Name, Type)], Substitution)
forall (m :: * -> *) a. Monad m => a -> m a
return (([Dom (Name, Type)], Substitution)
-> m ([Dom (Name, Type)], Substitution))
-> ([Dom (Name, Type)], Substitution)
-> m ([Dom (Name, Type)], Substitution)
forall a b. (a -> b) -> a -> b
$ ([Dom (Name, Type)]
xs , Int -> Term -> Substitution
forall a. DeBruijn a => Int -> a -> Substitution' a
singletonS Int
0 Term
t)
substContext Int
i Term
t (Dom (Name, Type)
x:[Dom (Name, Type)]
xs) | Int
i Int -> Int -> Bool
forall a. Ord a => a -> a -> Bool
> Int
0 = do
VerboseKey -> Int -> TCM Doc -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCM Doc -> m ()
reportSDoc VerboseKey
"conv.forall" Int
20 (TCM Doc -> m ()) -> TCM Doc -> m ()
forall a b. (a -> b) -> a -> b
$
[TCM Doc] -> TCM Doc
forall (m :: * -> *). Monad m => [m Doc] -> m Doc
fsep [TCM Doc
"substContext"
, VerboseKey -> TCM Doc
forall (m :: * -> *). Monad m => VerboseKey -> m Doc
text (Int -> VerboseKey
forall a. Show a => a -> VerboseKey
show (Int
iInt -> Int -> Int
forall a. Num a => a -> a -> a
-Int
1))
, Term -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Term
t
, [Dom (Name, Type)] -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM [Dom (Name, Type)]
xs
]
([Dom (Name, Type)]
c,Substitution
sigma) <- Int
-> Term
-> [Dom (Name, Type)]
-> m ([Dom (Name, Type)], Substitution)
forall (m :: * -> *).
MonadConversion m =>
Int
-> Term
-> [Dom (Name, Type)]
-> m ([Dom (Name, Type)], Substitution)
substContext (Int
iInt -> Int -> Int
forall a. Num a => a -> a -> a
-Int
1) Term
t [Dom (Name, Type)]
xs
let e :: Dom (Name, Type)
e = Substitution -> Dom (Name, Type) -> Dom (Name, Type)
forall t a. Subst t a => Substitution' t -> a -> a
applySubst Substitution
sigma Dom (Name, Type)
x
([Dom (Name, Type)], Substitution)
-> m ([Dom (Name, Type)], Substitution)
forall (m :: * -> *) a. Monad m => a -> m a
return (Dom (Name, Type)
eDom (Name, Type) -> [Dom (Name, Type)] -> [Dom (Name, Type)]
forall a. a -> [a] -> [a]
:[Dom (Name, Type)]
c, Int -> Substitution -> Substitution
forall a. Int -> Substitution' a -> Substitution' a
liftS Int
1 Substitution
sigma)
substContext Int
i Term
t (Dom (Name, Type)
x:[Dom (Name, Type)]
xs) = m ([Dom (Name, Type)], Substitution)
forall a. HasCallStack => a
__IMPOSSIBLE__
compareInterval :: MonadConversion m => Comparison -> Type -> Term -> Term -> m ()
compareInterval :: Comparison -> Type -> Term -> Term -> m ()
compareInterval Comparison
cmp Type
i Term
t Term
u = do
VerboseKey -> Int -> TCM Doc -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCM Doc -> m ()
reportSDoc VerboseKey
"tc.conv.interval" Int
15 (TCM Doc -> m ()) -> TCM Doc -> m ()
forall a b. (a -> b) -> a -> b
$
[TCM Doc] -> TCM Doc
forall (m :: * -> *). Monad m => [m Doc] -> m Doc
sep [ TCM Doc
"{ compareInterval" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Term -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Term
t TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> TCM Doc
"=" TCM Doc -> TCM Doc -> TCM Doc
forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> Term -> TCM Doc
forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Term
u ]
Blocked Term
tb <- Term -> m (Blocked Term)
forall a (m :: * -> *).
(Reduce a, MonadReduce m) =>
a -> m (Blocked a)
reduceB Term
t
Blocked Term
ub <- Term -> m (Blocked Term)
forall a (m :: * -> *).
(Reduce a, MonadReduce m) =>
a -> m (Blocked a)
reduceB Term
u
let t :: Term
t = Blocked Term -> Term
forall t. Blocked t -> t
ignoreBlocking Blocked Term
tb
u :: Term
u = Blocked Term -> Term
forall t. Blocked t -> t
ignoreBlocking Blocked Term
ub
[(Map Int (Set Bool), [Term])]
it <- Term -> m [(Map Int (Set Bool), [Term])]
forall (m :: * -> *).
HasBuiltins m =>
Term -> m [(Map Int (Set Bool), [Term])]
decomposeInterval' Term
t
[(Map Int (Set Bool), [Term])]
iu <- Term -> m [(Map Int (Set Bool), [Term])]
forall (m :: * -> *).
HasBuiltins m =>
Term -> m [(Map Int (Set Bool), [Term])]
decomposeInterval' Term
u
case () of
()
_ | Blocked Term -> Bool
blockedOrMeta Blocked Term
tb Bool -> Bool -> Bool
|| Blocked Term -> Bool
blockedOrMeta Blocked Term
ub -> do
Type
interval <- m Term -> m Type
forall (tcm :: * -> *). Functor tcm => tcm Term -> tcm Type
elInf (m Term -> m Type) -> m Term -> m Type
forall a b. (a -> b) -> a -> b
$ m Term
forall (m :: * -> *).
(HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) =>
m Term
primInterval
Comparison -> CompareAs -> Term -> Term -> m ()
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 <- [(Map Int (Set Bool), [Term])]
-> [(Map Int (Set Bool), [Term])] -> m Bool
forall (m :: * -> *).
MonadConversion m =>
[(Map Int (Set Bool), [Term])]
-> [(Map Int (Set Bool), [Term])] -> m Bool
leqInterval [(Map Int (Set Bool), [Term])]
it [(Map Int (Set Bool), [Term])]
iu
Bool
y <- [(Map Int (Set Bool), [Term])]
-> [(Map Int (Set Bool), [Term])] -> m Bool
forall (m :: * -> *).
MonadConversion m =>
[(Map Int (Set Bool), [Term])]
-> [(Map Int (Set Bool), [Term])] -> m Bool
leqInterval [(Map Int (Set Bool), [Term])]
iu [(Map Int (Set Bool), [Term])]
it
let final :: Bool
final = [(Map Int (Set Bool), [Term])] -> Bool
isCanonical [(Map Int (Set Bool), [Term])]
it Bool -> Bool -> Bool
&& [(Map Int (Set Bool), [Term])] -> Bool
isCanonical [(Map Int (Set Bool), [Term])]
iu
if Bool
x Bool -> Bool -> Bool
&& Bool
y then VerboseKey -> Int -> TCM Doc -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCM Doc -> m ()
reportSDoc VerboseKey
"tc.conv.interval" Int
15 (TCM Doc -> m ()) -> TCM Doc -> m ()
forall a b. (a -> b) -> a -> b
$ TCM Doc
"Ok! }" else
if Bool
final then TypeError -> m ()
forall (m :: * -> *) a.
(MonadTCEnv m, ReadTCState m, MonadError TCErr m) =>
TypeError -> m a
typeError (TypeError -> m ()) -> TypeError -> m ()
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
VerboseKey -> Int -> TCM Doc -> m ()
forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> Int -> TCM Doc -> m ()
reportSDoc VerboseKey
"tc.conv.interval" Int
15 (TCM Doc -> m ()) -> TCM Doc -> m ()
forall a b. (a -> b) -> a -> b
$ TCM Doc
"Giving up! }"
m ()
forall (m :: * -> *) a. MonadError TCErr m => m a
patternViolation
where
blockedOrMeta :: Blocked Term -> Bool
blockedOrMeta Blocked{} = Bool
True
blockedOrMeta (NotBlocked NotBlocked
_ (MetaV{})) = Bool
True
blockedOrMeta Blocked Term
_ = Bool
False
type Conj = (Map.Map Int (Set.Set Bool),[Term])
isCanonical :: [Conj] -> Bool
isCanonical :: [(Map Int (Set Bool), [Term])] -> Bool
isCanonical = ((Map Int (Set Bool), [Term]) -> Bool)
-> [(Map Int (Set Bool), [Term])] -> Bool
forall (t :: * -> *) a. Foldable t => (a -> Bool) -> t a -> Bool
all ([Term] -> Bool
forall (t :: * -> *) a. Foldable t => t a -> Bool
null ([Term] -> Bool)
-> ((Map Int (Set Bool), [Term]) -> [Term])
-> (Map Int (Set Bool), [Term])
-> Bool
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (Map Int (Set Bool), [Term]) -> [Term]
forall a b. (a, b) -> b
snd)
leqInterval :: MonadConversion m => [Conj] -> [Conj] -> m Bool
leqInterval :: [(Map Int (Set Bool), [Term])]
-> [(Map Int (Set Bool), [Term])] -> m Bool
leqInterval [(Map Int (Set Bool), [Term])]
r [(Map Int (Set Bool), [Term])]
q =
[Bool] -> Bool
forall (t :: * -> *). Foldable t => t Bool -> Bool
and ([Bool] -> Bool) -> m [Bool] -> m Bool
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> [(Map Int (Set Bool), [Term])]
-> ((Map Int (Set Bool), [Term]) -> m Bool) -> m [Bool]
forall (t :: * -> *) (m :: * -> *) a b.
(Traversable t, Monad m) =>
t a -> (a -> m b) -> m (t b)
forM [(Map Int (Set Bool), [Term])]
r (\ (Map Int (Set Bool), [Term])
r_i ->
[Bool] -> Bool
forall (t :: * -> *). Foldable t => t Bool -> Bool
or ([Bool] -> Bool) -> m [Bool] -> m Bool
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> [(Map Int (Set Bool), [Term])]
-> ((Map Int (Set Bool), [Term]) -> m Bool) -> m [Bool]
forall (t :: * -> *) (m :: * -> *) a b.
(Traversable t, Monad m) =>
t a -> (a -> m b) -> m (t b)
forM [(Map Int (Set Bool), [Term])]
q (\ (Map Int (Set Bool), [Term])
q_j -> (Map Int (Set Bool), [Term])
-> (Map Int (Set Bool), [Term]) -> m Bool
forall (m :: * -> *).
MonadConversion m =>
(Map Int (Set Bool), [Term])
-> (Map Int (Set Bool), [Term]) -> m Bool
leqConj (Map Int (Set Bool), [Term])
r_i (Map Int (Set Bool), [Term])
q_j))
leqConj :: MonadConversion m => Conj -> Conj -> m Bool
leqConj :: (Map Int (Set Bool), [Term])
-> (Map Int (Set Bool), [Term]) -> m Bool
leqConj (Map Int (Set Bool)
rs,[Term]
rst) (Map Int (Set Bool)
qs,[Term]
qst) = do
case Map Int (Set Bool) -> Set (Int, Bool)
forall a b. (Ord a, Ord b) => Map a (Set b) -> Set (a, b)
toSet Map Int (Set Bool)
qs Set (Int, Bool) -> Set (Int, Bool) -> Bool
forall a. Ord a => Set a -> Set a -> Bool
`Set.isSubsetOf` Map Int (Set Bool) -> Set (Int, Bool)
forall a b. (Ord a, Ord b) => Map a (Set b) -> Set (a, b)
toSet Map Int (Set Bool)
rs of
Bool
False -> Bool -> m Bool
forall (m :: * -> *) a. Monad m => a -> m a
return Bool
False
Bool
True -> do
Type
interval <- m Term -> m Type
forall (tcm :: * -> *). Functor tcm => tcm Term -> tcm Type
elInf (m Term -> m Type) -> m Term -> m Type
forall a b. (a -> b) -> a -> b
$ Term -> Maybe Term -> Term
forall a. a -> Maybe a -> a
fromMaybe Term
forall a. HasCallStack => a
__IMPOSSIBLE__ (Maybe Term -> Term) -> m (Maybe Term) -> m Term
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> VerboseKey -> m (Maybe Term)
forall (m :: * -> *). HasBuiltins m => VerboseKey -> m (Maybe Term)
getBuiltin' VerboseKey
builtinInterval
let eqT :: Term -> Term -> m Bool
eqT Term
t Term
u = m () -> m Bool
forall (m :: * -> *).
(MonadConstraint m, MonadWarning m, MonadError TCErr m,
MonadFresh ProblemId m) =>
m () -> m Bool
tryConversion (Comparison -> CompareAs -> Term -> Term -> m ()
forall (m :: * -> *).
MonadConversion m =>
Comparison -> CompareAs -> Term -> Term -> m ()
compareAtom Comparison
CmpEq (Type -> CompareAs
AsTermsOf Type
interval) Term
t Term
u)
let listSubset :: t Term -> t Term -> f Bool
listSubset t Term
ts t Term
us = t Bool -> Bool
forall (t :: * -> *). Foldable t => t Bool -> Bool
and (t Bool -> Bool) -> f (t Bool) -> f Bool
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> t Term -> (Term -> f Bool) -> f (t Bool)
forall (t :: * -> *) (m :: * -> *) a b.
(Traversable t, Monad m) =>
t a -> (a -> m b) -> m (t b)
forM t Term
ts (\ Term
t ->
t Bool -> Bool
forall (t :: * -> *). Foldable t => t Bool -> Bool
or (t Bool -> Bool) -> f (t Bool) -> f Bool
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> t Term -> (Term -> f Bool) -> f (t Bool)
forall (t :: * -> *) (m :: * -> *) a b.
(Traversable t, Monad m) =>
t a -> (a -> m b) -> m (t b)
forM t Term
us (\ Term
u -> Term -> Term -> f Bool
forall (m :: * -> *).
(MonadWarning m, MonadMetaSolver m, MonadStatistics m,
MonadFresh ProblemId m, MonadFresh Int m) =>
Term -> Term -> m Bool
eqT Term
t Term
u))
[Term] -> [Term] -> m Bool
forall (f :: * -> *) (t :: * -> *) (t :: * -> *).
(Traversable t, Traversable t, MonadWarning f, MonadMetaSolver f,
MonadStatistics f, MonadFresh ProblemId f, MonadFresh Int f) =>
t Term -> t Term -> f Bool
listSubset [Term]
qst [Term]
rst
where
toSet :: Map a (Set b) -> Set (a, b)
toSet Map a (Set b)
m = [(a, b)] -> Set (a, b)
forall a. Ord a => [a] -> Set a
Set.fromList [ (a
i,b
b) | (a
i,Set b
bs) <- Map a (Set b) -> [(a, Set b)]
forall k a. Map k a -> [(k, a)]
Map.toList Map a (Set b)
m, b
b <- Set b -> [b]
forall a. Set a -> [a]
Set.toList Set b
bs]
equalTermOnFace :: MonadConversion m => Term -> Type -> Term -> Term -> m ()
equalTermOnFace :: Term -> Type -> Term -> Term -> m ()
equalTermOnFace = Comparison -> Term -> Type -> Term -> Term -> m ()
forall (m :: * -> *).
MonadConversion m =>
Comparison -> Term -> Type -> Term -> Term -> m ()
compareTermOnFace Comparison
CmpEq
compareTermOnFace :: MonadConversion m => Comparison -> Term -> Type -> Term -> Term -> m ()
compareTermOnFace :: Comparison -> Term -> Type -> Term -> Term -> m ()
compareTermOnFace = (Comparison -> Type -> Term -> Term -> m ())
-> Comparison -> Term -> Type -> Term -> Term -> m ()
forall (m :: * -> *).
MonadConversion m =>
(Comparison -> Type -> Term -> Term -> m ())
-> Comparison -> Term -> Type -> Term -> Term -> m ()
compareTermOnFace' Comparison -> Type -> Term -> Term -> m ()
forall (m :: * -> *).
MonadConversion m =>
Comparison -> Type -> Term -> Term -> m ()
compareTerm
compareTermOnFace' :: MonadConversion m => (Comparison -> Type -> Term -> Term -> m ()) -> Comparison -> Term -> Type -> Term -> Term -> m ()
compareTermOnFace' :: (Comparison -> Type -> Term -> Term -> m ())
-> Comparison -> Term -> Type -> Term -> Term -> m ()
compareTermOnFace' Comparison -> Type -> Term -> Term -> m ()
k Comparison
cmp Term
phi Type
ty Term
u Term
v = do
Term
phi <- Term -> m Term
forall a (m :: * -> *). (Reduce a, MonadReduce m) => a -> m a
reduce Term
phi
[()]
_ <- Term
-> (Map Int Bool -> MetaId -> Term -> m ())
-> (Substitution -> m ())
-> m [()]
forall (m :: * -> *) a.
MonadConversion m =>
Term
-> (Map Int Bool -> MetaId -> Term -> m a)
-> (Substitution -> m a)
-> m [a]
forallFaceMaps Term
phi Map Int Bool -> MetaId -> Term -> m ()
forall (m :: * -> *) p.
(HasBuiltins m, MonadConstraint m) =>
Map Int Bool -> p -> Term -> m ()
postponed
((Substitution -> m ()) -> m [()])
-> (Substitution -> m ()) -> m [()]
forall a b. (a -> b) -> a -> b
$ \ Substitution
alpha -> Comparison -> Type -> Term -> Term -> m ()
k Comparison
cmp (Substitution -> Type -> Type
forall t a. Subst t a => Substitution' t -> a -> a
applySubst Substitution
alpha Type
ty) (Substitution -> Term -> Term
forall t a. Subst t a => Substitution' t -> a -> a
applySubst Substitution
alpha Term
u) (Substitution -> Term -> Term
forall t a. Subst t a => Substitution' t -> a -> a
applySubst Substitution
alpha Term
v)
() -> m ()
forall (m :: * -> *) a. Monad m => a -> m a
return ()
where
postponed :: Map Int Bool -> p -> Term -> m ()
postponed Map Int Bool
ms p
i Term
psi = do
Term
phi <- Names -> NamesT m Term -> m Term
forall (m :: * -> *) a. Names -> NamesT m a -> m a
runNamesT [] (NamesT m Term -> m Term) -> NamesT m Term -> m Term
forall a b. (a -> b) -> a -> b
$ do
Term
imin <- m Term -> NamesT m Term
forall (m :: * -> *) a. Monad m => m a -> NamesT m a
cl (m Term -> NamesT m Term) -> m Term -> NamesT m Term
forall a b. (a -> b) -> a -> b
$ VerboseKey -> m Term
forall (m :: * -> *).
(HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) =>
VerboseKey -> m Term
getPrimitiveTerm VerboseKey
"primIMin"
Term
ineg <- m Term -> NamesT m Term
forall (m :: * -> *) a. Monad m => m a -> NamesT m a
cl (m Term -> NamesT m Term) -> m Term -> NamesT m Term
forall a b. (a -> b) -> a -> b
$ VerboseKey -> m Term
forall (m :: * -> *).
(HasBuiltins m, MonadError TCErr m, MonadTCEnv m, ReadTCState m) =>
VerboseKey -> m Term
getPrimitiveTerm VerboseKey
"primINeg"
NamesT m Term
psi <- Term -> NamesT m (NamesT m Term)
forall (m :: * -> *) t a.
(MonadFail m, Subst t a) =>
a -> NamesT m (NamesT m a)
open Term
psi
let phi :: NamesT m Term
phi = ((Int, Bool) -> NamesT m Term -> NamesT m Term)
-> NamesT m Term -> [(Int, Bool)] -> NamesT m Term
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 <- Term -> NamesT m (NamesT m Term)
forall (m :: * -> *) t a.
(MonadFail m, Subst t a) =>
a -> NamesT m (NamesT m a)
open (Int -> Term
var Int
i); Term -> NamesT m Term
forall (f :: * -> *) a. Applicative f => a -> f a
pure Term
imin NamesT m Term -> NamesT m Term -> NamesT m Term
forall (tcm :: * -> *).
Monad tcm =>
tcm Term -> tcm Term -> tcm Term
<@> (if Bool
b then NamesT m Term
i else Term -> NamesT m Term
forall (f :: * -> *) a. Applicative f => a -> f a
pure Term
ineg NamesT m Term -> NamesT m Term -> NamesT m Term
forall (tcm :: * -> *).
Monad tcm =>
tcm Term -> tcm Term -> tcm Term
<@> NamesT m Term
i) NamesT m Term -> NamesT m Term -> NamesT m Term
forall (tcm :: * -> *).
Monad tcm =>
tcm Term -> tcm Term -> tcm Term
<@> NamesT m Term
r)
NamesT m Term
psi (Map Int Bool -> [(Int, Bool)]
forall k a. Map k a -> [(k, a)]
Map.toList Map Int Bool
ms)
NamesT m Term
phi
Constraint -> m ()
forall (m :: * -> *). MonadConstraint m => Constraint -> m ()
addConstraint (Comparison -> Term -> Type -> Term -> Term -> Constraint
ValueCmpOnFace Comparison
cmp Term
phi Type
ty Term
u Term
v)
bothAbsurd :: MonadConversion m => QName -> QName -> m Bool
bothAbsurd :: QName -> QName -> m Bool
bothAbsurd QName
f QName
f'
| QName -> Bool
isAbsurdLambdaName QName
f, QName -> Bool
isAbsurdLambdaName QName
f' = do
Definition
def <- QName -> m Definition
forall (m :: * -> *). HasConstInfo m => QName -> m Definition
getConstInfo QName
f
Definition
def' <- QName -> m Definition
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 }] }) -> Bool -> m Bool
forall (m :: * -> *) a. Monad m => a -> m a
return Bool
True
(Defn, Defn)
_ -> Bool -> m Bool
forall (m :: * -> *) a. Monad m => a -> m a
return Bool
False
| Bool
otherwise = Bool -> m Bool
forall (m :: * -> *) a. Monad m => a -> m a
return Bool
False