quox/lib/Quox/Typechecker.idr

281 lines
9.4 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
2023-01-26 13:54:46 -05:00
import Data.SnocVect
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
expectTYPE : CanTC' q _ m => Term q d n -> m Universe
2022-04-23 18:21:30 -04:00
expectTYPE s =
2023-02-10 15:40:44 -05:00
case whnfD !ask s of
Element (TYPE l) _ => pure l
_ => throwError $ ExpectedTYPE s
2022-04-23 18:21:30 -04:00
private covering %inline
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-02-10 15:40:44 -05:00
case whnfD !ask ty of
2023-01-20 20:34:28 -05:00
Element (Pi qty _ arg res) _ => pure (qty, arg, res)
_ => throwError $ ExpectedPi ty
2023-01-26 13:54:46 -05:00
private covering %inline
expectSig : CanTC' q _ m => Term q d n ->
m (Term q d n, ScopeTerm q d n)
expectSig ty =
2023-02-10 15:40:44 -05:00
case whnfD !ask ty of
2023-01-26 13:54:46 -05:00
Element (Sig _ arg res) _ => pure (arg, res)
_ => throwError $ ExpectedSig ty
2023-01-20 20:34:28 -05:00
private covering %inline
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-02-10 15:40:44 -05:00
case whnfD !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
2023-01-26 13:54:46 -05:00
private
popQs : HasErr q m => IsQty q =>
SnocVect s q -> QOutput q (s + n) -> m (QOutput q n)
popQs [<] qctx = pure qctx
popQs (pis :< pi) (qctx :< rh) = do expectCompatQ rh pi; popQs pis qctx
2022-04-23 18:21:30 -04:00
private %inline
2023-01-26 13:54:46 -05:00
popQ : HasErr q m => IsQty q => q -> QOutput q (S n) -> m (QOutput q n)
popQ pi = popQs [< pi]
2022-04-23 18:21:30 -04:00
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
2023-01-26 13:54:46 -05:00
lookupBound pi VZ (MkTyContext {tctx = tctx :< ty, _}) =
InfRes {type = weakT ty, qout = (zero <$ tctx) :< 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
2023-02-10 15:40:44 -05:00
lookupFree : CanTC' q g m => Name -> m (Definition' q g)
lookupFree x = lookupFree' !ask x
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
2023-01-08 14:44:25 -05:00
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
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 =
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-26 13:54:46 -05:00
||| `check0 ctx subj ty` checks a term in a zero context.
export covering %inline
check0 : TyContext q d n -> Term q d n -> Term q d n -> m (CheckResult q n)
check0 ctx = check (zeroed ctx) szero
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 =
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 ->
(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
2023-01-26 13:54:46 -05:00
-- if < ' then Ψ | Γ ⊢ Type · 0 ⇐ Type ' ⊳ 𝟎
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
2023-01-26 13:54:46 -05:00
check' ctx sg (Pi qty _ 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
2023-01-26 13:54:46 -05:00
-- if Ψ | Γ ⊢ A · 0 ⇐ Type 𝟎
ignore $ check0 ctx arg (TYPE l)
-- if Ψ | Γ, x · 0 : A ⊢ B · 0 ⇐ Type 𝟎
2022-04-27 15:58:09 -04:00
case res of
2023-01-26 13:54:46 -05:00
TUsed res => ignore $ check0 (extendTy arg zero ctx) res (TYPE l)
TUnused res => ignore $ check0 ctx res (TYPE l)
-- then Ψ | Γ ⊢ (x : A) → B · 0 ⇐ Type 𝟎
pure $ zeroFor ctx
2022-04-27 15:58:09 -04:00
2023-01-26 13:54:46 -05:00
check' ctx sg (Lam _ body) _ ty = do
2022-04-27 15:58:09 -04:00
(qty, arg, res) <- expectPi ty
2023-01-26 13:54:46 -05:00
-- if Ψ | Γ, x · πσ : A ⊢ t · σ ⇐ B ⊳ Σ, x · σπ
2023-01-20 20:34:28 -05:00
qout <- check (extendTy arg (sg.fst * qty) ctx) sg body.term res.term
2023-01-26 13:54:46 -05:00
-- then Ψ | Γ ⊢ λx. t · σ ⇐ (x · π : A) → B ⊳ Σ
popQ (sg.fst * qty) qout
check' ctx sg (Sig _ fst snd) _ ty = do
l <- expectTYPE ty
expectEqualQ zero sg.fst
-- if Ψ | Γ ⊢ A · 0 ⇐ Type 𝟎
ignore $ check0 ctx fst (TYPE l)
-- if Ψ | Γ, x · 0 : A ⊢ B · 0 ⇐ Type 𝟎
case snd of
TUsed snd => ignore $ check0 (extendTy fst zero ctx) snd (TYPE l)
TUnused snd => ignore $ check0 ctx snd (TYPE l)
-- then Ψ | Γ ⊢ (x : A) × B · 0 ⇐ Type 𝟎
pure $ zeroFor ctx
check' ctx sg (Pair fst snd) _ ty = do
(tfst, tsnd) <- expectSig ty
-- if Ψ | Γ ⊢ s · σ ⇐ A ⊳ Σ₁
qfst <- check ctx sg fst tfst
let tsnd = sub1 tsnd (fst :# tfst)
-- if Ψ | Γ ⊢ t · σ ⇐ B[s] ⊳ Σ₂
qsnd <- check ctx sg snd tsnd
-- then Ψ | Γ ⊢ (s, t) · σ ⇐ (x : A) × B ⊳ Σ₁ + Σ₂
pure $ qfst + qsnd
2022-04-27 15:58:09 -04:00
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
2023-01-26 13:54:46 -05:00
-- if Ψ, i | Γ ⊢ A · 0 ⇐ Type 𝟎
2023-01-20 20:34:28 -05:00
case t of
2023-01-26 13:54:46 -05:00
DUsed t => ignore $ check0 (extendDim ctx) t (TYPE u)
DUnused t => ignore $ check0 ctx t (TYPE u)
-- if Ψ | Γ ⊢ l · 0 ⇐ A0𝟎
ignore $ check0 ctx t.zero l
-- if Ψ | Γ ⊢ r · 0 ⇐ A1𝟎
ignore $ check0 ctx t.one r
-- then Ψ | Γ ⊢ Eq [i ⇒ A] l r ⇐ Type 𝟎
2023-01-20 20:34:28 -05:00
pure $ zeroFor ctx
check' ctx sg (DLam i body) _ ty = do
(ty, l, r) <- expectEq ty
2023-01-26 13:54:46 -05:00
-- if Ψ, i | Γ ⊢ t · σ ⇐ A ⊳ Σ
2023-01-20 20:34:28 -05:00
qout <- check (extendDim ctx) sg body.term ty.term
let eqs = makeDimEq ctx.dctx
2023-01-26 13:54:46 -05:00
-- if Ψ ⊢ t0 = l
2023-02-10 15:40:44 -05:00
equal eqs ctx.tctx ty.zero body.zero l
2023-01-26 13:54:46 -05:00
-- if Ψ ⊢ t1 = r
2023-02-10 15:40:44 -05:00
equal eqs ctx.tctx ty.one body.one r
2023-01-26 13:54:46 -05:00
-- then Ψ | Γ ⊢ (λᴰi ⇒ t) · σ ⇐ Eq [i ⇒ A] l r ⊳ Σ
2023-01-20 20:34:28 -05:00
pure qout
check' ctx sg (E e) _ ty = do
2023-01-26 13:54:46 -05:00
-- if Ψ | Γ ⊢ e · σ ⇒ A' ⊳ Σ
infres <- infer ctx sg e
2023-01-26 13:54:46 -05:00
-- if Ψ ⊢ A' <: A
2023-02-10 15:40:44 -05:00
subtype (makeDimEq ctx.dctx) ctx.tctx infres.type ty
2023-01-26 13:54:46 -05:00
-- then Ψ | Γ ⊢ e · σ ⇐ A ⊳ Σ
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 ->
(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
2023-01-26 13:54:46 -05:00
-- if x · π : A {≔ s} in global context
2023-01-20 20:34:28 -05:00
g <- lookupFree x
2023-01-26 13:54:46 -05:00
-- if σ ≤ π
expectCompatQ sg.fst g.qty
-- then Ψ | Γ ⊢ x ⇒ A ⊳ 𝟎
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-26 13:54:46 -05:00
-- if x : A ∈ Γ
-- then Ψ | Γ ⊢ x · σ ⇒ A ⊳ (𝟎, σ · x, 𝟎)
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
2023-01-26 13:54:46 -05:00
-- if Ψ | Γ ⊢ f · σ ⇒ (x · π : A) → B ⊳ Σ₁
funres <- infer ctx sg fun
2022-04-27 15:58:09 -04:00
(qty, argty, res) <- expectPi funres.type
2023-01-26 13:54:46 -05:00
-- if Ψ | Γ ⊢ s · σ∧π ⇐ A ⊳ Σ₂
-- (0∧π = σ∧0 = 0; σ∧π = σ otherwise)
2023-01-20 20:34:28 -05:00
argout <- check ctx (subjMult sg qty) arg argty
2023-01-26 13:54:46 -05:00
-- then Ψ | Γ ⊢ f s · σ ⇒ B[s] ⊳ Σ₁ + Σ₂
2023-01-20 20:34:28 -05:00
pure $ InfRes {
type = sub1 res $ arg :# argty,
qout = funres.qout + argout
}
2023-01-26 13:54:46 -05:00
infer' ctx sg (CasePair pi pair _ ret _ _ body) _ = do
-- if 1 ≤ π
expectCompatQ one pi
-- if Ψ | Γ ⊢ pair · 1 ⇒ (x : A) × B ⊳ Σ₁
pairres <- infer ctx sone pair
ignore $ check0 (extendTy pairres.type zero ctx) ret.term (TYPE UAny)
(tfst, tsnd) <- expectSig pairres.type
-- if Ψ | Γ, x · π : A, y · π : B ⊢ σ body ⇐ ret[(x, y)]
-- ⊳ Σ₂, x · π₁, y · π₂
-- if π₁, π₂ ≤ π
let bodyctx = extendTyN [< (tfst, pi), (tsnd.term, pi)] ctx
2023-02-10 15:40:44 -05:00
bodyty = substCasePairRet pairres.type ret
2023-01-26 13:54:46 -05:00
bodyout <- check bodyctx sg body.term bodyty >>= popQs [< pi, pi]
-- then Ψ | Γ ⊢ σ case ⋯ ⇒ ret[pair] ⊳ πΣ₁ + Σ₂
pure $ InfRes {
type = sub1 ret pair,
qout = pi * pairres.qout + bodyout
}
2023-01-20 20:34:28 -05:00
infer' ctx sg (fun :% dim) _ = do
2023-01-26 13:54:46 -05:00
-- if Ψ | Γ ⊢ f · σ ⇒ Eq [i ⇒ A] l r ⊳ Σ
2023-01-20 20:34:28 -05:00
InfRes {type, qout} <- infer ctx sg fun
(ty, _, _) <- expectEq type
2023-01-26 13:54:46 -05:00
-- then Ψ | Γ ⊢ f p · σ ⇒ Ap ⊳ Σ
2023-01-20 20:34:28 -05:00
pure $ InfRes {type = dsub1 ty dim, qout}
infer' ctx sg (term :# type) _ = do
2023-01-26 13:54:46 -05:00
-- if Ψ | Γ ⊢ A · 0 ⇐ Type 𝟎
ignore $ check0 ctx type (TYPE UAny)
-- if Ψ | Γ ⊢ s · σ ⇐ A ⊳ Σ
2023-01-20 20:34:28 -05:00
qout <- check ctx sg term type
2023-01-26 13:54:46 -05:00
-- then Ψ | Γ ⊢ (s ∷ A) · σ ⇒ A ⊳ Σ
2023-01-20 20:34:28 -05:00
pure $ InfRes {type, qout}