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
|
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
|
|
|
|
|
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)
|
2023-02-14 15:14:47 -05:00
|
|
|
|
popQs [<] qout = pure qout
|
|
|
|
|
popQs (pis :< pi) (qout :< rh) = do expectCompatQ rh pi; popQs pis qout
|
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
|
2023-02-14 15:14:47 -05:00
|
|
|
|
tail = {tctx $= tail}
|
2022-04-23 18:21:30 -04:00
|
|
|
|
|
|
|
|
|
|
|
|
|
|
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
|
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-02-14 15:14:47 -05:00
|
|
|
|
private %inline
|
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
|
|
|
|
|
2023-02-14 15:14:47 -05:00
|
|
|
|
||| ``sg `subjMult` pi`` is zero if either argument is, otherwise it
|
|
|
|
|
||| is equal to `sg`. (written as "σ ⨴ π" to emphasise its left bias.)
|
|
|
|
|
|||
|
|
|
|
|
||| only certain quantities can occur for a typing judgement to avoid losing
|
|
|
|
|
||| subject reduction [@qtt, §2.3], which is why this is not just `sg*pi`.
|
2022-04-27 16:57:56 -04:00
|
|
|
|
private %inline
|
2023-02-14 15:14:47 -05:00
|
|
|
|
subjMult : IsQty q => SQty q -> q -> SQty q
|
|
|
|
|
subjMult sg pi = if isYes $ isZero pi then szero else sg
|
2023-01-09 13:03:21 -05: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-26 13:54:46 -05:00
|
|
|
|
||| `check0 ctx subj ty` checks a term in a zero context.
|
|
|
|
|
export covering %inline
|
2023-02-13 16:06:03 -05:00
|
|
|
|
check0 : TyContext q d n -> Term q d n -> Term q d n -> m ()
|
2023-02-14 15:14:47 -05:00
|
|
|
|
check0 ctx tm ty = ignore $ check ctx szero tm ty
|
|
|
|
|
-- the output will always be 𝟎 because the subject quantity is 0
|
2023-01-26 13:54:46 -05: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-02-13 16:05:27 -05:00
|
|
|
|
check' ctx sg (TYPE k) _ ty = do
|
|
|
|
|
-- if 𝓀 < ℓ then Ψ | Γ ⊢₀ Type 𝓀 ⇐ Type ℓ
|
|
|
|
|
l <- expectTYPE !ask ty
|
2023-01-08 14:44:25 -05:00
|
|
|
|
expectEqualQ zero sg.fst
|
2023-02-13 16:05:27 -05:00
|
|
|
|
unless (k < l) $ throwError $ BadUniverse k l
|
2023-01-09 13:03:21 -05:00
|
|
|
|
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
|
2023-02-11 12:15:50 -05:00
|
|
|
|
l <- expectTYPE !ask ty
|
2023-01-08 14:44:25 -05:00
|
|
|
|
expectEqualQ zero sg.fst
|
2023-02-13 16:05:27 -05:00
|
|
|
|
-- if Ψ | Γ ⊢₀ A ⇐ Type ℓ
|
2023-02-13 16:06:03 -05:00
|
|
|
|
check0 ctx arg (TYPE l)
|
2023-02-13 16:05:27 -05:00
|
|
|
|
-- if Ψ | Γ, x : A ⊢₀ B ⇐ Type ℓ
|
2022-04-27 15:58:09 -04:00
|
|
|
|
case res of
|
2023-02-14 15:14:47 -05:00
|
|
|
|
TUsed res => check0 (extendTy arg ctx) res (TYPE l)
|
2023-02-13 16:06:03 -05:00
|
|
|
|
TUnused res => check0 ctx res (TYPE l)
|
2023-02-13 16:05:27 -05:00
|
|
|
|
-- then Ψ | Γ ⊢₀ (π·x : A) → B ⇐ Type ℓ
|
2023-01-09 13:03:21 -05:00
|
|
|
|
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
|
2023-02-11 12:15:50 -05:00
|
|
|
|
(qty, arg, res) <- expectPi !ask ty
|
2023-02-14 15:14:47 -05:00
|
|
|
|
-- if Ψ | Γ, x : A ⊢ σ · t ⇐ B ⊳ Σ, ρ·x
|
|
|
|
|
-- with ρ ≤ σπ
|
2023-02-13 16:06:03 -05:00
|
|
|
|
let qty' = sg.fst * qty
|
2023-02-14 15:14:47 -05:00
|
|
|
|
qout <- check (extendTy arg ctx) sg body.term res.term
|
2023-02-13 16:05:27 -05:00
|
|
|
|
-- then Ψ | Γ ⊢ σ · (λx ⇒ t) ⇐ (π·x : A) → B ⊳ Σ
|
2023-02-13 16:06:03 -05:00
|
|
|
|
popQ qty' qout
|
2023-01-26 13:54:46 -05:00
|
|
|
|
|
|
|
|
|
check' ctx sg (Sig _ fst snd) _ ty = do
|
2023-02-11 12:15:50 -05:00
|
|
|
|
l <- expectTYPE !ask ty
|
2023-01-26 13:54:46 -05:00
|
|
|
|
expectEqualQ zero sg.fst
|
2023-02-13 16:05:27 -05:00
|
|
|
|
-- if Ψ | Γ ⊢₀ A ⇐ Type ℓ
|
2023-02-13 16:06:03 -05:00
|
|
|
|
check0 ctx fst (TYPE l)
|
2023-02-13 16:05:27 -05:00
|
|
|
|
-- if Ψ | Γ, x : A ⊢₀ B ⇐ Type ℓ
|
2023-01-26 13:54:46 -05:00
|
|
|
|
case snd of
|
2023-02-14 15:14:47 -05:00
|
|
|
|
TUsed snd => check0 (extendTy fst ctx) snd (TYPE l)
|
2023-02-13 16:06:03 -05:00
|
|
|
|
TUnused snd => check0 ctx snd (TYPE l)
|
2023-02-13 16:05:27 -05:00
|
|
|
|
-- then Ψ | Γ ⊢₀ (x : A) × B ⇐ Type ℓ
|
2023-01-26 13:54:46 -05:00
|
|
|
|
pure $ zeroFor ctx
|
|
|
|
|
|
|
|
|
|
check' ctx sg (Pair fst snd) _ ty = do
|
2023-02-11 12:15:50 -05:00
|
|
|
|
(tfst, tsnd) <- expectSig !ask ty
|
2023-02-13 16:05:27 -05:00
|
|
|
|
-- if Ψ | Γ ⊢ σ · s ⇐ A ⊳ Σ₁
|
2023-01-26 13:54:46 -05:00
|
|
|
|
qfst <- check ctx sg fst tfst
|
|
|
|
|
let tsnd = sub1 tsnd (fst :# tfst)
|
2023-02-13 16:05:27 -05:00
|
|
|
|
-- if Ψ | Γ ⊢ σ · t ⇐ B[s] ⊳ Σ₂
|
2023-01-26 13:54:46 -05:00
|
|
|
|
qsnd <- check ctx sg snd tsnd
|
2023-02-13 16:05:27 -05:00
|
|
|
|
-- then Ψ | Γ ⊢ σ · (s, t) ⇐ (x : A) × B ⊳ Σ₁ + Σ₂
|
2023-01-26 13:54:46 -05:00
|
|
|
|
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
|
2023-02-11 12:15:50 -05:00
|
|
|
|
u <- expectTYPE !ask ty
|
2023-01-20 20:34:28 -05:00
|
|
|
|
expectEqualQ zero sg.fst
|
2023-02-13 16:05:27 -05:00
|
|
|
|
-- if Ψ, i | Γ ⊢₀ A ⇐ Type ℓ
|
2023-01-20 20:34:28 -05:00
|
|
|
|
case t of
|
2023-02-13 16:06:03 -05:00
|
|
|
|
DUsed t => check0 (extendDim ctx) t (TYPE u)
|
|
|
|
|
DUnused t => check0 ctx t (TYPE u)
|
2023-02-13 16:05:27 -05:00
|
|
|
|
-- if Ψ | Γ ⊢₀ l ⇐ A‹0›
|
2023-02-13 16:06:03 -05:00
|
|
|
|
check0 ctx t.zero l
|
2023-02-13 16:05:27 -05:00
|
|
|
|
-- if Ψ | Γ ⊢₀ r ⇐ A‹1›
|
2023-02-13 16:06:03 -05:00
|
|
|
|
check0 ctx t.one r
|
2023-02-13 16:05:27 -05:00
|
|
|
|
-- 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
|
2023-02-11 12:15:50 -05:00
|
|
|
|
(ty, l, r) <- expectEq !ask ty
|
2023-02-13 16:05:27 -05:00
|
|
|
|
-- if Ψ, i | Γ ⊢ σ · t ⇐ A ⊳ Σ
|
2023-01-20 20:34:28 -05:00
|
|
|
|
qout <- check (extendDim ctx) sg body.term ty.term
|
2023-02-13 16:05:27 -05:00
|
|
|
|
-- if Ψ | Γ ⊢ t‹0› = l : A‹0›
|
2023-02-14 15:29:04 -05:00
|
|
|
|
equal ctx.dctx ctx.tctx ty.zero body.zero l
|
2023-02-13 16:05:27 -05:00
|
|
|
|
-- if Ψ | Γ ⊢ t‹1› = r : A‹1›
|
2023-02-14 15:29:04 -05:00
|
|
|
|
equal ctx.dctx ctx.tctx ty.one body.one r
|
2023-02-13 16:05:27 -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-02-13 16:05:27 -05:00
|
|
|
|
-- if Ψ | Γ ⊢ σ · e ⇒ A' ⊳ Σ
|
2022-04-27 16:57:56 -04:00
|
|
|
|
infres <- infer ctx sg e
|
2023-02-13 16:05:27 -05:00
|
|
|
|
-- if Ψ | Γ ⊢ A' <: A
|
2023-02-14 15:29:04 -05:00
|
|
|
|
subtype ctx.dctx ctx.tctx infres.type ty
|
2023-02-13 16:05:27 -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 ->
|
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
|
2023-02-14 15:14:47 -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
|
2023-02-13 16:05:27 -05:00
|
|
|
|
-- then Ψ | Γ ⊢ σ · x ⇒ A ⊳ 𝟎
|
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-26 13:54:46 -05:00
|
|
|
|
-- if x : A ∈ Γ
|
2023-02-14 15:14:47 -05:00
|
|
|
|
-- 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-02-13 16:05:27 -05:00
|
|
|
|
-- if Ψ | Γ ⊢ σ · f ⇒ (π·x : A) → B ⊳ Σ₁
|
2022-04-27 16:57:56 -04:00
|
|
|
|
funres <- infer ctx sg fun
|
2023-02-11 12:15:50 -05:00
|
|
|
|
(qty, argty, res) <- expectPi !ask funres.type
|
2023-02-14 15:14:47 -05:00
|
|
|
|
-- if Ψ | Γ ⊢ σ ⨴ π · s ⇐ A ⊳ Σ₂
|
|
|
|
|
-- (σ ⨴ 0 = 0; σ ⨴ π = σ otherwise)
|
2023-01-20 20:34:28 -05:00
|
|
|
|
argout <- check ctx (subjMult sg qty) arg argty
|
2023-02-13 16:05:27 -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
|
2023-02-13 16:05:27 -05:00
|
|
|
|
-- if Ψ | Γ ⊢ 1 · pair ⇒ (x : A) × B ⊳ Σ₁
|
2023-01-26 13:54:46 -05:00
|
|
|
|
pairres <- infer ctx sone pair
|
2023-02-14 15:14:47 -05:00
|
|
|
|
check0 (extendTy pairres.type ctx) ret.term (TYPE UAny)
|
2023-02-11 12:15:50 -05:00
|
|
|
|
(tfst, tsnd) <- expectSig !ask pairres.type
|
2023-02-14 15:14:47 -05:00
|
|
|
|
-- if Ψ | Γ, x : A, y : B ⊢ σ · body ⇐ ret[(x, y)] ⊳ Σ₂, ρ₁·x, ρ₂·y
|
|
|
|
|
-- with ρ₁, ρ₂ ≤ π
|
|
|
|
|
let bodyctx = extendTyN [< tfst, tsnd.term] 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]
|
2023-02-13 16:05:27 -05:00
|
|
|
|
-- then Ψ | Γ ⊢ σ · case ⋯ ⇒ ret[pair] ⊳ πΣ₁ + Σ₂
|
2023-01-26 13:54:46 -05:00
|
|
|
|
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-02-13 16:05:27 -05:00
|
|
|
|
-- if Ψ | Γ ⊢ σ · f ⇒ Eq [i ⇒ A] l r ⊳ Σ
|
2023-01-20 20:34:28 -05:00
|
|
|
|
InfRes {type, qout} <- infer ctx sg fun
|
2023-02-11 12:15:50 -05:00
|
|
|
|
(ty, _, _) <- expectEq !ask type
|
2023-02-13 16:05:27 -05:00
|
|
|
|
-- then Ψ | Γ ⊢ σ · f p ⇒ A‹p› ⊳ Σ
|
2023-01-20 20:34:28 -05:00
|
|
|
|
pure $ InfRes {type = dsub1 ty dim, qout}
|
|
|
|
|
|
|
|
|
|
infer' ctx sg (term :# type) _ = do
|
2023-02-13 16:05:27 -05:00
|
|
|
|
-- if Ψ | Γ ⊢₀ A ⇐ Type ℓ
|
2023-02-13 16:06:03 -05:00
|
|
|
|
check0 ctx type (TYPE UAny)
|
2023-02-13 16:05:27 -05:00
|
|
|
|
-- if Ψ | Γ ⊢ σ · s ⇐ A ⊳ Σ
|
2023-01-20 20:34:28 -05:00
|
|
|
|
qout <- check ctx sg term type
|
2023-02-13 16:05:27 -05:00
|
|
|
|
-- then Ψ | Γ ⊢ σ · (s ∷ A) ⇒ A ⊳ Σ
|
2023-01-20 20:34:28 -05:00
|
|
|
|
pure $ InfRes {type, qout}
|