make uses of eff more consistent

This commit is contained in:
rhiannon morris 2023-08-24 19:55:57 +02:00
parent 8264a1bb81
commit 4b6b3853a1
5 changed files with 81 additions and 76 deletions

View file

@ -13,15 +13,16 @@ EqModeState : Type -> Type
EqModeState = State EqMode EqModeState = State EqMode
public export public export
Equal : Type -> Type Equal : List (Type -> Type)
Equal = Eff [ErrorEff, DefsReader, NameGen] Equal = [ErrorEff, DefsReader, NameGen]
public export public export
Equal_ : Type -> Type EqualInner : List (Type -> Type)
Equal_ = Eff [ErrorEff, NameGen, EqModeState] EqualInner = [ErrorEff, NameGen, EqModeState]
export export
runEqualWith_ : EqMode -> NameSuf -> Equal_ a -> (Either Error a, NameSuf) runEqualWith_ : EqMode -> NameSuf ->
Eff EqualInner a -> (Either Error a, NameSuf)
runEqualWith_ mode suf act = runEqualWith_ mode suf act =
extract $ extract $
runNameGenWith suf $ runNameGenWith suf $
@ -29,12 +30,13 @@ runEqualWith_ mode suf act =
evalState mode act evalState mode act
export export
runEqual_ : EqMode -> Equal_ a -> Either Error a runEqualInner : EqMode -> Eff EqualInner a -> Either Error a
runEqual_ mode act = fst $ runEqualWith_ mode 0 act runEqualInner mode act = fst $ runEqualWith_ mode 0 act
export export
runEqualWith : NameSuf -> Definitions -> Equal a -> (Either Error a, NameSuf) runEqualWith : NameSuf -> Definitions ->
Eff Equal a -> (Either Error a, NameSuf)
runEqualWith suf defs act = runEqualWith suf defs act =
extract $ extract $
runStateAt GEN suf $ runStateAt GEN suf $
@ -42,7 +44,7 @@ runEqualWith suf defs act =
runExcept act runExcept act
export export
runEqual : Definitions -> Equal a -> Either Error a runEqual : Definitions -> Eff Equal a -> Either Error a
runEqual defs act = fst $ runEqualWith 0 defs act runEqual defs act = fst $ runEqualWith 0 defs act
@ -53,19 +55,19 @@ mode = get
parameters (loc : Loc) (ctx : EqContext n) parameters (loc : Loc) (ctx : EqContext n)
private %inline private %inline
clashT : Term 0 n -> Term 0 n -> Term 0 n -> Equal_ a clashT : Term 0 n -> Term 0 n -> Term 0 n -> Eff EqualInner a
clashT ty s t = throw $ ClashT loc ctx !mode ty s t clashT ty s t = throw $ ClashT loc ctx !mode ty s t
private %inline private %inline
clashTy : Term 0 n -> Term 0 n -> Equal_ a clashTy : Term 0 n -> Term 0 n -> Eff EqualInner a
clashTy s t = throw $ ClashTy loc ctx !mode s t clashTy s t = throw $ ClashTy loc ctx !mode s t
private %inline private %inline
clashE : Elim 0 n -> Elim 0 n -> Equal_ a clashE : Elim 0 n -> Elim 0 n -> Eff EqualInner a
clashE e f = throw $ ClashE loc ctx !mode e f clashE e f = throw $ ClashE loc ctx !mode e f
private %inline private %inline
wrongType : Term 0 n -> Term 0 n -> Equal_ a wrongType : Term 0 n -> Term 0 n -> Eff EqualInner a
wrongType ty s = throw $ WrongType loc ctx ty s wrongType ty s = throw $ WrongType loc ctx ty s
@ -101,7 +103,8 @@ sameTyCon (E {}) _ = False
||| * an enum type is a subsingleton if it has zero or one tags. ||| * an enum type is a subsingleton if it has zero or one tags.
||| * a box type is a subsingleton if its content is ||| * a box type is a subsingleton if its content is
public export covering public export covering
isSubSing : {n : Nat} -> Definitions -> EqContext n -> Term 0 n -> Equal_ Bool isSubSing : {n : Nat} -> Definitions -> EqContext n -> Term 0 n ->
Eff EqualInner Bool
isSubSing defs ctx ty0 = do isSubSing defs ctx ty0 = do
Element ty0 nc <- whnf defs ctx ty0.loc ty0 Element ty0 nc <- whnf defs ctx ty0.loc ty0
case ty0 of case ty0 of
@ -141,7 +144,7 @@ ensureTyCon loc ctx t = case nchoose $ isTyConE t of
private covering private covering
computeElimTypeE : (defs : Definitions) -> EqContext n -> computeElimTypeE : (defs : Definitions) -> EqContext n ->
(e : Elim 0 n) -> (0 ne : NotRedex defs e) => (e : Elim 0 n) -> (0 ne : NotRedex defs e) =>
Equal_ (Term 0 n) Eff EqualInner (Term 0 n)
computeElimTypeE defs ectx e = computeElimTypeE defs ectx e =
let Val n = ectx.termLen in let Val n = ectx.termLen in
lift $ computeElimType defs (toWhnfContext ectx) e lift $ computeElimType defs (toWhnfContext ectx) e
@ -154,7 +157,7 @@ parameters (defs : Definitions)
||| |||
||| ⚠ **assumes that `s`, `t` have already been checked against `ty`**. ⚠ ||| ⚠ **assumes that `s`, `t` have already been checked against `ty`**. ⚠
export covering %inline export covering %inline
compare0 : EqContext n -> (ty, s, t : Term 0 n) -> Equal_ () compare0 : EqContext n -> (ty, s, t : Term 0 n) -> Eff EqualInner ()
compare0 ctx ty s t = compare0 ctx ty s t =
wrapErr (WhileComparingT ctx !mode ty s t) $ do wrapErr (WhileComparingT ctx !mode ty s t) $ do
let Val n = ctx.termLen let Val n = ctx.termLen
@ -175,7 +178,7 @@ parameters (defs : Definitions)
(ty, s, t : Term 0 n) -> (ty, s, t : Term 0 n) ->
(0 _ : NotRedex defs ty) => (0 _ : So (isTyConE ty)) => (0 _ : NotRedex defs ty) => (0 _ : So (isTyConE ty)) =>
(0 _ : NotRedex defs s) => (0 _ : NotRedex defs t) => (0 _ : NotRedex defs s) => (0 _ : NotRedex defs t) =>
Equal_ () Eff EqualInner ()
compare0' ctx (TYPE {}) s t = compareType ctx s t compare0' ctx (TYPE {}) s t = compareType ctx s t
compare0' ctx ty@(Pi {qty, arg, res, _}) s t {n} = local_ Equal $ compare0' ctx ty@(Pi {qty, arg, res, _}) s t {n} = local_ Equal $
@ -201,7 +204,7 @@ parameters (defs : Definitions)
ctx' : EqContext (S n) ctx' : EqContext (S n)
ctx' = extendTy qty res.name arg ctx ctx' = extendTy qty res.name arg ctx
eta : Loc -> Elim 0 n -> ScopeTerm 0 n -> Equal_ () eta : Loc -> Elim 0 n -> ScopeTerm 0 n -> Eff EqualInner ()
eta _ e (S _ (Y b)) = compare0 ctx' res.term (toLamBody e) b eta _ e (S _ (Y b)) = compare0 ctx' res.term (toLamBody e) b
eta loc e (S _ (N _)) = clashT loc ctx ty s t eta loc e (S _ (N _)) = clashT loc ctx ty s t
@ -297,7 +300,7 @@ parameters (defs : Definitions)
||| compares two types, using the current variance `mode` for universes. ||| compares two types, using the current variance `mode` for universes.
||| fails if they are not types, even if they would happen to be equal. ||| fails if they are not types, even if they would happen to be equal.
export covering %inline export covering %inline
compareType : EqContext n -> (s, t : Term 0 n) -> Equal_ () compareType : EqContext n -> (s, t : Term 0 n) -> Eff EqualInner ()
compareType ctx s t = do compareType ctx s t = do
let Val n = ctx.termLen let Val n = ctx.termLen
Element s' _ <- whnf defs ctx s.loc s Element s' _ <- whnf defs ctx s.loc s
@ -313,7 +316,7 @@ parameters (defs : Definitions)
(0 _ : NotRedex defs s) => (0 _ : So (isTyConE s)) => (0 _ : NotRedex defs s) => (0 _ : So (isTyConE s)) =>
(0 _ : NotRedex defs t) => (0 _ : So (isTyConE t)) => (0 _ : NotRedex defs t) => (0 _ : So (isTyConE t)) =>
(0 _ : So (sameTyCon s t)) => (0 _ : So (sameTyCon s t)) =>
Equal_ () Eff EqualInner ()
-- equality is the same as subtyping, except with the -- equality is the same as subtyping, except with the
-- "≤" in the TYPE rule being replaced with "=" -- "≤" in the TYPE rule being replaced with "="
compareType' ctx a@(TYPE k {}) (TYPE l {}) = compareType' ctx a@(TYPE k {}) (TYPE l {}) =
@ -384,7 +387,7 @@ parameters (defs : Definitions)
||| ⚠ **assumes that they have both been typechecked, and have ||| ⚠ **assumes that they have both been typechecked, and have
||| equal types.** ⚠ ||| equal types.** ⚠
export covering %inline export covering %inline
compare0 : EqContext n -> (e, f : Elim 0 n) -> Equal_ () compare0 : EqContext n -> (e, f : Elim 0 n) -> Eff EqualInner ()
compare0 ctx e f = compare0 ctx e f =
wrapErr (WhileComparingE ctx !mode e f) $ do wrapErr (WhileComparingE ctx !mode e f) $ do
let Val n = ctx.termLen let Val n = ctx.termLen
@ -397,7 +400,7 @@ parameters (defs : Definitions)
compare0' : EqContext n -> compare0' : EqContext n ->
(e, f : Elim 0 n) -> (e, f : Elim 0 n) ->
(0 ne : NotRedex defs e) -> (0 nf : NotRedex defs f) -> (0 ne : NotRedex defs e) -> (0 nf : NotRedex defs f) ->
Equal_ () Eff EqualInner ()
compare0' ctx e@(F x u _) f@(F y v _) _ _ = compare0' ctx e@(F x u _) f@(F y v _) _ _ =
unless (x == y && u == v) $ clashE e.loc ctx e f unless (x == y && u == v) $ clashE e.loc ctx e f
@ -457,7 +460,8 @@ parameters (defs : Definitions)
compare0 ctx (sub1 eret $ Ann (Tag t l.loc) ety l.loc) l r compare0 ctx (sub1 eret $ Ann (Tag t l.loc) ety l.loc) l r
expectEqualQ eloc epi fpi expectEqualQ eloc epi fpi
where where
lookupArm : Loc -> TagVal -> CaseEnumArms d n -> Equal_ (Term d n) lookupArm : Loc -> TagVal -> CaseEnumArms d n ->
Eff EqualInner (Term d n)
lookupArm loc t arms = case lookup t arms of lookupArm loc t arms = case lookup t arms of
Just arm => pure arm Just arm => pure arm
Nothing => throw $ TagNotIn loc t (fromList $ keys arms) Nothing => throw $ TagNotIn loc t (fromList $ keys arms)
@ -570,7 +574,7 @@ parameters (defs : Definitions)
(ret : Term 0 n) -> (u : Universe) -> (ret : Term 0 n) -> (u : Universe) ->
(b1, b2 : Maybe (TypeCaseArmBody k 0 n)) -> (b1, b2 : Maybe (TypeCaseArmBody k 0 n)) ->
(def : Term 0 n) -> (def : Term 0 n) ->
Equal_ () Eff EqualInner ()
compareArm {b1 = Nothing, b2 = Nothing, _} = pure () compareArm {b1 = Nothing, b2 = Nothing, _} = pure ()
compareArm ctx k ret u b1 b2 def = compareArm ctx k ret u b1 b2 def =
let def = SN def in let def = SN def in
@ -580,7 +584,7 @@ parameters (defs : Definitions)
compareArm_ : EqContext n -> (k : TyConKind) -> compareArm_ : EqContext n -> (k : TyConKind) ->
(ret : Term 0 n) -> (u : Universe) -> (ret : Term 0 n) -> (u : Universe) ->
(b1, b2 : TypeCaseArmBody k 0 n) -> (b1, b2 : TypeCaseArmBody k 0 n) ->
Equal_ () Eff EqualInner ()
compareArm_ ctx KTYPE ret u b1 b2 = compareArm_ ctx KTYPE ret u b1 b2 =
compare0 ctx ret b1.term b2.term compare0 ctx ret b1.term b2.term
@ -626,8 +630,8 @@ parameters (loc : Loc) (ctx : TyContext d n)
parameters (mode : EqMode) parameters (mode : EqMode)
private private
fromEqual_ : Equal_ a -> Equal a fromInner : Eff EqualInner a -> Eff Equal a
fromEqual_ act = lift $ evalState mode act fromInner act = lift $ evalState mode act
private private
eachFace : Applicative f => (EqContext n -> DSubst d 0 -> f ()) -> f () eachFace : Applicative f => (EqContext n -> DSubst d 0 -> f ()) -> f ()
@ -635,18 +639,22 @@ parameters (loc : Loc) (ctx : TyContext d n)
for_ (splits loc ctx.dctx) $ \th => act (makeEqContext ctx th) th for_ (splits loc ctx.dctx) $ \th => act (makeEqContext ctx th) th
private private
runCompare : (Definitions -> EqContext n -> DSubst d 0 -> Equal_ ()) -> CompareAction : Nat -> Nat -> Type
Equal () CompareAction d n =
runCompare act = fromEqual_ $ eachFace $ act !(askAt DEFS) Definitions -> EqContext n -> DSubst d 0 -> Eff EqualInner ()
private
runCompare : CompareAction d n -> Eff Equal ()
runCompare act = fromInner $ eachFace $ act !(askAt DEFS)
namespace Term namespace Term
export covering export covering
compare : (ty, s, t : Term d n) -> Equal () compare : (ty, s, t : Term d n) -> Eff Equal ()
compare ty s t = runCompare $ \defs, ectx, th => compare ty s t = runCompare $ \defs, ectx, th =>
compare0 defs ectx (ty // th) (s // th) (t // th) compare0 defs ectx (ty // th) (s // th) (t // th)
export covering export covering
compareType : (s, t : Term d n) -> Equal () compareType : (s, t : Term d n) -> Eff Equal ()
compareType s t = runCompare $ \defs, ectx, th => compareType s t = runCompare $ \defs, ectx, th =>
compareType defs ectx (s // th) (t // th) compareType defs ectx (s // th) (t // th)
@ -654,26 +662,26 @@ parameters (loc : Loc) (ctx : TyContext d n)
||| you don't have to pass the type in but the arguments must still be ||| you don't have to pass the type in but the arguments must still be
||| of the same type!! ||| of the same type!!
export covering export covering
compare : (e, f : Elim d n) -> Equal () compare : (e, f : Elim d n) -> Eff Equal ()
compare e f = runCompare $ \defs, ectx, th => compare e f = runCompare $ \defs, ectx, th =>
compare0 defs ectx (e // th) (f // th) compare0 defs ectx (e // th) (f // th)
namespace Term namespace Term
export covering %inline export covering %inline
equal, sub, super : (ty, s, t : Term d n) -> Equal () equal, sub, super : (ty, s, t : Term d n) -> Eff Equal ()
equal = compare Equal equal = compare Equal
sub = compare Sub sub = compare Sub
super = compare Super super = compare Super
export covering %inline export covering %inline
equalType, subtype, supertype : (s, t : Term d n) -> Equal () equalType, subtype, supertype : (s, t : Term d n) -> Eff Equal ()
equalType = compareType Equal equalType = compareType Equal
subtype = compareType Sub subtype = compareType Sub
supertype = compareType Super supertype = compareType Super
namespace Elim namespace Elim
export covering %inline export covering %inline
equal, sub, super : (e, f : Elim d n) -> Equal () equal, sub, super : (e, f : Elim d n) -> Eff Equal ()
equal = compare Equal equal = compare Equal
sub = compare Sub sub = compare Sub
super = compare Super super = compare Super

View file

@ -293,7 +293,7 @@ fromPBaseNameNS : PBaseName -> Eff [StateL NS Mods] Name
fromPBaseNameNS name = pure $ addMods !(getAt NS) $ fromPBaseName name fromPBaseNameNS name = pure $ addMods !(getAt NS) $ fromPBaseName name
private private
liftTC : TC a -> Eff FromParserPure a liftTC : Eff TC a -> Eff FromParserPure a
liftTC act = do liftTC act = do
res <- lift $ runExcept $ runReaderAt DEFS !(getAt DEFS) act res <- lift $ runExcept $ runReaderAt DEFS !(getAt DEFS) act
rethrow $ mapFst WrapTypeError res rethrow $ mapFst WrapTypeError res

View file

@ -13,25 +13,22 @@ import Quox.EffExtra
public export public export
0 TCEff : List (Type -> Type) 0 TC : List (Type -> Type)
TCEff = [ErrorEff, DefsReader, NameGen] TC = [ErrorEff, DefsReader, NameGen]
public export
0 TC : Type -> Type
TC = Eff TCEff
export export
runTCWith : NameSuf -> Definitions -> TC a -> (Either Error a, NameSuf) runTCWith : NameSuf -> Definitions -> Eff TC a -> (Either Error a, NameSuf)
runTCWith = runEqualWith runTCWith = runEqualWith
export export
runTC : Definitions -> TC a -> Either Error a runTC : Definitions -> Eff TC a -> Either Error a
runTC = runEqual runTC = runEqual
parameters (loc : Loc) parameters (loc : Loc)
export export
popQs : Has ErrorEff fs => QContext s -> QOutput (s + n) -> Eff fs (QOutput n) popQs : Has ErrorEff fs => QContext s -> QOutput (s + n) ->
Eff fs (QOutput n)
popQs [<] qout = pure qout popQs [<] qout = pure qout
popQs (pis :< pi) (qout :< rh) = do expectCompatQ loc rh pi; popQs pis qout popQs (pis :< pi) (qout :< rh) = do expectCompatQ loc rh pi; popQs pis qout
@ -91,28 +88,28 @@ mutual
||| doing any further work. ||| doing any further work.
export covering %inline export covering %inline
check : (ctx : TyContext d n) -> SQty -> Term d n -> Term d n -> check : (ctx : TyContext d n) -> SQty -> Term d n -> Term d n ->
TC (CheckResult ctx.dctx n) Eff TC (CheckResult ctx.dctx n)
check ctx sg subj ty = ifConsistent ctx.dctx $ checkC ctx sg subj ty check ctx sg subj ty = ifConsistent ctx.dctx $ checkC ctx sg subj ty
||| "Ψ | Γ ⊢₀ s ⇐ A" ||| "Ψ | Γ ⊢₀ s ⇐ A"
||| |||
||| `check0 ctx subj ty` checks a term (as `check`) in a zero context. ||| `check0 ctx subj ty` checks a term (as `check`) in a zero context.
export covering %inline export covering %inline
check0 : TyContext d n -> Term d n -> Term d n -> TC () check0 : TyContext d n -> Term d n -> Term d n -> Eff TC ()
check0 ctx tm ty = ignore $ check ctx szero tm ty check0 ctx tm ty = ignore $ check ctx szero tm ty
-- the output will always be 𝟎 because the subject quantity is 0 -- the output will always be 𝟎 because the subject quantity is 0
||| `check`, assuming the dimension context is consistent ||| `check`, assuming the dimension context is consistent
export covering %inline export covering %inline
checkC : (ctx : TyContext d n) -> SQty -> Term d n -> Term d n -> checkC : (ctx : TyContext d n) -> SQty -> Term d n -> Term d n ->
TC (CheckResult' n) Eff TC (CheckResult' n)
checkC ctx sg subj ty = checkC ctx sg subj ty =
wrapErr (WhileChecking ctx sg.fst subj ty) $ wrapErr (WhileChecking ctx sg.fst subj ty) $
checkCNoWrap ctx sg subj ty checkCNoWrap ctx sg subj ty
export covering %inline export covering %inline
checkCNoWrap : (ctx : TyContext d n) -> SQty -> Term d n -> Term d n -> checkCNoWrap : (ctx : TyContext d n) -> SQty -> Term d n -> Term d n ->
TC (CheckResult' n) Eff TC (CheckResult' n)
checkCNoWrap ctx sg subj ty = checkCNoWrap ctx sg subj ty =
let Element subj nc = pushSubsts subj in let Element subj nc = pushSubsts subj in
check' ctx sg subj ty check' ctx sg subj ty
@ -122,16 +119,16 @@ mutual
||| `checkType ctx subj ty` checks a type (in a zero context). sometimes the ||| `checkType ctx subj ty` checks a type (in a zero context). sometimes the
||| universe doesn't matter, only that a term is _a_ type, so it is optional. ||| universe doesn't matter, only that a term is _a_ type, so it is optional.
export covering %inline export covering %inline
checkType : TyContext d n -> Term d n -> Maybe Universe -> TC () checkType : TyContext d n -> Term d n -> Maybe Universe -> Eff TC ()
checkType ctx subj l = ignore $ ifConsistent ctx.dctx $ checkTypeC ctx subj l checkType ctx subj l = ignore $ ifConsistent ctx.dctx $ checkTypeC ctx subj l
export covering %inline export covering %inline
checkTypeC : TyContext d n -> Term d n -> Maybe Universe -> TC () checkTypeC : TyContext d n -> Term d n -> Maybe Universe -> Eff TC ()
checkTypeC ctx subj l = checkTypeC ctx subj l =
wrapErr (WhileCheckingTy ctx subj l) $ checkTypeNoWrap ctx subj l wrapErr (WhileCheckingTy ctx subj l) $ checkTypeNoWrap ctx subj l
export covering %inline export covering %inline
checkTypeNoWrap : TyContext d n -> Term d n -> Maybe Universe -> TC () checkTypeNoWrap : TyContext d n -> Term d n -> Maybe Universe -> Eff TC ()
checkTypeNoWrap ctx subj l = checkTypeNoWrap ctx subj l =
let Element subj nc = pushSubsts subj in let Element subj nc = pushSubsts subj in
checkType' ctx subj l checkType' ctx subj l
@ -145,13 +142,13 @@ mutual
||| doing any further work. ||| doing any further work.
export covering %inline export covering %inline
infer : (ctx : TyContext d n) -> SQty -> Elim d n -> infer : (ctx : TyContext d n) -> SQty -> Elim d n ->
TC (InferResult ctx.dctx d n) Eff TC (InferResult ctx.dctx d n)
infer ctx sg subj = ifConsistent ctx.dctx $ inferC ctx sg subj infer ctx sg subj = ifConsistent ctx.dctx $ inferC ctx sg subj
||| `infer`, assuming the dimension context is consistent ||| `infer`, assuming the dimension context is consistent
export covering %inline export covering %inline
inferC : (ctx : TyContext d n) -> SQty -> Elim d n -> inferC : (ctx : TyContext d n) -> SQty -> Elim d n ->
TC (InferResult' d n) Eff TC (InferResult' d n)
inferC ctx sg subj = inferC ctx sg subj =
wrapErr (WhileInferring ctx sg.fst subj) $ wrapErr (WhileInferring ctx sg.fst subj) $
let Element subj nc = pushSubsts subj in let Element subj nc = pushSubsts subj in
@ -161,7 +158,7 @@ mutual
private covering private covering
toCheckType : TyContext d n -> SQty -> toCheckType : TyContext d n -> SQty ->
(subj : Term d n) -> (0 nc : NotClo subj) => Term d n -> (subj : Term d n) -> (0 nc : NotClo subj) => Term d n ->
TC (CheckResult' n) Eff TC (CheckResult' n)
toCheckType ctx sg t ty = do toCheckType ctx sg t ty = do
u <- expectTYPE !(askAt DEFS) ctx ty.loc ty u <- expectTYPE !(askAt DEFS) ctx ty.loc ty
expectEqualQ t.loc Zero sg.fst expectEqualQ t.loc Zero sg.fst
@ -171,7 +168,7 @@ mutual
private covering private covering
check' : TyContext d n -> SQty -> check' : TyContext d n -> SQty ->
(subj : Term d n) -> (0 nc : NotClo subj) => Term d n -> (subj : Term d n) -> (0 nc : NotClo subj) => Term d n ->
TC (CheckResult' n) Eff TC (CheckResult' n)
check' ctx sg t@(TYPE {}) ty = toCheckType ctx sg t ty check' ctx sg t@(TYPE {}) ty = toCheckType ctx sg t ty
@ -253,7 +250,7 @@ mutual
private covering private covering
checkType' : TyContext d n -> checkType' : TyContext d n ->
(subj : Term d n) -> (0 nc : NotClo subj) => (subj : Term d n) -> (0 nc : NotClo subj) =>
Maybe Universe -> TC () Maybe Universe -> Eff TC ()
checkType' ctx (TYPE k loc) u = do checkType' ctx (TYPE k loc) u = do
-- if 𝓀 < then Ψ | Γ ⊢₀ Type 𝓀 ⇐ Type -- if 𝓀 < then Ψ | Γ ⊢₀ Type 𝓀 ⇐ Type
@ -320,7 +317,7 @@ mutual
private covering private covering
checkTypeScope : TyContext d n -> Term d n -> checkTypeScope : TyContext d n -> Term d n ->
ScopeTerm d n -> Maybe Universe -> TC () ScopeTerm d n -> Maybe Universe -> Eff TC ()
checkTypeScope ctx s (S _ (N body)) u = checkType ctx body u checkTypeScope ctx s (S _ (N body)) u = checkType ctx body u
checkTypeScope ctx s (S [< x] (Y body)) u = checkTypeScope ctx s (S [< x] (Y body)) u =
checkType (extendTy Zero x s ctx) body u checkType (extendTy Zero x s ctx) body u
@ -329,7 +326,7 @@ mutual
private covering private covering
infer' : TyContext d n -> SQty -> infer' : TyContext d n -> SQty ->
(subj : Elim d n) -> (0 nc : NotClo subj) => (subj : Elim d n) -> (0 nc : NotClo subj) =>
TC (InferResult' d n) Eff TC (InferResult' d n)
infer' ctx sg (F x u loc) = do infer' ctx sg (F x u loc) = do
-- if π·x : A {≔ s} in global context -- if π·x : A {≔ s} in global context

View file

@ -20,7 +20,7 @@ defGlobals = fromList
("eq-AB", ^mkPostulate gzero (^Eq0 (^TYPE 0) (^FT "A" 0) (^FT "B" 0))), ("eq-AB", ^mkPostulate gzero (^Eq0 (^TYPE 0) (^FT "A" 0) (^FT "B" 0))),
("two", ^mkDef gany (^Nat) (^Succ (^Succ (^Zero))))] ("two", ^mkDef gany (^Nat) (^Succ (^Succ (^Zero))))]
parameters (label : String) (act : Equal ()) parameters (label : String) (act : Eff Equal ())
{default defGlobals globals : Definitions} {default defGlobals globals : Definitions}
testEq : Test testEq : Test
testEq = test label $ runEqual globals act testEq = test label $ runEqual globals act
@ -30,13 +30,13 @@ parameters (label : String) (act : Equal ())
parameters (ctx : TyContext d n) parameters (ctx : TyContext d n)
subT, equalT : Term d n -> Term d n -> Term d n -> TC () subT, equalT : Term d n -> Term d n -> Term d n -> Eff TC ()
subT ty s t = lift $ Term.sub noLoc ctx ty s t subT ty s t = lift $ Term.sub noLoc ctx ty s t
equalT ty s t = lift $ Term.equal noLoc ctx ty s t equalT ty s t = lift $ Term.equal noLoc ctx ty s t
equalTy : Term d n -> Term d n -> TC () equalTy : Term d n -> Term d n -> Eff TC ()
equalTy s t = lift $ Term.equalType noLoc ctx s t equalTy s t = lift $ Term.equalType noLoc ctx s t
subE, equalE : Elim d n -> Elim d n -> TC () subE, equalE : Elim d n -> Elim d n -> Eff TC ()
subE e f = lift $ Elim.sub noLoc ctx e f subE e f = lift $ Elim.sub noLoc ctx e f
equalE e f = lift $ Elim.equal noLoc ctx e f equalE e f = lift $ Elim.equal noLoc ctx e f

View file

@ -30,10 +30,10 @@ ToInfo Error' where
("wanted", show good), ("wanted", show good),
("wanted", show bad)] ("wanted", show bad)]
0 M : Type -> Type 0 Test : List (Type -> Type)
M = Eff [Except Error', DefsReader] Test = [Except Error', DefsReader]
inj : TC a -> M a inj : Eff TC a -> Eff Test a
inj act = rethrow $ mapFst TCError $ runTC !(askAt DEFS) act inj act = rethrow $ mapFst TCError $ runTC !(askAt DEFS) act
@ -109,7 +109,7 @@ defGlobals = fromList
("fst", ^mkDef gany fstTy fstDef), ("fst", ^mkDef gany fstTy fstDef),
("snd", ^mkDef gany sndTy sndDef)] ("snd", ^mkDef gany sndTy sndDef)]
parameters (label : String) (act : Lazy (M ())) parameters (label : String) (act : Lazy (Eff Test ()))
{default defGlobals globals : Definitions} {default defGlobals globals : Definitions}
testTC : Test testTC : Test
testTC = test label {e = Error', a = ()} $ testTC = test label {e = Error', a = ()} $
@ -120,22 +120,22 @@ parameters (label : String) (act : Lazy (M ()))
(extract $ runExcept $ runReaderAt DEFS globals act) $> "()" (extract $ runExcept $ runReaderAt DEFS globals act) $> "()"
inferredTypeEq : TyContext d n -> (exp, got : Term d n) -> M () inferredTypeEq : TyContext d n -> (exp, got : Term d n) -> Eff Test ()
inferredTypeEq ctx exp got = inferredTypeEq ctx exp got =
wrapErr (const $ WrongInfer ctx.dnames ctx.tnames exp got) $ inj $ lift $ wrapErr (const $ WrongInfer ctx.dnames ctx.tnames exp got) $ inj $ lift $
equalType noLoc ctx exp got equalType noLoc ctx exp got
qoutEq : (exp, got : QOutput n) -> M () qoutEq : (exp, got : QOutput n) -> Eff Test ()
qoutEq qout res = unless (qout == res) $ throw $ WrongQOut qout res qoutEq qout res = unless (qout == res) $ throw $ WrongQOut qout res
inferAs : TyContext d n -> (sg : SQty) -> Elim d n -> Term d n -> M () inferAs : TyContext d n -> (sg : SQty) -> Elim d n -> Term d n -> Eff Test ()
inferAs ctx@(MkTyContext {dctx, _}) sg e ty = do inferAs ctx@(MkTyContext {dctx, _}) sg e ty = do
case !(inj $ infer ctx sg e) of case !(inj $ infer ctx sg e) of
Just res => inferredTypeEq ctx ty res.type Just res => inferredTypeEq ctx ty res.type
Nothing => pure () Nothing => pure ()
inferAsQ : TyContext d n -> (sg : SQty) -> inferAsQ : TyContext d n -> (sg : SQty) ->
Elim d n -> Term d n -> QOutput n -> M () Elim d n -> Term d n -> QOutput n -> Eff Test ()
inferAsQ ctx@(MkTyContext {dctx, _}) sg e ty qout = do inferAsQ ctx@(MkTyContext {dctx, _}) sg e ty qout = do
case !(inj $ infer ctx sg e) of case !(inj $ infer ctx sg e) of
Just res => do Just res => do
@ -143,20 +143,20 @@ inferAsQ ctx@(MkTyContext {dctx, _}) sg e ty qout = do
qoutEq qout res.qout qoutEq qout res.qout
Nothing => pure () Nothing => pure ()
infer_ : TyContext d n -> (sg : SQty) -> Elim d n -> M () infer_ : TyContext d n -> (sg : SQty) -> Elim d n -> Eff Test ()
infer_ ctx sg e = ignore $ inj $ infer ctx sg e infer_ ctx sg e = ignore $ inj $ infer ctx sg e
checkQ : TyContext d n -> SQty -> checkQ : TyContext d n -> SQty ->
Term d n -> Term d n -> QOutput n -> M () Term d n -> Term d n -> QOutput n -> Eff Test ()
checkQ ctx@(MkTyContext {dctx, _}) sg s ty qout = do checkQ ctx@(MkTyContext {dctx, _}) sg s ty qout = do
case !(inj $ check ctx sg s ty) of case !(inj $ check ctx sg s ty) of
Just res => qoutEq qout res Just res => qoutEq qout res
Nothing => pure () Nothing => pure ()
check_ : TyContext d n -> SQty -> Term d n -> Term d n -> M () check_ : TyContext d n -> SQty -> Term d n -> Term d n -> Eff Test ()
check_ ctx sg s ty = ignore $ inj $ check ctx sg s ty check_ ctx sg s ty = ignore $ inj $ check ctx sg s ty
checkType_ : TyContext d n -> Term d n -> Maybe Universe -> M () checkType_ : TyContext d n -> Term d n -> Maybe Universe -> Eff Test ()
checkType_ ctx s u = inj $ checkType ctx s u checkType_ ctx s u = inj $ checkType ctx s u