quox/lib/Quox/Typechecker.idr

156 lines
4.7 KiB
Idris
Raw Normal View History

2022-04-23 18:21:30 -04:00
module Quox.Typechecker
import public Quox.Syntax
import public Quox.Typing
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
private covering %inline
2023-01-08 14:44:25 -05:00
expectTYPE : HasErr q m => Term q d n -> m Universe
2022-04-23 18:21:30 -04:00
expectTYPE s =
case (whnfT s).fst of
TYPE l => pure l
_ => throwError $ ExpectedTYPE s
2022-04-23 18:21:30 -04:00
private covering %inline
2023-01-08 14:44:25 -05:00
expectPi : HasErr q m => Term q d n ->
m (q, Term q d n, ScopeTerm q d n)
2022-04-23 18:21:30 -04:00
expectPi ty =
case (whnfT ty).fst of
2022-04-27 15:58:09 -04:00
Pi qty _ arg res => pure (qty, arg, res)
_ => throwError $ ExpectedPi 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 ()
expectEqualQ pi rh =
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
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
private %inline
2023-01-08 14:44:25 -05:00
subjMult : IsQty q => (sg : SQty q) -> q -> SQty q
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
public export
2023-01-08 14:44:25 -05:00
0 CanTC' : (q : Type) -> (q -> Type) -> (Type -> Type) -> Type
CanTC' q isGlobal m = (HasErr q m, MonadReader (Definitions' q isGlobal) m)
2023-01-08 14:44:25 -05:00
public export
0 CanTC : (q : Type) -> IsQty q => (Type -> Type) -> Type
CanTC q = CanTC' q IsGlobal
parameters {auto _ : IsQty q} {auto _ : CanTC q m}
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
export covering %inline
check : (ctx : TyContext q d n) -> (sg : SQty q) ->
2023-01-08 14:44:25 -05:00
(subj : Term q d n) -> (ty : Term q d n) ->
m (CheckResult q n)
check ctx sg subj ty = check' ctx sg (pushSubstsT subj) ty
2022-04-23 18:21:30 -04:00
export covering %inline
infer : (ctx : TyContext q d n) -> (sg : SQty q) ->
2023-01-08 14:44:25 -05:00
(subj : Elim q d n) ->
m (InferResult q d n)
infer ctx sg subj = infer' ctx sg (pushSubstsE subj)
2022-04-23 18:21:30 -04:00
export covering
check' : (ctx : TyContext q d n) -> (sg : SQty q) ->
2023-01-08 14:44:25 -05:00
(subj : NotCloTerm q d n) -> (ty : Term q d n) ->
m (CheckResult q n)
2022-04-23 18:21:30 -04:00
check' ctx sg (Element (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
unless (l < l') $ throwError $ BadUniverse l l'
pure $ zeroFor ctx
2022-04-23 18:21:30 -04:00
check' ctx sg (Element (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-08 14:44:25 -05:00
TUsed res =>
ignore $ check (extendTy arg zero ctx) szero res (TYPE l)
TUnused res =>
ignore $ check ctx szero res (TYPE l)
pure $ zeroFor ctx
2022-04-27 15:58:09 -04:00
check' ctx sg (Element (Lam x body) _) ty = do
2022-04-27 15:58:09 -04:00
(qty, arg, res) <- expectPi ty
2022-04-23 18:21:30 -04:00
-- [todo] do this properly?
let body = fromScopeTerm body; res = fromScopeTerm res
2023-01-08 14:44:25 -05:00
qout <- check (extendTy arg (sg.fst * qty) ctx) sg body res
2022-04-27 15:58:09 -04:00
popQ qty qout
check' ctx sg (Element (E e) _) ty = do
infres <- infer ctx sg e
2023-01-08 14:44:25 -05:00
ignore $ check ctx szero ty (TYPE UAny)
subTWith (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
infer' : (ctx : TyContext q d n) -> (sg : SQty q) ->
2023-01-08 14:44:25 -05:00
(subj : NotCloElim q d n) ->
m (InferResult q d n)
2022-04-23 18:21:30 -04:00
infer' ctx sg (Element (F x) _) = do
Just g <- asks $ lookup x | Nothing => throwError $ NotInScope x
2023-01-08 14:44:25 -05:00
when (isZero g) $ expectEqualQ sg.fst zero
pure $ InfRes {type = g.type.get, qout = zeroFor ctx}
2022-04-23 18:21:30 -04:00
infer' ctx sg (Element (B i) _) =
2023-01-08 14:44:25 -05:00
pure $ lookupBound sg.fst i ctx
2022-04-27 15:58:09 -04:00
infer' ctx sg (Element (fun :@ arg) _) = do
funres <- infer ctx sg fun
2022-04-27 15:58:09 -04:00
(qty, argty, res) <- expectPi funres.type
2023-01-08 14:44:25 -05:00
let sg' = subjMult sg qty
argout <- check ctx sg' arg argty
2022-04-27 15:58:09 -04:00
pure $ InfRes {type = fromScopeTerm res //. ((arg :# argty) ::: id),
qout = funres.qout + argout}
infer' ctx sg (Element (tm :# ty) _) = do
2023-01-08 14:44:25 -05:00
ignore $ check ctx szero ty (TYPE UAny)
qout <- check ctx sg tm ty
2022-04-27 15:58:09 -04:00
pure $ InfRes {type = ty, qout}