2022-04-23 18:21:30 -04:00
|
|
|
module Quox.Typechecker
|
|
|
|
|
|
|
|
import public Quox.Syntax
|
|
|
|
import public Quox.Typing
|
2022-08-22 23:43:23 -04:00
|
|
|
import public Quox.Equal
|
|
|
|
import public Control.Monad.Either
|
2023-01-08 14:44:25 -05:00
|
|
|
import Decidable.Decidable
|
2022-04-23 18:21:30 -04:00
|
|
|
|
|
|
|
%default total
|
|
|
|
|
|
|
|
|
2023-01-20 20:34:28 -05:00
|
|
|
public export
|
|
|
|
0 CanTC' : (q : Type) -> (q -> Type) -> (Type -> Type) -> Type
|
|
|
|
CanTC' q isGlobal m = (HasErr q m, MonadReader (Definitions' q isGlobal) m)
|
|
|
|
|
|
|
|
public export
|
|
|
|
0 CanTC : (q : Type) -> IsQty q => (Type -> Type) -> Type
|
|
|
|
CanTC q = CanTC' q IsGlobal
|
|
|
|
|
|
|
|
|
2022-04-23 18:21:30 -04:00
|
|
|
private covering %inline
|
2023-01-22 18:53:34 -05:00
|
|
|
expectTYPE : CanTC' q _ m => Term q d n -> m Universe
|
2022-04-23 18:21:30 -04:00
|
|
|
expectTYPE s =
|
2023-01-22 18:53:34 -05:00
|
|
|
case whnf !ask s of
|
|
|
|
Element (TYPE l) _ => pure l
|
|
|
|
_ => throwError $ ExpectedTYPE s
|
2022-04-23 18:21:30 -04:00
|
|
|
|
|
|
|
private covering %inline
|
2023-01-22 18:53:34 -05:00
|
|
|
expectPi : CanTC' q _ m => Term q d n ->
|
2023-01-08 14:44:25 -05:00
|
|
|
m (q, Term q d n, ScopeTerm q d n)
|
2022-04-23 18:21:30 -04:00
|
|
|
expectPi ty =
|
2023-01-22 18:53:34 -05:00
|
|
|
case whnf !ask ty of
|
2023-01-20 20:34:28 -05:00
|
|
|
Element (Pi qty _ arg res) _ => pure (qty, arg, res)
|
|
|
|
_ => throwError $ ExpectedPi ty
|
|
|
|
|
|
|
|
private covering %inline
|
2023-01-22 18:53:34 -05:00
|
|
|
expectEq : CanTC' q _ m => Term q d n ->
|
2023-01-20 20:34:28 -05:00
|
|
|
m (DScopeTerm q d n, Term q d n, Term q d n)
|
|
|
|
expectEq ty =
|
2023-01-22 18:53:34 -05:00
|
|
|
case whnf !ask ty of
|
2023-01-20 20:34:28 -05:00
|
|
|
Element (Eq _ ty l r) _ => pure (ty, l, r)
|
|
|
|
_ => throwError $ ExpectedEq ty
|
2022-04-23 18:21:30 -04:00
|
|
|
|
|
|
|
private %inline
|
2023-01-08 14:44:25 -05:00
|
|
|
expectEqualQ : HasErr q m => Eq q =>
|
|
|
|
(expect, actual : q) -> m ()
|
2022-05-06 15:23:58 -04:00
|
|
|
expectEqualQ pi rh =
|
2022-08-22 23:43:23 -04:00
|
|
|
unless (pi == rh) $ throwError $ ClashQ pi rh
|
2022-04-23 18:21:30 -04:00
|
|
|
|
|
|
|
|
|
|
|
private %inline
|
2023-01-08 14:44:25 -05:00
|
|
|
popQ : HasErr q m => Eq q => q -> QOutput q (S n) -> m (QOutput q n)
|
2022-04-23 18:21:30 -04:00
|
|
|
popQ pi (qctx :< rh) = expectEqualQ pi rh $> qctx
|
|
|
|
|
|
|
|
|
|
|
|
private %inline
|
2023-01-08 14:44:25 -05:00
|
|
|
tail : TyContext q d (S n) -> TyContext q d n
|
2022-04-23 18:21:30 -04:00
|
|
|
tail = {tctx $= tail, qctx $= tail}
|
|
|
|
|
|
|
|
|
|
|
|
private %inline
|
2023-01-08 14:44:25 -05:00
|
|
|
weakI : IsQty q => InferResult q d n -> InferResult q d (S n)
|
2022-04-27 15:58:09 -04:00
|
|
|
weakI = {type $= weakT, qout $= (:< zero)}
|
2022-04-23 18:21:30 -04:00
|
|
|
|
|
|
|
private
|
2023-01-09 13:03:21 -05:00
|
|
|
lookupBound : IsQty q => q -> Var n -> TyContext q d n -> InferResult q d n
|
|
|
|
lookupBound pi VZ ctx@(MkTyContext {tctx = _ :< ty, _}) =
|
|
|
|
InfRes {type = weakT ty, qout = zeroFor (tail ctx) :< pi}
|
2022-04-27 15:58:09 -04:00
|
|
|
lookupBound pi (VS i) ctx =
|
|
|
|
weakI $ lookupBound pi i (tail ctx)
|
2022-04-23 18:21:30 -04:00
|
|
|
|
2023-01-20 20:34:28 -05:00
|
|
|
private
|
|
|
|
lookupFree : IsQty q => CanTC q m => Name -> m (Definition q)
|
|
|
|
lookupFree x =
|
|
|
|
case lookup x !ask of
|
|
|
|
Just d => pure d
|
|
|
|
Nothing => throwError $ NotInScope x
|
2022-04-23 18:21:30 -04:00
|
|
|
|
2022-04-27 16:57:56 -04:00
|
|
|
private %inline
|
2023-01-08 14:44:25 -05:00
|
|
|
subjMult : IsQty q => (sg : SQty q) -> q -> SQty q
|
2023-01-09 13:03:21 -05:00
|
|
|
subjMult sg qty = if isYes $ isZero qty then szero else sg
|
|
|
|
|
|
|
|
|
|
|
|
export
|
|
|
|
makeDimEq : DContext d -> DimEq d
|
|
|
|
makeDimEq DNil = zeroEq
|
|
|
|
makeDimEq (DBind dctx) = makeDimEq dctx :<? Nothing
|
|
|
|
makeDimEq (DEq p q dctx) = set p q $ makeDimEq dctx
|
2022-04-27 16:57:56 -04:00
|
|
|
|
|
|
|
|
2023-01-08 14:44:25 -05:00
|
|
|
parameters {auto _ : IsQty q} {auto _ : CanTC q m}
|
2022-08-22 23:43:23 -04:00
|
|
|
mutual
|
2022-04-23 18:21:30 -04:00
|
|
|
-- [todo] it seems like the options here for dealing with substitutions are
|
|
|
|
-- to either push them or parametrise the whole typechecker over ambient
|
|
|
|
-- substitutions. both of them seem like the same amount of work for the
|
|
|
|
-- computer but pushing is less work for the me
|
|
|
|
|
2023-01-20 20:34:28 -05:00
|
|
|
||| `check ctx sg subj ty` checks that in the context `ctx`, the term
|
|
|
|
||| `subj` has the type `ty`, with quantity `sg`. if so, returns the
|
|
|
|
||| quantities of all bound variables that it used.
|
2022-04-23 18:21:30 -04:00
|
|
|
export covering %inline
|
2023-01-20 20:34:28 -05:00
|
|
|
check : TyContext q d n -> SQty q -> Term q d n -> Term q d n ->
|
2023-01-08 14:44:25 -05:00
|
|
|
m (CheckResult q n)
|
2023-01-20 20:34:28 -05:00
|
|
|
check ctx sg subj ty =
|
2023-01-22 18:53:34 -05:00
|
|
|
let Element subj nc = pushSubsts subj in
|
2023-01-20 20:34:28 -05:00
|
|
|
check' ctx sg subj nc ty
|
2022-04-23 18:21:30 -04:00
|
|
|
|
2023-01-20 20:34:28 -05:00
|
|
|
||| `infer ctx sg subj` infers the type of `subj` in the context `ctx`,
|
|
|
|
||| and returns its type and the bound variables it used.
|
2022-04-23 18:21:30 -04:00
|
|
|
export covering %inline
|
2023-01-20 20:34:28 -05:00
|
|
|
infer : TyContext q d n -> SQty q -> Elim q d n -> m (InferResult q d n)
|
|
|
|
infer ctx sg subj =
|
2023-01-22 18:53:34 -05:00
|
|
|
let Element subj nc = pushSubsts subj in
|
2023-01-20 20:34:28 -05:00
|
|
|
infer' ctx sg subj nc
|
2022-04-23 18:21:30 -04:00
|
|
|
|
|
|
|
|
|
|
|
export covering
|
2023-01-20 20:34:28 -05:00
|
|
|
check' : TyContext q d n -> SQty q ->
|
2023-01-22 18:53:34 -05:00
|
|
|
(subj : Term q d n) -> (0 nc : NotClo subj) -> Term q d n ->
|
2023-01-08 14:44:25 -05:00
|
|
|
m (CheckResult q n)
|
2022-04-23 18:21:30 -04:00
|
|
|
|
2023-01-20 20:34:28 -05:00
|
|
|
check' ctx sg (TYPE l) _ ty = do
|
2022-04-23 18:21:30 -04:00
|
|
|
l' <- expectTYPE ty
|
2023-01-08 14:44:25 -05:00
|
|
|
expectEqualQ zero sg.fst
|
2022-05-06 15:23:58 -04:00
|
|
|
unless (l < l') $ throwError $ BadUniverse l l'
|
2023-01-09 13:03:21 -05:00
|
|
|
pure $ zeroFor ctx
|
2022-04-23 18:21:30 -04:00
|
|
|
|
2023-01-20 20:34:28 -05:00
|
|
|
check' ctx sg (Pi qty x arg res) _ ty = do
|
2022-04-23 18:21:30 -04:00
|
|
|
l <- expectTYPE ty
|
2023-01-08 14:44:25 -05:00
|
|
|
expectEqualQ zero sg.fst
|
|
|
|
ignore $ check ctx szero arg (TYPE l)
|
2022-04-27 15:58:09 -04:00
|
|
|
case res of
|
2023-01-20 20:34:28 -05:00
|
|
|
TUsed res => ignore $ check (extendTy arg zero ctx) szero res (TYPE l)
|
|
|
|
TUnused res => ignore $ check ctx szero res (TYPE l)
|
2023-01-09 13:03:21 -05:00
|
|
|
pure $ zeroFor ctx
|
2022-04-27 15:58:09 -04:00
|
|
|
|
2023-01-20 20:34:28 -05:00
|
|
|
check' ctx sg (Lam x body) _ ty = do
|
2022-04-27 15:58:09 -04:00
|
|
|
(qty, arg, res) <- expectPi ty
|
2023-01-20 20:34:28 -05:00
|
|
|
qout <- check (extendTy arg (sg.fst * qty) ctx) sg body.term res.term
|
2022-04-27 15:58:09 -04:00
|
|
|
popQ qty qout
|
|
|
|
|
2023-01-20 20:34:28 -05:00
|
|
|
check' ctx sg (Eq i t l r) _ ty = do
|
|
|
|
u <- expectTYPE ty
|
|
|
|
expectEqualQ zero sg.fst
|
|
|
|
case t of
|
|
|
|
DUsed t => ignore $ check (extendDim ctx) sg t (TYPE u)
|
|
|
|
DUnused t => ignore $ check ctx sg t (TYPE u)
|
|
|
|
ignore $ check ctx sg t.zero l
|
|
|
|
ignore $ check ctx sg t.one r
|
|
|
|
pure $ zeroFor ctx
|
|
|
|
|
|
|
|
check' ctx sg (DLam i body) _ ty = do
|
|
|
|
(ty, l, r) <- expectEq ty
|
|
|
|
qout <- check (extendDim ctx) sg body.term ty.term
|
|
|
|
let eqs = makeDimEq ctx.dctx
|
2023-01-22 18:53:34 -05:00
|
|
|
equal !ask eqs body.zero l
|
|
|
|
equal !ask eqs body.one r
|
2023-01-20 20:34:28 -05:00
|
|
|
pure qout
|
|
|
|
|
|
|
|
check' ctx sg (E e) _ ty = do
|
2022-04-27 16:57:56 -04:00
|
|
|
infres <- infer ctx sg e
|
2023-01-08 14:44:25 -05:00
|
|
|
ignore $ check ctx szero ty (TYPE UAny)
|
2023-01-22 18:53:34 -05:00
|
|
|
sub !ask (makeDimEq ctx.dctx) infres.type ty
|
2022-04-27 15:58:09 -04:00
|
|
|
pure infres.qout
|
2022-04-23 18:21:30 -04:00
|
|
|
|
|
|
|
export covering
|
2023-01-20 20:34:28 -05:00
|
|
|
infer' : TyContext q d n -> SQty q ->
|
2023-01-22 18:53:34 -05:00
|
|
|
(subj : Elim q d n) -> (0 nc : NotClo subj) ->
|
2023-01-08 14:44:25 -05:00
|
|
|
m (InferResult q d n)
|
2022-04-23 18:21:30 -04:00
|
|
|
|
2023-01-20 20:34:28 -05:00
|
|
|
infer' ctx sg (F x) _ = do
|
|
|
|
g <- lookupFree x
|
|
|
|
when (isYes $ isZero g) $ expectEqualQ sg.fst zero
|
2023-01-09 13:03:21 -05:00
|
|
|
pure $ InfRes {type = g.type.get, qout = zeroFor ctx}
|
2022-04-23 18:21:30 -04:00
|
|
|
|
2023-01-20 20:34:28 -05:00
|
|
|
infer' ctx sg (B i) _ =
|
2023-01-08 14:44:25 -05:00
|
|
|
pure $ lookupBound sg.fst i ctx
|
2022-04-27 15:58:09 -04:00
|
|
|
|
2023-01-20 20:34:28 -05:00
|
|
|
infer' ctx sg (fun :@ arg) _ = do
|
2022-04-27 16:57:56 -04:00
|
|
|
funres <- infer ctx sg fun
|
2022-04-27 15:58:09 -04:00
|
|
|
(qty, argty, res) <- expectPi funres.type
|
2023-01-20 20:34:28 -05:00
|
|
|
argout <- check ctx (subjMult sg qty) arg argty
|
|
|
|
pure $ InfRes {
|
|
|
|
type = sub1 res $ arg :# argty,
|
|
|
|
qout = funres.qout + argout
|
|
|
|
}
|
|
|
|
|
|
|
|
infer' ctx sg (fun :% dim) _ = do
|
|
|
|
InfRes {type, qout} <- infer ctx sg fun
|
|
|
|
(ty, _, _) <- expectEq type
|
|
|
|
pure $ InfRes {type = dsub1 ty dim, qout}
|
|
|
|
|
|
|
|
infer' ctx sg (term :# type) _ = do
|
|
|
|
ignore $ check ctx szero type (TYPE UAny)
|
|
|
|
qout <- check ctx sg term type
|
|
|
|
pure $ InfRes {type, qout}
|