{-# LANGUAGE NondecreasingIndentation #-}

{-| Compile-time irrelevance.

In type theory with compile-time irrelevance à la Pfenning (LiCS 2001),
variables in the context are annotated with relevance attributes.
@@
  Γ = r₁x₁:A₁, ..., rⱼxⱼ:Aⱼ
@@
To handle irrelevant projections, we also record the current relevance
attribute in the judgement.  For instance, this attribute is equal to
to 'Irrelevant' if we are in an irrelevant position, like an
irrelevant argument.
@@
  Γ ⊢r t : A
@@
Only relevant variables can be used:
@@

  (Relevant x : A) ∈ Γ
  --------------------
  Γ  ⊢r  x  :  A
@@
Irrelevant global declarations can only be used if @r = Irrelevant@.

When we enter a @r'@-relevant function argument, we compose the @r@ with @r'@
and (left-)divide the attributes in the context by @r'@.
@@
  Γ  ⊢r  t  :  (r' x : A) → B      r' \ Γ  ⊢(r'·r)  u  :  A
  ---------------------------------------------------------
  Γ  ⊢r  t u  :  B[u/x]
@@
No surprises for abstraction:
@@

  Γ, (r' x : A)  ⊢r  :  B
  -----------------------------
  Γ  ⊢r  λxt  :  (r' x : A) → B
@@

This is different for runtime irrelevance (erasure) which is ``flat'',
meaning that once one is in an irrelevant context, all new assumptions will
be usable, since they are turned relevant once entering the context.
See Conor McBride (WadlerFest 2016), for a type system in this spirit:

We use such a rule for runtime-irrelevance:
@@
  Γ, (q \ q') x : A  ⊢q  t  :  B
  ------------------------------
  Γ  ⊢q  λxt  :  (q' x : A) → B
@@

Conor's system is however set up differently, with a very different
variable rule:

@@

  (q x : A) ∈ Γ
  --------------
  Γ  ⊢q  x  :  A

  Γ, (q·p) x : A  ⊢q  t  :  B
  -----------------------------
  Γ  ⊢q  λxt  :  (p x : A) → B

  Γ  ⊢q  t  :  (p x : A) → B       Γ'  ⊢qp  u  :  A
  -------------------------------------------------
  Γ + Γ'  ⊢q  t u  :  B[u/x]
@@


-}

module Agda.TypeChecking.Irrelevance where

import Control.Arrow (second)

import Control.Monad.Except

import qualified Data.Map as Map

import Agda.Interaction.Options

import Agda.Syntax.Common
import Agda.Syntax.Internal
import Agda.Syntax.Concrete.Pretty

import Agda.TypeChecking.Monad
import Agda.TypeChecking.Pretty
import Agda.TypeChecking.Reduce
import Agda.TypeChecking.Substitute.Class

import Agda.Utils.Function
import Agda.Utils.Lens
import Agda.Utils.Maybe
import Agda.Utils.Monad
import Agda.Utils.WithDefault
import qualified Agda.Utils.Null as Null
import Agda.Utils.WithDefault (collapseDefault)

-- | data 'Relevance'
--   see "Agda.Syntax.Common".

-- * Operations on 'Dom'.

-- | Prepare parts of a parameter telescope for abstraction in constructors
--   and projections.
hideAndRelParams :: (LensHiding a, LensRelevance a) => a -> a
hideAndRelParams :: forall a. (LensHiding a, LensRelevance a) => a -> a
hideAndRelParams = forall a. LensHiding a => a -> a
hideOrKeepInstance forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall a. LensRelevance a => (Relevance -> Relevance) -> a -> a
mapRelevance Relevance -> Relevance
nonStrictToIrr

-- * Operations on 'Context'.

-- | Modify the context whenever going from the l.h.s. (term side)
--   of the typing judgement to the r.h.s. (type side).
workOnTypes :: (MonadTCEnv m, HasOptions m, MonadDebug m)
            => m a -> m a
workOnTypes :: forall (m :: * -> *) a.
(MonadTCEnv m, HasOptions m, MonadDebug m) =>
m a -> m a
workOnTypes m a
cont = do
  Bool
allowed <- PragmaOptions -> Bool
optExperimentalIrrelevance forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> forall (m :: * -> *). HasOptions m => m PragmaOptions
pragmaOptions
  forall (m :: * -> *) a.
MonadDebug m =>
VerboseKey -> VerboseLevel -> VerboseKey -> m a -> m a
verboseBracket VerboseKey
"tc.irr" VerboseLevel
60 VerboseKey
"workOnTypes" forall a b. (a -> b) -> a -> b
$ forall (m :: * -> *) a. MonadTCEnv m => Bool -> m a -> m a
workOnTypes' Bool
allowed m a
cont

-- | Internal workhorse, expects value of --experimental-irrelevance flag
--   as argument.
workOnTypes' :: (MonadTCEnv m) => Bool -> m a -> m a
workOnTypes' :: forall (m :: * -> *) a. MonadTCEnv m => Bool -> m a -> m a
workOnTypes' Bool
experimental
  = (if Bool
experimental
     then forall (tcm :: * -> *) a.
MonadTCEnv tcm =>
(forall e. Dom e -> Dom e) -> tcm a -> tcm a
modifyContextInfo (forall a. LensRelevance a => (Relevance -> Relevance) -> a -> a
mapRelevance Relevance -> Relevance
irrToNonStrict)
     else forall a. a -> a
id)
  forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall (tcm :: * -> *) q a.
(MonadTCEnv tcm, LensQuantity q) =>
q -> tcm a -> tcm a
applyQuantityToContext Quantity
zeroQuantity
  forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall (m :: * -> *) a. MonadTCEnv m => m a -> m a
typeLevelReductions
  forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall (m :: * -> *) a.
MonadTCEnv m =>
(TCEnv -> TCEnv) -> m a -> m a
localTC (\ TCEnv
e -> TCEnv
e { envWorkingOnTypes :: Bool
envWorkingOnTypes = Bool
True })

-- | (Conditionally) wake up irrelevant variables and make them relevant.
--   For instance,
--   in an irrelevant function argument otherwise irrelevant variables
--   may be used, so they are awoken before type checking the argument.
--
--   Also allow the use of irrelevant definitions.
applyRelevanceToContext :: (MonadTCEnv tcm, LensRelevance r) => r -> tcm a -> tcm a
applyRelevanceToContext :: forall (tcm :: * -> *) r a.
(MonadTCEnv tcm, LensRelevance r) =>
r -> tcm a -> tcm a
applyRelevanceToContext r
thing =
  case forall a. LensRelevance a => a -> Relevance
getRelevance r
thing of
    Relevance
Relevant -> forall a. a -> a
id
    Relevance
rel      -> forall (tcm :: * -> *) a.
MonadTCEnv tcm =>
Relevance -> tcm a -> tcm a
applyRelevanceToContextOnly   Relevance
rel
              forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall (tcm :: * -> *) a.
MonadTCEnv tcm =>
Relevance -> tcm a -> tcm a
applyRelevanceToJudgementOnly Relevance
rel

-- | (Conditionally) wake up irrelevant variables and make them relevant.
--   For instance,
--   in an irrelevant function argument otherwise irrelevant variables
--   may be used, so they are awoken before type checking the argument.
--
--   Precondition: @Relevance /= Relevant@
applyRelevanceToContextOnly :: (MonadTCEnv tcm) => Relevance -> tcm a -> tcm a
applyRelevanceToContextOnly :: forall (tcm :: * -> *) a.
MonadTCEnv tcm =>
Relevance -> tcm a -> tcm a
applyRelevanceToContextOnly Relevance
rel = forall (m :: * -> *) a.
MonadTCEnv m =>
(TCEnv -> TCEnv) -> m a -> m a
localTC
  forall a b. (a -> b) -> a -> b
$ forall i o. Lens' i o -> LensMap i o
over Lens' Context TCEnv
eContext     (forall a b. (a -> b) -> [a] -> [b]
map forall a b. (a -> b) -> a -> b
$ forall a. LensRelevance a => Relevance -> a -> a
inverseApplyRelevance Relevance
rel)
  forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall i o. Lens' i o -> LensMap i o
over Lens' LetBindings TCEnv
eLetBindings (forall a b k. (a -> b) -> Map k a -> Map k b
Map.map forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
fmap forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall (a :: * -> * -> *) b c d.
Arrow a =>
a b c -> a (d, b) (d, c)
second forall a b. (a -> b) -> a -> b
$ forall a. LensRelevance a => Relevance -> a -> a
inverseApplyRelevance Relevance
rel)

-- | Apply relevance @rel@ the the relevance annotation of the (typing/equality)
--   judgement.  This is part of the work done when going into a @rel@-context.
--
--   Precondition: @Relevance /= Relevant@
applyRelevanceToJudgementOnly :: (MonadTCEnv tcm) => Relevance -> tcm a -> tcm a
applyRelevanceToJudgementOnly :: forall (tcm :: * -> *) a.
MonadTCEnv tcm =>
Relevance -> tcm a -> tcm a
applyRelevanceToJudgementOnly = forall (m :: * -> *) a.
MonadTCEnv m =>
(TCEnv -> TCEnv) -> m a -> m a
localTC forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall i o. Lens' i o -> LensMap i o
over Lens' Relevance TCEnv
eRelevance forall b c a. (b -> c) -> (a -> b) -> a -> c
. Relevance -> Relevance -> Relevance
composeRelevance

-- | Like 'applyRelevanceToContext', but only act on context if
--   @--irrelevant-projections@.
--   See issue #2170.
applyRelevanceToContextFunBody :: (MonadTCM tcm, LensRelevance r) => r -> tcm a -> tcm a
applyRelevanceToContextFunBody :: forall (tcm :: * -> *) r a.
(MonadTCM tcm, LensRelevance r) =>
r -> tcm a -> tcm a
applyRelevanceToContextFunBody r
thing tcm a
cont =
  case forall a. LensRelevance a => a -> Relevance
getRelevance r
thing of
    Relevance
Relevant -> tcm a
cont
    Relevance
rel -> forall (m :: * -> *) a.
Monad m =>
m Bool -> (m a -> m a) -> m a -> m a
applyWhenM (PragmaOptions -> Bool
optIrrelevantProjections forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> forall (m :: * -> *). HasOptions m => m PragmaOptions
pragmaOptions)
      (forall (tcm :: * -> *) a.
MonadTCEnv tcm =>
Relevance -> tcm a -> tcm a
applyRelevanceToContextOnly Relevance
rel) forall a b. (a -> b) -> a -> b
$    -- enable local irr. defs only when option
      forall (tcm :: * -> *) a.
MonadTCEnv tcm =>
Relevance -> tcm a -> tcm a
applyRelevanceToJudgementOnly Relevance
rel tcm a
cont -- enable global irr. defs alway

-- | Sets the current quantity (unless the given quantity is 1).
applyQuantityToContext :: (MonadTCEnv tcm, LensQuantity q) => q -> tcm a -> tcm a
applyQuantityToContext :: forall (tcm :: * -> *) q a.
(MonadTCEnv tcm, LensQuantity q) =>
q -> tcm a -> tcm a
applyQuantityToContext q
thing =
  case forall a. LensQuantity a => a -> Quantity
getQuantity q
thing of
    Quantity1{} -> forall a. a -> a
id
    Quantity
q           -> forall (tcm :: * -> *) a.
MonadTCEnv tcm =>
Quantity -> tcm a -> tcm a
applyQuantityToJudgementOnly Quantity
q

-- | Apply quantity @q@ the the quantity annotation of the (typing/equality)
--   judgement.  This is part of the work done when going into a @q@-context.
--
--   Precondition: @Quantity /= Quantity1@
applyQuantityToJudgementOnly :: (MonadTCEnv tcm) => Quantity -> tcm a -> tcm a
applyQuantityToJudgementOnly :: forall (tcm :: * -> *) a.
MonadTCEnv tcm =>
Quantity -> tcm a -> tcm a
applyQuantityToJudgementOnly = forall (m :: * -> *) a.
MonadTCEnv m =>
(TCEnv -> TCEnv) -> m a -> m a
localTC forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall i o. Lens' i o -> LensMap i o
over Lens' Quantity TCEnv
eQuantity forall b c a. (b -> c) -> (a -> b) -> a -> c
. Quantity -> Quantity -> Quantity
composeQuantity

-- | Apply inverse composition with the given cohesion to the typing context.
applyCohesionToContext :: (MonadTCEnv tcm, LensCohesion m) => m -> tcm a -> tcm a
applyCohesionToContext :: forall (tcm :: * -> *) m a.
(MonadTCEnv tcm, LensCohesion m) =>
m -> tcm a -> tcm a
applyCohesionToContext m
thing =
  case forall a. LensCohesion a => a -> Cohesion
getCohesion m
thing of
    Cohesion
m | Cohesion
m forall a. Eq a => a -> a -> Bool
== Cohesion
unitCohesion -> forall a. a -> a
id
      | Bool
otherwise         -> forall (tcm :: * -> *) a.
MonadTCEnv tcm =>
Cohesion -> tcm a -> tcm a
applyCohesionToContextOnly   Cohesion
m
                             -- Cohesion does not apply to the judgment.

applyCohesionToContextOnly :: (MonadTCEnv tcm) => Cohesion -> tcm a -> tcm a
applyCohesionToContextOnly :: forall (tcm :: * -> *) a.
MonadTCEnv tcm =>
Cohesion -> tcm a -> tcm a
applyCohesionToContextOnly Cohesion
q = forall (m :: * -> *) a.
MonadTCEnv m =>
(TCEnv -> TCEnv) -> m a -> m a
localTC
  forall a b. (a -> b) -> a -> b
$ forall i o. Lens' i o -> LensMap i o
over Lens' Context TCEnv
eContext     (forall a b. (a -> b) -> [a] -> [b]
map forall a b. (a -> b) -> a -> b
$ forall a. LensCohesion a => Cohesion -> a -> a
inverseApplyCohesion Cohesion
q)
  forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall i o. Lens' i o -> LensMap i o
over Lens' LetBindings TCEnv
eLetBindings (forall a b k. (a -> b) -> Map k a -> Map k b
Map.map forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
fmap forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall (a :: * -> * -> *) b c d.
Arrow a =>
a b c -> a (d, b) (d, c)
second forall a b. (a -> b) -> a -> b
$ forall a. LensCohesion a => Cohesion -> a -> a
inverseApplyCohesion Cohesion
q)

-- | Can we split on arguments of the given cohesion?
splittableCohesion :: (HasOptions m, LensCohesion a) => a -> m Bool
splittableCohesion :: forall (m :: * -> *) a.
(HasOptions m, LensCohesion a) =>
a -> m Bool
splittableCohesion a
a = do
  let c :: Cohesion
c = forall a. LensCohesion a => a -> Cohesion
getCohesion a
a
  forall (f :: * -> *) a. Applicative f => a -> f a
pure (forall a. LensCohesion a => a -> Bool
usableCohesion Cohesion
c) forall (m :: * -> *). Monad m => m Bool -> m Bool -> m Bool
`and2M` (forall (f :: * -> *) a. Applicative f => a -> f a
pure (Cohesion
c forall a. Eq a => a -> a -> Bool
/= Cohesion
Flat) forall (m :: * -> *). Monad m => m Bool -> m Bool -> m Bool
`or2M` do forall (b :: Bool). KnownBool b => WithDefault b -> Bool
collapseDefault forall b c a. (b -> c) -> (a -> b) -> a -> c
. PragmaOptions -> WithDefault 'False
optFlatSplit forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> forall (m :: * -> *). HasOptions m => m PragmaOptions
pragmaOptions)

-- | (Conditionally) wake up irrelevant variables and make them relevant.
--   For instance,
--   in an irrelevant function argument otherwise irrelevant variables
--   may be used, so they are awoken before type checking the argument.
--
--   Also allow the use of irrelevant definitions.
--
--   This function might also do something for other modalities.
applyModalityToContext :: (MonadTCEnv tcm, LensModality m) => m -> tcm a -> tcm a
applyModalityToContext :: forall (tcm :: * -> *) m a.
(MonadTCEnv tcm, LensModality m) =>
m -> tcm a -> tcm a
applyModalityToContext m
thing =
  case forall a. LensModality a => a -> Modality
getModality m
thing of
    Modality
m | Modality
m forall a. Eq a => a -> a -> Bool
== Modality
unitModality -> forall a. a -> a
id
      | Bool
otherwise         -> forall (tcm :: * -> *) a.
MonadTCEnv tcm =>
Modality -> tcm a -> tcm a
applyModalityToContextOnly   Modality
m
                           forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall (tcm :: * -> *) a.
MonadTCEnv tcm =>
Modality -> tcm a -> tcm a
applyModalityToJudgementOnly Modality
m

-- | (Conditionally) wake up irrelevant variables and make them relevant.
--   For instance,
--   in an irrelevant function argument otherwise irrelevant variables
--   may be used, so they are awoken before type checking the
--   argument.
--
--   This function might also do something for other modalities, but
--   not for quantities.
--
--   Precondition: @Modality /= Relevant@
applyModalityToContextOnly :: (MonadTCEnv tcm) => Modality -> tcm a -> tcm a
applyModalityToContextOnly :: forall (tcm :: * -> *) a.
MonadTCEnv tcm =>
Modality -> tcm a -> tcm a
applyModalityToContextOnly Modality
m = forall (m :: * -> *) a.
MonadTCEnv m =>
(TCEnv -> TCEnv) -> m a -> m a
localTC
  forall a b. (a -> b) -> a -> b
$ forall i o. Lens' i o -> LensMap i o
over Lens' Context TCEnv
eContext (forall a b. (a -> b) -> [a] -> [b]
map forall a b. (a -> b) -> a -> b
$ forall a. LensModality a => Modality -> a -> a
inverseApplyModalityButNotQuantity Modality
m)
  forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall i o. Lens' i o -> LensMap i o
over Lens' LetBindings TCEnv
eLetBindings
      (forall a b k. (a -> b) -> Map k a -> Map k b
Map.map forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
fmap forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall (a :: * -> * -> *) b c d.
Arrow a =>
a b c -> a (d, b) (d, c)
second forall a b. (a -> b) -> a -> b
$ forall a. LensModality a => Modality -> a -> a
inverseApplyModalityButNotQuantity Modality
m)

-- | Apply modality @m@ the the modality annotation of the (typing/equality)
--   judgement.  This is part of the work done when going into a @m@-context.
--
--   Precondition: @Modality /= Relevant@
applyModalityToJudgementOnly :: (MonadTCEnv tcm) => Modality -> tcm a -> tcm a
applyModalityToJudgementOnly :: forall (tcm :: * -> *) a.
MonadTCEnv tcm =>
Modality -> tcm a -> tcm a
applyModalityToJudgementOnly = forall (m :: * -> *) a.
MonadTCEnv m =>
(TCEnv -> TCEnv) -> m a -> m a
localTC forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall i o. Lens' i o -> LensMap i o
over Lens' Modality TCEnv
eModality forall b c a. (b -> c) -> (a -> b) -> a -> c
. Modality -> Modality -> Modality
composeModality

-- | Like 'applyModalityToContext', but only act on context (for Relevance) if
--   @--irrelevant-projections@.
--   See issue #2170.
applyModalityToContextFunBody :: (MonadTCM tcm, LensModality r) => r -> tcm a -> tcm a
applyModalityToContextFunBody :: forall (tcm :: * -> *) r a.
(MonadTCM tcm, LensModality r) =>
r -> tcm a -> tcm a
applyModalityToContextFunBody r
thing tcm a
cont = do
    forall (m :: * -> *) a. Monad m => m Bool -> m a -> m a -> m a
ifM (PragmaOptions -> Bool
optIrrelevantProjections forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> forall (m :: * -> *). HasOptions m => m PragmaOptions
pragmaOptions)
      {-then-} (forall (tcm :: * -> *) m a.
(MonadTCEnv tcm, LensModality m) =>
m -> tcm a -> tcm a
applyModalityToContext Modality
m tcm a
cont)                -- enable global irr. defs always
      {-else-} (forall (tcm :: * -> *) r a.
(MonadTCM tcm, LensRelevance r) =>
r -> tcm a -> tcm a
applyRelevanceToContextFunBody (forall a. LensRelevance a => a -> Relevance
getRelevance Modality
m)
               forall a b. (a -> b) -> a -> b
$ forall (tcm :: * -> *) m a.
(MonadTCEnv tcm, LensCohesion m) =>
m -> tcm a -> tcm a
applyCohesionToContext (forall a. LensCohesion a => a -> Cohesion
getCohesion Modality
m)
               forall a b. (a -> b) -> a -> b
$ forall (tcm :: * -> *) q a.
(MonadTCEnv tcm, LensQuantity q) =>
q -> tcm a -> tcm a
applyQuantityToContext (forall a. LensQuantity a => a -> Quantity
getQuantity Modality
m) tcm a
cont) -- enable local irr. defs only when option
  where
    m :: Modality
m = forall a. LensModality a => a -> Modality
getModality r
thing

-- | Wake up irrelevant variables and make them relevant. This is used
--   when type checking terms in a hole, in which case you want to be able to
--   (for instance) infer the type of an irrelevant variable. In the course
--   of type checking an irrelevant function argument 'applyRelevanceToContext'
--   is used instead, which also sets the context relevance to 'Irrelevant'.
--   This is not the right thing to do when type checking interactively in a
--   hole since it also marks all metas created during type checking as
--   irrelevant (issue #2568).
--
--   Also set the current quantity to 0.
wakeIrrelevantVars :: (MonadTCEnv tcm) => tcm a -> tcm a
wakeIrrelevantVars :: forall (m :: * -> *) a. MonadTCEnv m => m a -> m a
wakeIrrelevantVars
  = forall (tcm :: * -> *) a.
MonadTCEnv tcm =>
Relevance -> tcm a -> tcm a
applyRelevanceToContextOnly Relevance
Irrelevant
  forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall (tcm :: * -> *) a.
MonadTCEnv tcm =>
Quantity -> tcm a -> tcm a
applyQuantityToJudgementOnly Quantity
zeroQuantity

-- | Check whether something can be used in a position of the given relevance.
--
--   This is a substitute for double-checking that only makes sure
--   relevances are correct.  See issue #2640.
--
--   Used in unifier (@ unifyStep Solution{}@).
--
--   At the moment, this implements McBride-style irrelevance,
--   where Pfenning-style would be the most accurate thing.
--   However, these two notions only differ how they handle
--   bound variables in a term.  Here, we are only concerned
--   in the free variables, used meta-variables, and used
--   (irrelevant) definitions.
--
class UsableRelevance a where
  usableRel
    :: (ReadTCState m, HasConstInfo m, MonadTCEnv m, MonadAddContext m, MonadDebug m)
    => Relevance -> a -> m Bool

instance UsableRelevance Term where
  usableRel :: forall (m :: * -> *).
(ReadTCState m, HasConstInfo m, MonadTCEnv m, MonadAddContext m,
 MonadDebug m) =>
Relevance -> Term -> m Bool
usableRel Relevance
rel = \case
    Var VerboseLevel
i Elims
vs -> do
      Relevance
irel <- forall a. LensRelevance a => a -> Relevance
getRelevance forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> forall (m :: * -> *).
(Applicative m, MonadFail m, MonadTCEnv m) =>
VerboseLevel -> m (Dom Type)
domOfBV VerboseLevel
i
      let ok :: Bool
ok = Relevance
irel Relevance -> Relevance -> Bool
`moreRelevant` Relevance
rel
      forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> VerboseLevel -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"tc.irr" VerboseLevel
50 forall a b. (a -> b) -> a -> b
$
        TCMT IO Doc
"Variable" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM (VerboseLevel -> Term
var VerboseLevel
i) forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+>
        forall (m :: * -> *). Applicative m => VerboseKey -> m Doc
text (VerboseKey
"has relevance " forall a. [a] -> [a] -> [a]
++ forall a. Show a => a -> VerboseKey
show Relevance
irel forall a. [a] -> [a] -> [a]
++ VerboseKey
", which is " forall a. [a] -> [a] -> [a]
++
              (if Bool
ok then VerboseKey
"" else VerboseKey
"NOT ") forall a. [a] -> [a] -> [a]
++ VerboseKey
"more relevant than " forall a. [a] -> [a] -> [a]
++ forall a. Show a => a -> VerboseKey
show Relevance
rel)
      forall (m :: * -> *) a. Monad m => a -> m a
return Bool
ok forall (m :: * -> *). Monad m => m Bool -> m Bool -> m Bool
`and2M` forall a (m :: * -> *).
(UsableRelevance a, ReadTCState m, HasConstInfo m, MonadTCEnv m,
 MonadAddContext m, MonadDebug m) =>
Relevance -> a -> m Bool
usableRel Relevance
rel Elims
vs
    Def QName
f Elims
vs -> do
      Relevance
frel <- forall (m :: * -> *). HasConstInfo m => QName -> m Relevance
relOfConst QName
f
      forall (m :: * -> *) a. Monad m => a -> m a
return (Relevance
frel Relevance -> Relevance -> Bool
`moreRelevant` Relevance
rel) forall (m :: * -> *). Monad m => m Bool -> m Bool -> m Bool
`and2M` forall a (m :: * -> *).
(UsableRelevance a, ReadTCState m, HasConstInfo m, MonadTCEnv m,
 MonadAddContext m, MonadDebug m) =>
Relevance -> a -> m Bool
usableRel Relevance
rel Elims
vs
    Con ConHead
c ConInfo
_ Elims
vs -> forall a (m :: * -> *).
(UsableRelevance a, ReadTCState m, HasConstInfo m, MonadTCEnv m,
 MonadAddContext m, MonadDebug m) =>
Relevance -> a -> m Bool
usableRel Relevance
rel Elims
vs
    Lit Literal
l    -> forall (m :: * -> *) a. Monad m => a -> m a
return Bool
True
    Lam ArgInfo
_ Abs Term
v  -> forall a (m :: * -> *).
(UsableRelevance a, ReadTCState m, HasConstInfo m, MonadTCEnv m,
 MonadAddContext m, MonadDebug m) =>
Relevance -> a -> m Bool
usableRel Relevance
rel Abs Term
v
    Pi Dom Type
a Abs Type
b   -> forall a (m :: * -> *).
(UsableRelevance a, ReadTCState m, HasConstInfo m, MonadTCEnv m,
 MonadAddContext m, MonadDebug m) =>
Relevance -> a -> m Bool
usableRel Relevance
rel (Dom Type
a,Abs Type
b)
    Sort Sort
s   -> forall a (m :: * -> *).
(UsableRelevance a, ReadTCState m, HasConstInfo m, MonadTCEnv m,
 MonadAddContext m, MonadDebug m) =>
Relevance -> a -> m Bool
usableRel Relevance
rel Sort
s
    Level Level
l  -> forall (m :: * -> *) a. Monad m => a -> m a
return Bool
True
    MetaV MetaId
m Elims
vs -> do
      Relevance
mrel <- forall a. LensRelevance a => a -> Relevance
getRelevance forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> forall (m :: * -> *). ReadTCState m => MetaId -> m Modality
lookupMetaModality MetaId
m
      forall (m :: * -> *) a. Monad m => a -> m a
return (Relevance
mrel Relevance -> Relevance -> Bool
`moreRelevant` Relevance
rel) forall (m :: * -> *). Monad m => m Bool -> m Bool -> m Bool
`and2M` forall a (m :: * -> *).
(UsableRelevance a, ReadTCState m, HasConstInfo m, MonadTCEnv m,
 MonadAddContext m, MonadDebug m) =>
Relevance -> a -> m Bool
usableRel Relevance
rel Elims
vs
    DontCare Term
v -> forall a (m :: * -> *).
(UsableRelevance a, ReadTCState m, HasConstInfo m, MonadTCEnv m,
 MonadAddContext m, MonadDebug m) =>
Relevance -> a -> m Bool
usableRel Relevance
rel Term
v -- TODO: allow irrelevant things to be used in DontCare position?
    Dummy{}  -> forall (m :: * -> *) a. Monad m => a -> m a
return Bool
True

instance UsableRelevance a => UsableRelevance (Type' a) where
  usableRel :: forall (m :: * -> *).
(ReadTCState m, HasConstInfo m, MonadTCEnv m, MonadAddContext m,
 MonadDebug m) =>
Relevance -> Type' a -> m Bool
usableRel Relevance
rel (El Sort
_ a
t) = forall a (m :: * -> *).
(UsableRelevance a, ReadTCState m, HasConstInfo m, MonadTCEnv m,
 MonadAddContext m, MonadDebug m) =>
Relevance -> a -> m Bool
usableRel Relevance
rel a
t

instance UsableRelevance Sort where
  usableRel :: forall (m :: * -> *).
(ReadTCState m, HasConstInfo m, MonadTCEnv m, MonadAddContext m,
 MonadDebug m) =>
Relevance -> Sort -> m Bool
usableRel Relevance
rel = \case
    Type Level
l -> forall a (m :: * -> *).
(UsableRelevance a, ReadTCState m, HasConstInfo m, MonadTCEnv m,
 MonadAddContext m, MonadDebug m) =>
Relevance -> a -> m Bool
usableRel Relevance
rel Level
l
    Prop Level
l -> forall a (m :: * -> *).
(UsableRelevance a, ReadTCState m, HasConstInfo m, MonadTCEnv m,
 MonadAddContext m, MonadDebug m) =>
Relevance -> a -> m Bool
usableRel Relevance
rel Level
l
    Inf IsFibrant
f Integer
n -> forall (m :: * -> *) a. Monad m => a -> m a
return Bool
True
    SSet Level
l -> forall a (m :: * -> *).
(UsableRelevance a, ReadTCState m, HasConstInfo m, MonadTCEnv m,
 MonadAddContext m, MonadDebug m) =>
Relevance -> a -> m Bool
usableRel Relevance
rel Level
l
    Sort
SizeUniv -> forall (m :: * -> *) a. Monad m => a -> m a
return Bool
True
    Sort
LockUniv -> forall (m :: * -> *) a. Monad m => a -> m a
return Bool
True
    Sort
IntervalUniv -> forall (m :: * -> *) a. Monad m => a -> m a
return Bool
True
    PiSort Dom' Term Term
a Sort
s1 Abs Sort
s2 -> forall a (m :: * -> *).
(UsableRelevance a, ReadTCState m, HasConstInfo m, MonadTCEnv m,
 MonadAddContext m, MonadDebug m) =>
Relevance -> a -> m Bool
usableRel Relevance
rel (Dom' Term Term
a,Sort
s1,Abs Sort
s2)
    FunSort Sort
s1 Sort
s2 -> forall a (m :: * -> *).
(UsableRelevance a, ReadTCState m, HasConstInfo m, MonadTCEnv m,
 MonadAddContext m, MonadDebug m) =>
Relevance -> a -> m Bool
usableRel Relevance
rel (Sort
s1,Sort
s2)
    UnivSort Sort
s -> forall a (m :: * -> *).
(UsableRelevance a, ReadTCState m, HasConstInfo m, MonadTCEnv m,
 MonadAddContext m, MonadDebug m) =>
Relevance -> a -> m Bool
usableRel Relevance
rel Sort
s
    MetaS MetaId
x Elims
es -> forall a (m :: * -> *).
(UsableRelevance a, ReadTCState m, HasConstInfo m, MonadTCEnv m,
 MonadAddContext m, MonadDebug m) =>
Relevance -> a -> m Bool
usableRel Relevance
rel Elims
es
    DefS QName
d Elims
es  -> forall a (m :: * -> *).
(UsableRelevance a, ReadTCState m, HasConstInfo m, MonadTCEnv m,
 MonadAddContext m, MonadDebug m) =>
Relevance -> a -> m Bool
usableRel Relevance
rel forall a b. (a -> b) -> a -> b
$ QName -> Elims -> Term
Def QName
d Elims
es
    DummyS{} -> forall (m :: * -> *) a. Monad m => a -> m a
return Bool
True

instance UsableRelevance Level where
  usableRel :: forall (m :: * -> *).
(ReadTCState m, HasConstInfo m, MonadTCEnv m, MonadAddContext m,
 MonadDebug m) =>
Relevance -> Level -> m Bool
usableRel Relevance
rel (Max Integer
_ [PlusLevel' Term]
ls) = forall a (m :: * -> *).
(UsableRelevance a, ReadTCState m, HasConstInfo m, MonadTCEnv m,
 MonadAddContext m, MonadDebug m) =>
Relevance -> a -> m Bool
usableRel Relevance
rel [PlusLevel' Term]
ls

instance UsableRelevance PlusLevel where
  usableRel :: forall (m :: * -> *).
(ReadTCState m, HasConstInfo m, MonadTCEnv m, MonadAddContext m,
 MonadDebug m) =>
Relevance -> PlusLevel' Term -> m Bool
usableRel Relevance
rel (Plus Integer
_ Term
l) = forall a (m :: * -> *).
(UsableRelevance a, ReadTCState m, HasConstInfo m, MonadTCEnv m,
 MonadAddContext m, MonadDebug m) =>
Relevance -> a -> m Bool
usableRel Relevance
rel Term
l

instance UsableRelevance a => UsableRelevance [a] where
  usableRel :: forall (m :: * -> *).
(ReadTCState m, HasConstInfo m, MonadTCEnv m, MonadAddContext m,
 MonadDebug m) =>
Relevance -> [a] -> m Bool
usableRel Relevance
rel = forall (f :: * -> *) (m :: * -> *).
(Foldable f, Monad m) =>
f (m Bool) -> m Bool
andM forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall a b. (a -> b) -> [a] -> [b]
map (forall a (m :: * -> *).
(UsableRelevance a, ReadTCState m, HasConstInfo m, MonadTCEnv m,
 MonadAddContext m, MonadDebug m) =>
Relevance -> a -> m Bool
usableRel Relevance
rel)

instance (UsableRelevance a, UsableRelevance b) => UsableRelevance (a,b) where
  usableRel :: forall (m :: * -> *).
(ReadTCState m, HasConstInfo m, MonadTCEnv m, MonadAddContext m,
 MonadDebug m) =>
Relevance -> (a, b) -> m Bool
usableRel Relevance
rel (a
a,b
b) = forall a (m :: * -> *).
(UsableRelevance a, ReadTCState m, HasConstInfo m, MonadTCEnv m,
 MonadAddContext m, MonadDebug m) =>
Relevance -> a -> m Bool
usableRel Relevance
rel a
a forall (m :: * -> *). Monad m => m Bool -> m Bool -> m Bool
`and2M` forall a (m :: * -> *).
(UsableRelevance a, ReadTCState m, HasConstInfo m, MonadTCEnv m,
 MonadAddContext m, MonadDebug m) =>
Relevance -> a -> m Bool
usableRel Relevance
rel b
b

instance (UsableRelevance a, UsableRelevance b, UsableRelevance c) => UsableRelevance (a,b,c) where
  usableRel :: forall (m :: * -> *).
(ReadTCState m, HasConstInfo m, MonadTCEnv m, MonadAddContext m,
 MonadDebug m) =>
Relevance -> (a, b, c) -> m Bool
usableRel Relevance
rel (a
a,b
b,c
c) = forall a (m :: * -> *).
(UsableRelevance a, ReadTCState m, HasConstInfo m, MonadTCEnv m,
 MonadAddContext m, MonadDebug m) =>
Relevance -> a -> m Bool
usableRel Relevance
rel a
a forall (m :: * -> *). Monad m => m Bool -> m Bool -> m Bool
`and2M` forall a (m :: * -> *).
(UsableRelevance a, ReadTCState m, HasConstInfo m, MonadTCEnv m,
 MonadAddContext m, MonadDebug m) =>
Relevance -> a -> m Bool
usableRel Relevance
rel b
b forall (m :: * -> *). Monad m => m Bool -> m Bool -> m Bool
`and2M` forall a (m :: * -> *).
(UsableRelevance a, ReadTCState m, HasConstInfo m, MonadTCEnv m,
 MonadAddContext m, MonadDebug m) =>
Relevance -> a -> m Bool
usableRel Relevance
rel c
c

instance UsableRelevance a => UsableRelevance (Elim' a) where
  usableRel :: forall (m :: * -> *).
(ReadTCState m, HasConstInfo m, MonadTCEnv m, MonadAddContext m,
 MonadDebug m) =>
Relevance -> Elim' a -> m Bool
usableRel Relevance
rel (Apply Arg a
a) = forall a (m :: * -> *).
(UsableRelevance a, ReadTCState m, HasConstInfo m, MonadTCEnv m,
 MonadAddContext m, MonadDebug m) =>
Relevance -> a -> m Bool
usableRel Relevance
rel Arg a
a
  usableRel Relevance
rel (Proj ProjOrigin
_ QName
p) = do
    Relevance
prel <- forall (m :: * -> *). HasConstInfo m => QName -> m Relevance
relOfConst QName
p
    forall (m :: * -> *) a. Monad m => a -> m a
return forall a b. (a -> b) -> a -> b
$ Relevance
prel Relevance -> Relevance -> Bool
`moreRelevant` Relevance
rel
  usableRel Relevance
rel (IApply a
x a
y a
v) = forall a (m :: * -> *).
(UsableRelevance a, ReadTCState m, HasConstInfo m, MonadTCEnv m,
 MonadAddContext m, MonadDebug m) =>
Relevance -> a -> m Bool
usableRel Relevance
rel a
v

instance UsableRelevance a => UsableRelevance (Arg a) where
  usableRel :: forall (m :: * -> *).
(ReadTCState m, HasConstInfo m, MonadTCEnv m, MonadAddContext m,
 MonadDebug m) =>
Relevance -> Arg a -> m Bool
usableRel Relevance
rel (Arg ArgInfo
info a
u) =
    let rel' :: Relevance
rel' = forall a. LensRelevance a => a -> Relevance
getRelevance ArgInfo
info
    in  forall a (m :: * -> *).
(UsableRelevance a, ReadTCState m, HasConstInfo m, MonadTCEnv m,
 MonadAddContext m, MonadDebug m) =>
Relevance -> a -> m Bool
usableRel (Relevance
rel Relevance -> Relevance -> Relevance
`composeRelevance` Relevance
rel') a
u

instance UsableRelevance a => UsableRelevance (Dom a) where
  usableRel :: forall (m :: * -> *).
(ReadTCState m, HasConstInfo m, MonadTCEnv m, MonadAddContext m,
 MonadDebug m) =>
Relevance -> Dom a -> m Bool
usableRel Relevance
rel Dom{unDom :: forall t e. Dom' t e -> e
unDom = a
u} = forall a (m :: * -> *).
(UsableRelevance a, ReadTCState m, HasConstInfo m, MonadTCEnv m,
 MonadAddContext m, MonadDebug m) =>
Relevance -> a -> m Bool
usableRel Relevance
rel a
u

instance (Subst a, UsableRelevance a) => UsableRelevance (Abs a) where
  usableRel :: forall (m :: * -> *).
(ReadTCState m, HasConstInfo m, MonadTCEnv m, MonadAddContext m,
 MonadDebug m) =>
Relevance -> Abs a -> m Bool
usableRel Relevance
rel Abs a
abs = forall a (m :: * -> *) b.
(Subst a, MonadAddContext m) =>
Abs a -> (a -> m b) -> m b
underAbstraction_ Abs a
abs forall a b. (a -> b) -> a -> b
$ \a
u -> forall a (m :: * -> *).
(UsableRelevance a, ReadTCState m, HasConstInfo m, MonadTCEnv m,
 MonadAddContext m, MonadDebug m) =>
Relevance -> a -> m Bool
usableRel Relevance
rel a
u

-- | Check whether something can be used in a position of the given modality.
--
--   This is a substitute for double-checking that only makes sure
--   modalities are correct.  See issue #2640.
--
--   Used in unifier (@ unifyStep Solution{}@).
--
--   This uses McBride-style modality checking.
--   It does not differ from Pfenning-style if we
--   are only interested in the modality of the
--   free variables, used meta-variables, and used
--   definitions.
--
class UsableModality a where
  usableMod
    :: (ReadTCState m, HasConstInfo m, MonadTCEnv m, MonadAddContext m, MonadDebug m, MonadReduce m, MonadError Blocker m)
    => Modality -> a -> m Bool

instance UsableModality Term where
  usableMod :: forall (m :: * -> *).
(ReadTCState m, HasConstInfo m, MonadTCEnv m, MonadAddContext m,
 MonadDebug m, MonadReduce m, MonadError Blocker m) =>
Modality -> Term -> m Bool
usableMod Modality
mod Term
u = do
   case Term
u of
    Var VerboseLevel
i Elims
vs -> do
      Modality
imod <- forall a. LensModality a => a -> Modality
getModality forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> forall (m :: * -> *).
(Applicative m, MonadFail m, MonadTCEnv m) =>
VerboseLevel -> m (Dom Type)
domOfBV VerboseLevel
i
      let ok :: Bool
ok = Modality
imod Modality -> Modality -> Bool
`moreUsableModality` Modality
mod
      forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> VerboseLevel -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"tc.irr" VerboseLevel
50 forall a b. (a -> b) -> a -> b
$
        TCMT IO Doc
"Variable" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM (VerboseLevel -> Term
var VerboseLevel
i) forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+>
        forall (m :: * -> *). Applicative m => VerboseKey -> m Doc
text (VerboseKey
"has modality " forall a. [a] -> [a] -> [a]
++ forall a. Show a => a -> VerboseKey
show Modality
imod forall a. [a] -> [a] -> [a]
++ VerboseKey
", which is a " forall a. [a] -> [a] -> [a]
++
              (if Bool
ok then VerboseKey
"" else VerboseKey
"NOT ") forall a. [a] -> [a] -> [a]
++ VerboseKey
"more usable modality than " forall a. [a] -> [a] -> [a]
++ forall a. Show a => a -> VerboseKey
show Modality
mod)
      forall (m :: * -> *) a. Monad m => a -> m a
return Bool
ok forall (m :: * -> *). Monad m => m Bool -> m Bool -> m Bool
`and2M` forall a (m :: * -> *).
(UsableModality a, ReadTCState m, HasConstInfo m, MonadTCEnv m,
 MonadAddContext m, MonadDebug m, MonadReduce m,
 MonadError Blocker m) =>
Modality -> a -> m Bool
usableMod Modality
mod Elims
vs
    Def QName
f Elims
vs -> do
      Modality
fmod <- forall (m :: * -> *). HasConstInfo m => QName -> m Modality
modalityOfConst QName
f
      let ok :: Bool
ok = Modality
fmod Modality -> Modality -> Bool
`moreUsableModality` Modality
mod
      forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> VerboseLevel -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"tc.irr" VerboseLevel
50 forall a b. (a -> b) -> a -> b
$
        TCMT IO Doc
"Definition" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM (QName -> Elims -> Term
Def QName
f []) forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+>
        forall (m :: * -> *). Applicative m => VerboseKey -> m Doc
text (VerboseKey
"has modality " forall a. [a] -> [a] -> [a]
++ forall a. Show a => a -> VerboseKey
show Modality
fmod forall a. [a] -> [a] -> [a]
++ VerboseKey
", which is a " forall a. [a] -> [a] -> [a]
++
              (if Bool
ok then VerboseKey
"" else VerboseKey
"NOT ") forall a. [a] -> [a] -> [a]
++ VerboseKey
"more usable modality than " forall a. [a] -> [a] -> [a]
++ forall a. Show a => a -> VerboseKey
show Modality
mod)
      forall (m :: * -> *) a. Monad m => a -> m a
return Bool
ok forall (m :: * -> *). Monad m => m Bool -> m Bool -> m Bool
`and2M` forall a (m :: * -> *).
(UsableModality a, ReadTCState m, HasConstInfo m, MonadTCEnv m,
 MonadAddContext m, MonadDebug m, MonadReduce m,
 MonadError Blocker m) =>
Modality -> a -> m Bool
usableMod Modality
mod Elims
vs
    Con ConHead
c ConInfo
o Elims
vs -> do
      Modality
cmod <- forall (m :: * -> *). HasConstInfo m => QName -> m Modality
modalityOfConst (ConHead -> QName
conName ConHead
c)
      let ok :: Bool
ok = Modality
cmod Modality -> Modality -> Bool
`moreUsableModality` Modality
mod
      forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> VerboseLevel -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"tc.irr" VerboseLevel
50 forall a b. (a -> b) -> a -> b
$
        TCMT IO Doc
"The constructor" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM (ConHead -> ConInfo -> Elims -> Term
Con ConHead
c ConInfo
o []) forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+>
        forall (m :: * -> *). Applicative m => VerboseKey -> m Doc
text (VerboseKey
"has the modality " forall a. [a] -> [a] -> [a]
++ forall a. Show a => a -> VerboseKey
show Modality
cmod forall a. [a] -> [a] -> [a]
++ VerboseKey
", which is " forall a. [a] -> [a] -> [a]
++
              (if Bool
ok then VerboseKey
"" else VerboseKey
"NOT ") forall a. [a] -> [a] -> [a]
++
              VerboseKey
"more usable than the modality " forall a. [a] -> [a] -> [a]
++ forall a. Show a => a -> VerboseKey
show Modality
mod forall a. [a] -> [a] -> [a]
++ VerboseKey
".")
      forall (m :: * -> *) a. Monad m => a -> m a
return Bool
ok forall (m :: * -> *). Monad m => m Bool -> m Bool -> m Bool
`and2M` forall a (m :: * -> *).
(UsableModality a, ReadTCState m, HasConstInfo m, MonadTCEnv m,
 MonadAddContext m, MonadDebug m, MonadReduce m,
 MonadError Blocker m) =>
Modality -> a -> m Bool
usableMod Modality
mod Elims
vs
    Lit Literal
l    -> forall (m :: * -> *) a. Monad m => a -> m a
return Bool
True
    Lam ArgInfo
info Abs Term
v  -> forall a (m :: * -> *).
(Subst a, MonadAddContext m, UsableModality a, ReadTCState m,
 HasConstInfo m, MonadReduce m, MonadError Blocker m) =>
ArgInfo -> Modality -> Abs a -> m Bool
usableModAbs ArgInfo
info Modality
mod Abs Term
v
    -- Even if Pi contains Type, here we check it as a constructor for terms in the universe.
    Pi Dom Type
a Abs Type
b   -> forall a (m :: * -> *).
(UsableModality a, ReadTCState m, HasConstInfo m, MonadTCEnv m,
 MonadAddContext m, MonadDebug m, MonadReduce m,
 MonadError Blocker m) =>
Modality -> a -> m Bool
usableMod Modality
domMod (forall t a. Type'' t a -> a
unEl forall a b. (a -> b) -> a -> b
$ forall t e. Dom' t e -> e
unDom Dom Type
a) forall (m :: * -> *). Monad m => m Bool -> m Bool -> m Bool
`and2M` forall a (m :: * -> *).
(Subst a, MonadAddContext m, UsableModality a, ReadTCState m,
 HasConstInfo m, MonadReduce m, MonadError Blocker m) =>
ArgInfo -> Modality -> Abs a -> m Bool
usableModAbs (forall a. LensArgInfo a => a -> ArgInfo
getArgInfo Dom Type
a) Modality
mod (forall t a. Type'' t a -> a
unEl forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> Abs Type
b)
      where
        domMod :: Modality
domMod = forall a. LensQuantity a => (Quantity -> Quantity) -> a -> a
mapQuantity (Quantity -> Quantity -> Quantity
composeQuantity forall a b. (a -> b) -> a -> b
$ forall a. LensQuantity a => a -> Quantity
getQuantity Dom Type
a) forall a b. (a -> b) -> a -> b
$
                 forall a. LensCohesion a => (Cohesion -> Cohesion) -> a -> a
mapCohesion (Cohesion -> Cohesion -> Cohesion
composeCohesion forall a b. (a -> b) -> a -> b
$ forall a. LensCohesion a => a -> Cohesion
getCohesion Dom Type
a) Modality
mod
    -- Andrea 15/10/2020 not updating these cases yet, but they are quite suspicious,
    -- do we have special typing rules for Sort and Level?
    Sort Sort
s   -> forall a (m :: * -> *).
(UsableModality a, ReadTCState m, HasConstInfo m, MonadTCEnv m,
 MonadAddContext m, MonadDebug m, MonadReduce m,
 MonadError Blocker m) =>
Modality -> a -> m Bool
usableMod Modality
mod Sort
s
    Level Level
l  -> forall (m :: * -> *) a. Monad m => a -> m a
return Bool
True
    MetaV MetaId
m Elims
vs -> do
      Modality
mmod <- forall (m :: * -> *). ReadTCState m => MetaId -> m Modality
lookupMetaModality MetaId
m
      let ok :: Bool
ok = Modality
mmod Modality -> Modality -> Bool
`moreUsableModality` Modality
mod
      forall (m :: * -> *).
MonadDebug m =>
VerboseKey -> VerboseLevel -> TCMT IO Doc -> m ()
reportSDoc VerboseKey
"tc.irr" VerboseLevel
50 forall a b. (a -> b) -> a -> b
$
        TCMT IO Doc
"Metavariable" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM (MetaId -> Elims -> Term
MetaV MetaId
m []) forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+>
        forall (m :: * -> *). Applicative m => VerboseKey -> m Doc
text (VerboseKey
"has modality " forall a. [a] -> [a] -> [a]
++ forall a. Show a => a -> VerboseKey
show Modality
mmod forall a. [a] -> [a] -> [a]
++ VerboseKey
", which is a " forall a. [a] -> [a] -> [a]
++
              (if Bool
ok then VerboseKey
"" else VerboseKey
"NOT ") forall a. [a] -> [a] -> [a]
++ VerboseKey
"more usable modality than " forall a. [a] -> [a] -> [a]
++ forall a. Show a => a -> VerboseKey
show Modality
mod)
      (forall (m :: * -> *) a. Monad m => a -> m a
return Bool
ok forall (m :: * -> *). Monad m => m Bool -> m Bool -> m Bool
`and2M` forall a (m :: * -> *).
(UsableModality a, ReadTCState m, HasConstInfo m, MonadTCEnv m,
 MonadAddContext m, MonadDebug m, MonadReduce m,
 MonadError Blocker m) =>
Modality -> a -> m Bool
usableMod Modality
mod Elims
vs) forall (m :: * -> *). Monad m => m Bool -> m Bool -> m Bool
`or2M` do
        Term
u <- forall a (m :: * -> *). (Instantiate a, MonadReduce m) => a -> m a
instantiate Term
u
        forall a b. Maybe a -> b -> (a -> b) -> b
caseMaybe (forall a. IsMeta a => a -> Maybe MetaId
isMeta Term
u) (forall a (m :: * -> *).
(UsableModality a, ReadTCState m, HasConstInfo m, MonadTCEnv m,
 MonadAddContext m, MonadDebug m, MonadReduce m,
 MonadError Blocker m) =>
Modality -> a -> m Bool
usableMod Modality
mod Term
u) forall a b. (a -> b) -> a -> b
$ \ MetaId
m -> forall e (m :: * -> *) a. MonadError e m => e -> m a
throwError (MetaId -> Blocker
UnblockOnMeta MetaId
m)
    DontCare Term
v -> forall a (m :: * -> *).
(UsableModality a, ReadTCState m, HasConstInfo m, MonadTCEnv m,
 MonadAddContext m, MonadDebug m, MonadReduce m,
 MonadError Blocker m) =>
Modality -> a -> m Bool
usableMod Modality
mod Term
v
    Dummy{}  -> forall (m :: * -> *) a. Monad m => a -> m a
return Bool
True

usableModAbs :: (Subst a, MonadAddContext m, UsableModality a,
                       ReadTCState m, HasConstInfo m, MonadReduce m, MonadError Blocker m) =>
                      ArgInfo -> Modality -> Abs a -> m Bool
usableModAbs :: forall a (m :: * -> *).
(Subst a, MonadAddContext m, UsableModality a, ReadTCState m,
 HasConstInfo m, MonadReduce m, MonadError Blocker m) =>
ArgInfo -> Modality -> Abs a -> m Bool
usableModAbs ArgInfo
info Modality
mod Abs a
abs = forall a (m :: * -> *) b.
(Subst a, MonadAddContext m) =>
Dom Type -> Abs a -> (a -> m b) -> m b
underAbstraction (forall a. LensArgInfo a => ArgInfo -> a -> a
setArgInfo ArgInfo
info forall a b. (a -> b) -> a -> b
$ HasCallStack => Dom Type
__DUMMY_DOM__) Abs a
abs forall a b. (a -> b) -> a -> b
$ \ a
u -> forall a (m :: * -> *).
(UsableModality a, ReadTCState m, HasConstInfo m, MonadTCEnv m,
 MonadAddContext m, MonadDebug m, MonadReduce m,
 MonadError Blocker m) =>
Modality -> a -> m Bool
usableMod Modality
mod a
u

instance UsableRelevance a => UsableModality (Type' a) where
  usableMod :: forall (m :: * -> *).
(ReadTCState m, HasConstInfo m, MonadTCEnv m, MonadAddContext m,
 MonadDebug m, MonadReduce m, MonadError Blocker m) =>
Modality -> Type' a -> m Bool
usableMod Modality
mod (El Sort
_ a
t) = forall a (m :: * -> *).
(UsableRelevance a, ReadTCState m, HasConstInfo m, MonadTCEnv m,
 MonadAddContext m, MonadDebug m) =>
Relevance -> a -> m Bool
usableRel (forall a. LensRelevance a => a -> Relevance
getRelevance Modality
mod) a
t

instance UsableModality Sort where
  usableMod :: forall (m :: * -> *).
(ReadTCState m, HasConstInfo m, MonadTCEnv m, MonadAddContext m,
 MonadDebug m, MonadReduce m, MonadError Blocker m) =>
Modality -> Sort -> m Bool
usableMod Modality
mod Sort
s = forall a (m :: * -> *).
(UsableRelevance a, ReadTCState m, HasConstInfo m, MonadTCEnv m,
 MonadAddContext m, MonadDebug m) =>
Relevance -> a -> m Bool
usableRel (forall a. LensRelevance a => a -> Relevance
getRelevance Modality
mod) Sort
s
  -- usableMod mod s = case s of
  --   Type l -> usableMod mod l
  --   Prop l -> usableMod mod l
  --   Inf    -> return True
  --   SizeUniv -> return True
  --   PiSort a s -> usableMod mod (a,s)
  --   UnivSort s -> usableMod mod s
  --   MetaS x es -> usableMod mod es
  --   DummyS{} -> return True

instance UsableModality Level where
  usableMod :: forall (m :: * -> *).
(ReadTCState m, HasConstInfo m, MonadTCEnv m, MonadAddContext m,
 MonadDebug m, MonadReduce m, MonadError Blocker m) =>
Modality -> Level -> m Bool
usableMod Modality
mod (Max Integer
_ [PlusLevel' Term]
ls) = forall a (m :: * -> *).
(UsableRelevance a, ReadTCState m, HasConstInfo m, MonadTCEnv m,
 MonadAddContext m, MonadDebug m) =>
Relevance -> a -> m Bool
usableRel (forall a. LensRelevance a => a -> Relevance
getRelevance Modality
mod) [PlusLevel' Term]
ls

-- instance UsableModality PlusLevel where
--   usableMod mod ClosedLevel{} = return True
--   usableMod mod (Plus _ l)    = usableMod mod l

instance UsableModality a => UsableModality [a] where
  usableMod :: forall (m :: * -> *).
(ReadTCState m, HasConstInfo m, MonadTCEnv m, MonadAddContext m,
 MonadDebug m, MonadReduce m, MonadError Blocker m) =>
Modality -> [a] -> m Bool
usableMod Modality
mod = forall (f :: * -> *) (m :: * -> *).
(Foldable f, Monad m) =>
f (m Bool) -> m Bool
andM forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall a b. (a -> b) -> [a] -> [b]
map (forall a (m :: * -> *).
(UsableModality a, ReadTCState m, HasConstInfo m, MonadTCEnv m,
 MonadAddContext m, MonadDebug m, MonadReduce m,
 MonadError Blocker m) =>
Modality -> a -> m Bool
usableMod Modality
mod)

instance (UsableModality a, UsableModality b) => UsableModality (a,b) where
  usableMod :: forall (m :: * -> *).
(ReadTCState m, HasConstInfo m, MonadTCEnv m, MonadAddContext m,
 MonadDebug m, MonadReduce m, MonadError Blocker m) =>
Modality -> (a, b) -> m Bool
usableMod Modality
mod (a
a,b
b) = forall a (m :: * -> *).
(UsableModality a, ReadTCState m, HasConstInfo m, MonadTCEnv m,
 MonadAddContext m, MonadDebug m, MonadReduce m,
 MonadError Blocker m) =>
Modality -> a -> m Bool
usableMod Modality
mod a
a forall (m :: * -> *). Monad m => m Bool -> m Bool -> m Bool
`and2M` forall a (m :: * -> *).
(UsableModality a, ReadTCState m, HasConstInfo m, MonadTCEnv m,
 MonadAddContext m, MonadDebug m, MonadReduce m,
 MonadError Blocker m) =>
Modality -> a -> m Bool
usableMod Modality
mod b
b

instance UsableModality a => UsableModality (Elim' a) where
  usableMod :: forall (m :: * -> *).
(ReadTCState m, HasConstInfo m, MonadTCEnv m, MonadAddContext m,
 MonadDebug m, MonadReduce m, MonadError Blocker m) =>
Modality -> Elim' a -> m Bool
usableMod Modality
mod (Apply Arg a
a) = forall a (m :: * -> *).
(UsableModality a, ReadTCState m, HasConstInfo m, MonadTCEnv m,
 MonadAddContext m, MonadDebug m, MonadReduce m,
 MonadError Blocker m) =>
Modality -> a -> m Bool
usableMod Modality
mod Arg a
a
  usableMod Modality
mod (Proj ProjOrigin
_ QName
p) = do
    Modality
pmod <- forall (m :: * -> *). HasConstInfo m => QName -> m Modality
modalityOfConst QName
p
    forall (m :: * -> *) a. Monad m => a -> m a
return forall a b. (a -> b) -> a -> b
$ Modality
pmod Modality -> Modality -> Bool
`moreUsableModality` Modality
mod
  usableMod Modality
mod (IApply a
x a
y a
v) = forall a (m :: * -> *).
(UsableModality a, ReadTCState m, HasConstInfo m, MonadTCEnv m,
 MonadAddContext m, MonadDebug m, MonadReduce m,
 MonadError Blocker m) =>
Modality -> a -> m Bool
usableMod Modality
mod a
v

instance UsableModality a => UsableModality (Arg a) where
  usableMod :: forall (m :: * -> *).
(ReadTCState m, HasConstInfo m, MonadTCEnv m, MonadAddContext m,
 MonadDebug m, MonadReduce m, MonadError Blocker m) =>
Modality -> Arg a -> m Bool
usableMod Modality
mod (Arg ArgInfo
info a
u) =
    let mod' :: Modality
mod' = forall a. LensModality a => a -> Modality
getModality ArgInfo
info
    in  forall a (m :: * -> *).
(UsableModality a, ReadTCState m, HasConstInfo m, MonadTCEnv m,
 MonadAddContext m, MonadDebug m, MonadReduce m,
 MonadError Blocker m) =>
Modality -> a -> m Bool
usableMod (Modality
mod Modality -> Modality -> Modality
`composeModality` Modality
mod') a
u

instance UsableModality a => UsableModality (Dom a) where
  usableMod :: forall (m :: * -> *).
(ReadTCState m, HasConstInfo m, MonadTCEnv m, MonadAddContext m,
 MonadDebug m, MonadReduce m, MonadError Blocker m) =>
Modality -> Dom a -> m Bool
usableMod Modality
mod Dom{unDom :: forall t e. Dom' t e -> e
unDom = a
u} = forall a (m :: * -> *).
(UsableModality a, ReadTCState m, HasConstInfo m, MonadTCEnv m,
 MonadAddContext m, MonadDebug m, MonadReduce m,
 MonadError Blocker m) =>
Modality -> a -> m Bool
usableMod Modality
mod a
u

usableAtModality'
  :: MonadConstraint TCM
  -- Note: This weird-looking constraint is to trick GHC into accepting
  -- that an instance of MonadConstraint TCM will exist, even if we
  -- can't import the module in which it is defined.
  => Maybe Sort -> WhyCheckModality -> Modality -> Term -> TCM ()
usableAtModality' :: MonadConstraint (TCMT IO) =>
Maybe Sort -> WhyCheckModality -> Modality -> Term -> TCM ()
usableAtModality' Maybe Sort
ms WhyCheckModality
why Modality
mod Term
t =
  forall (m :: * -> *).
MonadConstraint m =>
Constraint -> m () -> m ()
catchConstraint (WhyCheckModality -> Maybe Sort -> Modality -> Term -> Constraint
UsableAtModality WhyCheckModality
why Maybe Sort
ms Modality
mod Term
t) forall a b. (a -> b) -> a -> b
$ do
    forall (m :: * -> *). Monad m => m Bool -> m () -> m ()
whenM (forall b a. b -> (a -> b) -> Maybe a -> b
maybe (forall (f :: * -> *) a. Applicative f => a -> f a
pure Bool
True) forall a (m :: * -> *).
(LensSort a, PureTCM m, MonadBlock m) =>
a -> m Bool
isFibrant Maybe Sort
ms) forall a b. (a -> b) -> a -> b
$ do
      Either Blocker Bool
res <- forall e (m :: * -> *) a. ExceptT e m a -> m (Either e a)
runExceptT forall a b. (a -> b) -> a -> b
$ forall a (m :: * -> *).
(UsableModality a, ReadTCState m, HasConstInfo m, MonadTCEnv m,
 MonadAddContext m, MonadDebug m, MonadReduce m,
 MonadError Blocker m) =>
Modality -> a -> m Bool
usableMod Modality
mod Term
t
      case Either Blocker Bool
res of
        Right Bool
b -> forall (f :: * -> *). Applicative f => Bool -> f () -> f ()
unless Bool
b forall a b. (a -> b) -> a -> b
$
          forall (m :: * -> *) a.
(HasCallStack, MonadTCError m) =>
TypeError -> m a
typeError forall b c a. (b -> c) -> (a -> b) -> a -> c
. Doc -> TypeError
GenericDocError forall (m :: * -> *) a b. Monad m => (a -> m b) -> m a -> m b
=<< TCMT IO Doc
formatWhy
        Left Blocker
blocker -> forall (m :: * -> *) a. MonadBlock m => Blocker -> m a
patternViolation Blocker
blocker
  where
    formatWhy :: TCMT IO Doc
formatWhy = do
      Bool
compatible <- forall (b :: Bool). KnownBool b => WithDefault b -> Bool
collapseDefault forall b c a. (b -> c) -> (a -> b) -> a -> c
. PragmaOptions -> WithDefault 'False
optCubicalCompatible forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> forall (m :: * -> *). HasOptions m => m PragmaOptions
pragmaOptions
      Bool
cubical <- forall a. Maybe a -> Bool
isJust forall b c a. (b -> c) -> (a -> b) -> a -> c
. PragmaOptions -> Maybe Cubical
optCubical forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> forall (m :: * -> *). HasOptions m => m PragmaOptions
pragmaOptions
      let
        context :: VerboseKey
context
          | Bool
cubical    = VerboseKey
"in Cubical Agda,"
          | Bool
compatible = VerboseKey
"to maintain compatibility with Cubical Agda,"
          | Bool
otherwise  = VerboseKey
"when --without-K is enabled,"

        explanation :: VerboseKey -> [TCMT IO Doc]
explanation VerboseKey
what
          | Bool
cubical Bool -> Bool -> Bool
|| Bool
compatible =
            [ TCMT IO Doc
""
            , forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
fsep ( TCMT IO Doc
"Note:"forall a. a -> [a] -> [a]
:forall (m :: * -> *). Applicative m => VerboseKey -> [m Doc]
pwords VerboseKey
context
                  forall a. [a] -> [a] -> [a]
++ forall (m :: * -> *). Applicative m => VerboseKey -> [m Doc]
pwords VerboseKey
what forall a. [a] -> [a] -> [a]
++ forall (m :: * -> *). Applicative m => VerboseKey -> [m Doc]
pwords VerboseKey
"must be usable at the modality"
                  forall a. [a] -> [a] -> [a]
++ forall (m :: * -> *). Applicative m => VerboseKey -> [m Doc]
pwords VerboseKey
"in which the function was defined, since it will be"
                  forall a. [a] -> [a] -> [a]
++ forall (m :: * -> *). Applicative m => VerboseKey -> [m Doc]
pwords VerboseKey
"used for computing transports"
                  )
            , TCMT IO Doc
""
            ]
          | Bool
otherwise = []

      case WhyCheckModality
why of
        WhyCheckModality
IndexedClause ->
          forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
vcat forall a b. (a -> b) -> a -> b
$
            ( forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
fsep ( forall (m :: * -> *). Applicative m => VerboseKey -> [m Doc]
pwords VerboseKey
"This clause has target type"
                  forall a. [a] -> [a] -> [a]
++ [forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Term
t]
                  forall a. [a] -> [a] -> [a]
++ forall (m :: * -> *). Applicative m => VerboseKey -> [m Doc]
pwords VerboseKey
"which is not usable at the required modality"
                  forall a. [a] -> [a] -> [a]
++ [forall (f :: * -> *) a. Applicative f => a -> f a
pure (Modality -> Doc
attributesForModality Modality
mod) forall a. Semigroup a => a -> a -> a
<> TCMT IO Doc
"."]
                   )
            forall a. a -> [a] -> [a]
: VerboseKey -> [TCMT IO Doc]
explanation VerboseKey
"the target type")

        -- Arguments sometimes need to be transported too:
        IndexedClauseArg Name
forced Name
the_arg ->
          forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
vcat forall a b. (a -> b) -> a -> b
$
            ( forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
fsep (forall (m :: * -> *). Applicative m => VerboseKey -> [m Doc]
pwords VerboseKey
"The argument" forall a. [a] -> [a] -> [a]
++ [forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Name
the_arg] forall a. [a] -> [a] -> [a]
++ forall (m :: * -> *). Applicative m => VerboseKey -> [m Doc]
pwords VerboseKey
"has type")
            forall a. a -> [a] -> [a]
: forall (m :: * -> *). Functor m => VerboseLevel -> m Doc -> m Doc
nest VerboseLevel
2 (forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Term
t)
            forall a. a -> [a] -> [a]
: forall (m :: * -> *) (t :: * -> *).
(Applicative m, Foldable t) =>
t (m Doc) -> m Doc
fsep ( forall (m :: * -> *). Applicative m => VerboseKey -> [m Doc]
pwords VerboseKey
"which is not usable at the required modality"
                  forall a. [a] -> [a] -> [a]
++ [forall (f :: * -> *) a. Applicative f => a -> f a
pure (Modality -> Doc
attributesForModality Modality
mod) forall a. Semigroup a => a -> a -> a
<> TCMT IO Doc
"."] )
            forall a. a -> [a] -> [a]
: VerboseKey -> [TCMT IO Doc]
explanation VerboseKey
"this argument's type")

        -- Note: if a generated clause is modality-incorrect, that's a
        -- bug in the LHS modality check
        WhyCheckModality
GeneratedClause ->
          forall (m :: * -> *) a.
(HasCallStack, MonadDebug m) =>
VerboseKey -> m a
__IMPOSSIBLE_VERBOSE__ forall b c a. (b -> c) -> (a -> b) -> a -> c
. forall a. Show a => a -> VerboseKey
show forall (m :: * -> *) a b. Monad m => (a -> m b) -> m a -> m b
=<<
                   forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Term
t
              forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> TCMT IO Doc
"is not usable at the required modality"
              forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall (f :: * -> *) a. Applicative f => a -> f a
pure (Modality -> Doc
attributesForModality Modality
mod)
        WhyCheckModality
_ -> forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM Term
t forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> TCMT IO Doc
"is not usable at the required modality"
         forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall (f :: * -> *) a. Applicative f => a -> f a
pure (Modality -> Doc
attributesForModality Modality
mod)


usableAtModality :: MonadConstraint TCM => WhyCheckModality -> Modality -> Term -> TCM ()
usableAtModality :: MonadConstraint (TCMT IO) =>
WhyCheckModality -> Modality -> Term -> TCM ()
usableAtModality = MonadConstraint (TCMT IO) =>
Maybe Sort -> WhyCheckModality -> Modality -> Term -> TCM ()
usableAtModality' forall a. Maybe a
Nothing


-- * Propositions

-- | Is a type a proposition?  (Needs reduction.)

isPropM
  :: (LensSort a, PrettyTCM a, PureTCM m, MonadBlock m)
  => a -> m Bool
isPropM :: forall a (m :: * -> *).
(LensSort a, PrettyTCM a, PureTCM m, MonadBlock m) =>
a -> m Bool
isPropM a
a = do
  forall (m :: * -> *) a.
MonadDebug m =>
VerboseKey -> VerboseLevel -> TCMT IO Doc -> m a -> m a
traceSDoc VerboseKey
"tc.prop" VerboseLevel
80 (TCMT IO Doc
"Is " forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM a
a forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> TCMT IO Doc
"of sort" forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> forall a (m :: * -> *). (PrettyTCM a, MonadPretty m) => a -> m Doc
prettyTCM (forall a. LensSort a => a -> Sort
getSort a
a) forall (m :: * -> *). Applicative m => m Doc -> m Doc -> m Doc
<+> TCMT IO Doc
"in Prop?") forall a b. (a -> b) -> a -> b
$ do
  forall (m :: * -> *) t.
(MonadReduce m, MonadBlock m, IsMeta t, Reduce t) =>
t -> m t
abortIfBlocked (forall a. LensSort a => a -> Sort
getSort a
a) forall (m :: * -> *) a b. Functor m => m a -> (a -> b) -> m b
<&> \case
    Prop{} -> Bool
True
    Sort
_      -> Bool
False

isIrrelevantOrPropM
  :: (LensRelevance a, LensSort a, PrettyTCM a, PureTCM m, MonadBlock m)
  => a -> m Bool
isIrrelevantOrPropM :: forall a (m :: * -> *).
(LensRelevance a, LensSort a, PrettyTCM a, PureTCM m,
 MonadBlock m) =>
a -> m Bool
isIrrelevantOrPropM a
x = forall (m :: * -> *) a. Monad m => a -> m a
return (forall a. LensRelevance a => a -> Bool
isIrrelevant a
x) forall (m :: * -> *). Monad m => m Bool -> m Bool -> m Bool
`or2M` forall a (m :: * -> *).
(LensSort a, PrettyTCM a, PureTCM m, MonadBlock m) =>
a -> m Bool
isPropM a
x

-- * Fibrant types

-- | Is a type fibrant (i.e. Type, Prop)?

isFibrant
  :: (LensSort a, PureTCM m, MonadBlock m)
  => a -> m Bool
isFibrant :: forall a (m :: * -> *).
(LensSort a, PureTCM m, MonadBlock m) =>
a -> m Bool
isFibrant a
a = forall (m :: * -> *) t.
(MonadReduce m, MonadBlock m, IsMeta t, Reduce t) =>
t -> m t
abortIfBlocked (forall a. LensSort a => a -> Sort
getSort a
a) forall (m :: * -> *) a b. Functor m => m a -> (a -> b) -> m b
<&> \case
  Type{}     -> Bool
True
  Prop{}     -> Bool
True
  Inf IsFibrant
f Integer
_    -> IsFibrant
f forall a. Eq a => a -> a -> Bool
== IsFibrant
IsFibrant
  SSet{}     -> Bool
False
  SizeUniv{} -> Bool
False
  LockUniv{} -> Bool
False
  IntervalUniv{} -> Bool
False
  PiSort{}   -> Bool
False
  FunSort{}  -> Bool
False
  UnivSort{} -> Bool
False
  MetaS{}    -> Bool
False
  DefS{}     -> Bool
False
  DummyS{}   -> Bool
False


-- | Cofibrant types are those that could be the domain of a fibrant
--   pi type. (Notion by C. Sattler).
isCoFibrantSort :: (LensSort a, PureTCM m, MonadBlock m) => a -> m Bool
isCoFibrantSort :: forall a (m :: * -> *).
(LensSort a, PureTCM m, MonadBlock m) =>
a -> m Bool
isCoFibrantSort a
a = forall (m :: * -> *) t.
(MonadReduce m, MonadBlock m, IsMeta t, Reduce t) =>
t -> m t
abortIfBlocked (forall a. LensSort a => a -> Sort
getSort a
a) forall (m :: * -> *) a b. Functor m => m a -> (a -> b) -> m b
<&> \case
  Type{}     -> Bool
True
  Prop{}     -> Bool
True
  Inf IsFibrant
f Integer
_    -> IsFibrant
f forall a. Eq a => a -> a -> Bool
== IsFibrant
IsFibrant
  SSet{}     -> Bool
False
  SizeUniv{} -> Bool
False
  LockUniv{} -> Bool
True
  IntervalUniv{} -> Bool
True
  PiSort{}   -> Bool
False
  FunSort{}  -> Bool
False
  UnivSort{} -> Bool
False
  MetaS{}    -> Bool
False
  DefS{}     -> Bool
False
  DummyS{}   -> Bool
False