2022-04-23 18:21:30 -04:00
|
|
|
|
module Quox.Typechecker
|
|
|
|
|
|
|
|
|
|
import public Quox.Typing
|
2022-08-22 23:43:23 -04:00
|
|
|
|
import public Quox.Equal
|
2022-04-23 18:21:30 -04:00
|
|
|
|
|
2023-03-13 22:22:26 -04:00
|
|
|
|
import Data.List
|
2023-03-26 10:09:47 -04:00
|
|
|
|
import Data.SnocVect
|
2023-03-26 18:01:32 -04:00
|
|
|
|
import Data.List1
|
2023-03-31 13:23:30 -04:00
|
|
|
|
import Quox.EffExtra
|
2023-03-13 22:22:26 -04:00
|
|
|
|
|
2022-04-23 18:21:30 -04:00
|
|
|
|
%default total
|
|
|
|
|
|
|
|
|
|
|
2023-01-20 20:34:28 -05:00
|
|
|
|
public export
|
2023-04-01 13:16:43 -04:00
|
|
|
|
0 TCEff : List (Type -> Type)
|
2023-04-18 16:55:23 -04:00
|
|
|
|
TCEff = [ErrorEff, DefsReader]
|
2023-01-20 20:34:28 -05:00
|
|
|
|
|
|
|
|
|
public export
|
2023-04-01 13:16:43 -04:00
|
|
|
|
0 TC : Type -> Type
|
|
|
|
|
TC = Eff TCEff
|
2023-03-31 13:23:30 -04:00
|
|
|
|
|
|
|
|
|
export
|
2023-04-18 16:55:23 -04:00
|
|
|
|
runTC : Definitions -> TC a -> Either Error a
|
|
|
|
|
runTC defs =
|
|
|
|
|
extract . runExcept . runReaderAt DEFS defs
|
2023-01-20 20:34:28 -05:00
|
|
|
|
|
|
|
|
|
|
2023-03-26 18:01:32 -04:00
|
|
|
|
export
|
2023-04-20 13:29:57 -04:00
|
|
|
|
popQs : Has ErrorEff fs => QContext s -> QOutput (s + n) -> Eff fs (QOutput 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
|
|
|
|
|
2023-03-26 18:01:32 -04:00
|
|
|
|
export %inline
|
2023-04-01 13:16:43 -04:00
|
|
|
|
popQ : Has ErrorEff fs => Qty -> QOutput (S n) -> Eff fs (QOutput n)
|
2023-01-26 13:54:46 -05:00
|
|
|
|
popQ pi = popQs [< pi]
|
|
|
|
|
|
2023-03-26 18:01:32 -04:00
|
|
|
|
export
|
2023-04-01 13:16:43 -04:00
|
|
|
|
lubs1 : List1 (QOutput n) -> Maybe (QOutput n)
|
2023-03-26 18:01:32 -04:00
|
|
|
|
lubs1 ([<] ::: _) = Just [<]
|
|
|
|
|
lubs1 ((qs :< p) ::: pqs) =
|
|
|
|
|
let (qss, ps) = unzip $ map unsnoc pqs in
|
|
|
|
|
[|lubs1 (qs ::: qss) :< foldlM lub p ps|]
|
|
|
|
|
|
|
|
|
|
export
|
2023-04-01 13:16:43 -04:00
|
|
|
|
lubs : TyContext d n -> List (QOutput n) -> Maybe (QOutput n)
|
2023-03-26 18:01:32 -04:00
|
|
|
|
lubs ctx [] = Just $ zeroFor ctx
|
|
|
|
|
lubs ctx (x :: xs) = lubs1 $ x ::: xs
|
|
|
|
|
|
2022-04-23 18:21:30 -04:00
|
|
|
|
|
2023-04-03 11:46:23 -04:00
|
|
|
|
||| context extension with no names or quantities
|
|
|
|
|
private
|
|
|
|
|
CtxExtension0' : Nat -> Nat -> Nat -> Type
|
|
|
|
|
CtxExtension0' s d n = Context (Term d . (+ n)) s
|
|
|
|
|
|
|
|
|
|
private
|
|
|
|
|
addNames0 : Context (Term d . (+ n)) s -> NContext s -> CtxExtension d n (s + n)
|
|
|
|
|
addNames0 [<] [<] = [<]
|
|
|
|
|
addNames0 (ts :< t) (xs :< x) = addNames0 ts xs :< (Zero, x, t)
|
|
|
|
|
|
2023-04-15 09:13:01 -04:00
|
|
|
|
export
|
|
|
|
|
typecaseTel : (k : TyConKind) -> Universe -> CtxExtension0' (arity k) d n
|
|
|
|
|
typecaseTel k u = case k of
|
|
|
|
|
KTYPE => [<]
|
|
|
|
|
-- A : ★ᵤ, B : 0.A → ★ᵤ
|
|
|
|
|
KPi => [< TYPE u, Arr Zero (BVT 0) (TYPE u)]
|
|
|
|
|
KSig => [< TYPE u, Arr Zero (BVT 0) (TYPE u)]
|
|
|
|
|
KEnum => [<]
|
|
|
|
|
-- A₀ : ★ᵤ, A₁ : ★ᵤ, A : (A₀ ≡ A₁ : ★ᵤ), L : A₀, R : A₀
|
|
|
|
|
KEq => [< TYPE u, TYPE u, Eq0 (TYPE u) (BVT 1) (BVT 0), BVT 2, BVT 2]
|
|
|
|
|
KNat => [<]
|
|
|
|
|
-- A : ★ᵤ
|
|
|
|
|
KBOX => [< TYPE u]
|
|
|
|
|
|
2023-04-03 11:46:23 -04:00
|
|
|
|
|
2023-04-01 13:16:43 -04:00
|
|
|
|
mutual
|
2023-02-19 11:04:57 -05:00
|
|
|
|
||| "Ψ | Γ ⊢ σ · s ⇐ A ⊳ Σ"
|
|
|
|
|
|||
|
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.
|
2023-02-19 11:51:44 -05:00
|
|
|
|
|||
|
|
|
|
|
||| if the dimension context is inconsistent, then return `Nothing`, without
|
|
|
|
|
||| doing any further work.
|
2022-04-23 18:21:30 -04:00
|
|
|
|
export covering %inline
|
2023-04-01 13:16:43 -04:00
|
|
|
|
check : (ctx : TyContext d n) -> SQty -> Term d n -> Term d n ->
|
|
|
|
|
TC (CheckResult ctx.dctx n)
|
2023-02-20 15:38:47 -05:00
|
|
|
|
check ctx sg subj ty = ifConsistent ctx.dctx $ checkC ctx sg subj ty
|
2022-04-23 18:21:30 -04:00
|
|
|
|
|
2023-02-19 11:04:57 -05:00
|
|
|
|
||| "Ψ | Γ ⊢₀ s ⇐ A"
|
|
|
|
|
|||
|
|
|
|
|
||| `check0 ctx subj ty` checks a term (as `check`) in a zero context.
|
2023-01-26 13:54:46 -05:00
|
|
|
|
export covering %inline
|
2023-04-01 13:16:43 -04:00
|
|
|
|
check0 : TyContext d n -> Term d n -> Term d n -> TC ()
|
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-02-19 11:51:44 -05:00
|
|
|
|
||| `check`, assuming the dimension context is consistent
|
|
|
|
|
export covering %inline
|
2023-04-01 13:16:43 -04:00
|
|
|
|
checkC : (ctx : TyContext d n) -> SQty -> Term d n -> Term d n ->
|
|
|
|
|
TC (CheckResult' n)
|
2023-02-19 11:51:44 -05:00
|
|
|
|
checkC ctx sg subj ty =
|
2023-02-19 11:54:39 -05:00
|
|
|
|
wrapErr (WhileChecking ctx sg.fst subj ty) $
|
|
|
|
|
let Element subj nc = pushSubsts subj in
|
2023-02-20 15:42:21 -05:00
|
|
|
|
check' ctx sg subj ty
|
2023-02-19 11:51:44 -05:00
|
|
|
|
|
2023-03-05 07:14:25 -05:00
|
|
|
|
||| "Ψ | Γ ⊢₀ s ⇐ ★ᵢ"
|
|
|
|
|
|||
|
|
|
|
|
||| `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.
|
|
|
|
|
export covering %inline
|
2023-04-01 13:16:43 -04:00
|
|
|
|
checkType : TyContext d n -> Term d n -> Maybe Universe -> TC ()
|
2023-03-05 07:14:25 -05:00
|
|
|
|
checkType ctx subj l = ignore $ ifConsistent ctx.dctx $ checkTypeC ctx subj l
|
|
|
|
|
|
|
|
|
|
export covering %inline
|
2023-04-01 13:16:43 -04:00
|
|
|
|
checkTypeC : TyContext d n -> Term d n -> Maybe Universe -> TC ()
|
2023-03-05 07:14:25 -05:00
|
|
|
|
checkTypeC ctx subj l =
|
|
|
|
|
wrapErr (WhileCheckingTy ctx subj l) $ checkTypeNoWrap ctx subj l
|
|
|
|
|
|
|
|
|
|
export covering %inline
|
2023-04-01 13:16:43 -04:00
|
|
|
|
checkTypeNoWrap : TyContext d n -> Term d n -> Maybe Universe -> TC ()
|
2023-03-05 07:14:25 -05:00
|
|
|
|
checkTypeNoWrap ctx subj l =
|
|
|
|
|
let Element subj nc = pushSubsts subj in
|
|
|
|
|
checkType' ctx subj l
|
2023-02-19 11:51:44 -05:00
|
|
|
|
|
2023-02-19 11:04:57 -05:00
|
|
|
|
||| "Ψ | Γ ⊢ σ · e ⇒ A ⊳ Σ"
|
|
|
|
|
|||
|
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.
|
2023-02-19 11:51:44 -05:00
|
|
|
|
|||
|
|
|
|
|
||| if the dimension context is inconsistent, then return `Nothing`, without
|
|
|
|
|
||| doing any further work.
|
|
|
|
|
export covering %inline
|
2023-04-01 13:16:43 -04:00
|
|
|
|
infer : (ctx : TyContext d n) -> SQty -> Elim d n ->
|
|
|
|
|
TC (InferResult ctx.dctx d n)
|
2023-02-20 15:38:47 -05:00
|
|
|
|
infer ctx sg subj = ifConsistent ctx.dctx $ inferC ctx sg subj
|
2023-02-19 11:51:44 -05:00
|
|
|
|
|
|
|
|
|
||| `infer`, assuming the dimension context is consistent
|
2022-04-23 18:21:30 -04:00
|
|
|
|
export covering %inline
|
2023-04-01 13:16:43 -04:00
|
|
|
|
inferC : (ctx : TyContext d n) -> SQty -> Elim d n ->
|
|
|
|
|
TC (InferResult' d n)
|
2023-02-19 11:51:44 -05:00
|
|
|
|
inferC ctx sg subj =
|
2023-02-19 11:54:39 -05:00
|
|
|
|
wrapErr (WhileInferring ctx sg.fst subj) $
|
|
|
|
|
let Element subj nc = pushSubsts subj in
|
2023-02-20 15:42:21 -05:00
|
|
|
|
infer' ctx sg subj
|
2022-04-23 18:21:30 -04:00
|
|
|
|
|
|
|
|
|
|
2023-03-05 07:14:25 -05:00
|
|
|
|
private covering
|
2023-04-01 13:16:43 -04:00
|
|
|
|
toCheckType : TyContext d n -> SQty ->
|
|
|
|
|
(subj : Term d n) -> (0 nc : NotClo subj) => Term d n ->
|
|
|
|
|
TC (CheckResult' n)
|
2023-03-05 07:14:25 -05:00
|
|
|
|
toCheckType ctx sg t ty = do
|
2023-04-18 16:55:23 -04:00
|
|
|
|
u <- expectTYPE !defs ctx ty
|
2023-04-01 13:16:43 -04:00
|
|
|
|
expectEqualQ Zero sg.fst
|
2023-03-05 07:14:25 -05:00
|
|
|
|
checkTypeNoWrap ctx t (Just u)
|
|
|
|
|
pure $ zeroFor ctx
|
|
|
|
|
|
2023-02-19 11:51:44 -05:00
|
|
|
|
private covering
|
2023-04-01 13:16:43 -04:00
|
|
|
|
check' : TyContext d n -> SQty ->
|
|
|
|
|
(subj : Term d n) -> (0 nc : NotClo subj) => Term d n ->
|
|
|
|
|
TC (CheckResult' n)
|
2022-04-23 18:21:30 -04:00
|
|
|
|
|
2023-03-05 07:14:25 -05:00
|
|
|
|
check' ctx sg t@(TYPE _) ty = toCheckType ctx sg t ty
|
2022-04-23 18:21:30 -04:00
|
|
|
|
|
2023-03-05 07:14:25 -05:00
|
|
|
|
check' ctx sg t@(Pi {}) ty = toCheckType ctx sg t ty
|
2022-04-27 15:58:09 -04:00
|
|
|
|
|
2023-02-22 01:40:19 -05:00
|
|
|
|
check' ctx sg (Lam body) ty = do
|
2023-04-18 16:55:23 -04:00
|
|
|
|
(qty, arg, res) <- expectPi !defs ctx 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-03-15 10:54:51 -04:00
|
|
|
|
qout <- checkC (extendTy qty' body.name 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
|
|
|
|
|
2023-03-05 07:14:25 -05:00
|
|
|
|
check' ctx sg t@(Sig {}) ty = toCheckType ctx sg t ty
|
2023-01-26 13:54:46 -05:00
|
|
|
|
|
2023-02-20 15:42:21 -05:00
|
|
|
|
check' ctx sg (Pair fst snd) ty = do
|
2023-04-18 16:55:23 -04:00
|
|
|
|
(tfst, tsnd) <- expectSig !defs ctx ty
|
2023-02-13 16:05:27 -05:00
|
|
|
|
-- if Ψ | Γ ⊢ σ · s ⇐ A ⊳ Σ₁
|
2023-02-19 11:51:44 -05:00
|
|
|
|
qfst <- checkC ctx sg fst tfst
|
2023-01-26 13:54:46 -05:00
|
|
|
|
let tsnd = sub1 tsnd (fst :# tfst)
|
2023-02-13 16:05:27 -05:00
|
|
|
|
-- if Ψ | Γ ⊢ σ · t ⇐ B[s] ⊳ Σ₂
|
2023-02-19 11:51:44 -05:00
|
|
|
|
qsnd <- checkC 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-03-05 07:14:25 -05:00
|
|
|
|
check' ctx sg t@(Enum _) ty = toCheckType ctx sg t ty
|
2023-02-22 01:45:10 -05:00
|
|
|
|
|
|
|
|
|
check' ctx sg (Tag t) ty = do
|
2023-04-18 16:55:23 -04:00
|
|
|
|
tags <- expectEnum !defs ctx ty
|
2023-02-22 01:45:10 -05:00
|
|
|
|
-- if t ∈ ts
|
2023-03-31 13:23:30 -04:00
|
|
|
|
unless (t `elem` tags) $ throw $ TagNotIn t tags
|
2023-02-22 01:45:10 -05:00
|
|
|
|
-- then Ψ | Γ ⊢ σ · t ⇐ {ts} ⊳ 𝟎
|
|
|
|
|
pure $ zeroFor ctx
|
|
|
|
|
|
2023-03-05 07:14:25 -05:00
|
|
|
|
check' ctx sg t@(Eq {}) ty = toCheckType ctx sg t ty
|
2023-01-20 20:34:28 -05:00
|
|
|
|
|
2023-02-22 01:40:19 -05:00
|
|
|
|
check' ctx sg (DLam body) ty = do
|
2023-04-18 16:55:23 -04:00
|
|
|
|
(ty, l, r) <- expectEq !defs ctx ty
|
2023-04-17 15:55:12 -04:00
|
|
|
|
let ctx' = extendDim body.name ctx
|
|
|
|
|
ty = ty.term
|
|
|
|
|
body = body.term
|
2023-02-13 16:05:27 -05:00
|
|
|
|
-- if Ψ, i | Γ ⊢ σ · t ⇐ A ⊳ Σ
|
2023-04-17 15:55:12 -04:00
|
|
|
|
qout <- checkC ctx' sg body ty
|
|
|
|
|
-- if Ψ, i, i = 0 | Γ ⊢ t = l : A
|
|
|
|
|
equal (eqDim (BV 0) (K Zero) ctx') ty body (dweakT 1 l)
|
|
|
|
|
-- if Ψ, i, i = 1 | Γ ⊢ t = r : A
|
|
|
|
|
equal (eqDim (BV 0) (K One) ctx') ty body (dweakT 1 r)
|
2023-02-25 13:14:11 -05:00
|
|
|
|
-- then Ψ | Γ ⊢ σ · (δ i ⇒ t) ⇐ Eq [i ⇒ A] l r ⊳ Σ
|
2023-01-20 20:34:28 -05:00
|
|
|
|
pure qout
|
|
|
|
|
|
2023-03-26 08:40:54 -04:00
|
|
|
|
check' ctx sg Nat ty = toCheckType ctx sg Nat ty
|
|
|
|
|
|
|
|
|
|
check' ctx sg Zero ty = do
|
2023-04-18 16:55:23 -04:00
|
|
|
|
expectNat !defs ctx ty
|
2023-03-26 08:40:54 -04:00
|
|
|
|
pure $ zeroFor ctx
|
|
|
|
|
|
|
|
|
|
check' ctx sg (Succ n) ty = do
|
2023-04-18 16:55:23 -04:00
|
|
|
|
expectNat !defs ctx ty
|
2023-03-26 08:40:54 -04:00
|
|
|
|
checkC ctx sg n Nat
|
|
|
|
|
|
2023-03-31 13:11:35 -04:00
|
|
|
|
check' ctx sg t@(BOX {}) ty = toCheckType ctx sg t ty
|
|
|
|
|
|
|
|
|
|
check' ctx sg (Box val) ty = do
|
2023-04-18 16:55:23 -04:00
|
|
|
|
(q, ty) <- expectBOX !defs ctx ty
|
2023-03-31 13:11:35 -04:00
|
|
|
|
-- if Ψ | Γ ⊢ σ · s ⇐ A ⊳ Σ
|
|
|
|
|
valout <- checkC ctx sg val ty
|
|
|
|
|
-- then Ψ | Γ ⊢ σ · [s] ⇐ [π.A] ⊳ πΣ
|
|
|
|
|
pure $ q * valout
|
|
|
|
|
|
2023-02-20 15:42:21 -05:00
|
|
|
|
check' ctx sg (E e) ty = do
|
2023-02-13 16:05:27 -05:00
|
|
|
|
-- if Ψ | Γ ⊢ σ · e ⇒ A' ⊳ Σ
|
2023-02-19 11:51:44 -05:00
|
|
|
|
infres <- inferC ctx sg e
|
2023-02-13 16:05:27 -05:00
|
|
|
|
-- if Ψ | Γ ⊢ A' <: A
|
2023-02-14 16:28:10 -05:00
|
|
|
|
subtype ctx 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
|
|
|
|
|
2023-03-05 07:14:25 -05:00
|
|
|
|
private covering
|
2023-04-01 13:16:43 -04:00
|
|
|
|
checkType' : TyContext d n ->
|
|
|
|
|
(subj : Term d n) -> (0 nc : NotClo subj) =>
|
|
|
|
|
Maybe Universe -> TC ()
|
2023-03-05 07:14:25 -05:00
|
|
|
|
|
|
|
|
|
checkType' ctx (TYPE k) u = do
|
|
|
|
|
-- if 𝓀 < ℓ then Ψ | Γ ⊢₀ Type 𝓀 ⇐ Type ℓ
|
|
|
|
|
case u of
|
2023-03-31 13:23:30 -04:00
|
|
|
|
Just l => unless (k < l) $ throw $ BadUniverse k l
|
2023-03-05 07:14:25 -05:00
|
|
|
|
Nothing => pure ()
|
|
|
|
|
|
|
|
|
|
checkType' ctx (Pi qty arg res) u = do
|
|
|
|
|
-- if Ψ | Γ ⊢₀ A ⇐ Type ℓ
|
|
|
|
|
checkTypeC ctx arg u
|
|
|
|
|
-- if Ψ | Γ, x : A ⊢₀ B ⇐ Type ℓ
|
2023-04-03 11:46:23 -04:00
|
|
|
|
checkTypeScope ctx arg res u
|
2023-03-05 07:14:25 -05:00
|
|
|
|
-- then Ψ | Γ ⊢₀ (π·x : A) → B ⇐ Type ℓ
|
|
|
|
|
|
|
|
|
|
checkType' ctx t@(Lam {}) u =
|
2023-03-31 13:23:30 -04:00
|
|
|
|
throw $ NotType ctx t
|
2023-03-05 07:14:25 -05:00
|
|
|
|
|
|
|
|
|
checkType' ctx (Sig fst snd) u = do
|
|
|
|
|
-- if Ψ | Γ ⊢₀ A ⇐ Type ℓ
|
|
|
|
|
checkTypeC ctx fst u
|
|
|
|
|
-- if Ψ | Γ, x : A ⊢₀ B ⇐ Type ℓ
|
2023-04-03 11:46:23 -04:00
|
|
|
|
checkTypeScope ctx fst snd u
|
2023-03-05 07:14:25 -05:00
|
|
|
|
-- then Ψ | Γ ⊢₀ (x : A) × B ⇐ Type ℓ
|
|
|
|
|
|
|
|
|
|
checkType' ctx t@(Pair {}) u =
|
2023-03-31 13:23:30 -04:00
|
|
|
|
throw $ NotType ctx t
|
2023-03-05 07:14:25 -05:00
|
|
|
|
|
|
|
|
|
checkType' ctx (Enum _) u = pure ()
|
|
|
|
|
-- Ψ | Γ ⊢₀ {ts} ⇐ Type ℓ
|
|
|
|
|
|
|
|
|
|
checkType' ctx t@(Tag {}) u =
|
2023-03-31 13:23:30 -04:00
|
|
|
|
throw $ NotType ctx t
|
2023-03-05 07:14:25 -05:00
|
|
|
|
|
|
|
|
|
checkType' ctx (Eq t l r) u = do
|
|
|
|
|
-- if Ψ, i | Γ ⊢₀ A ⇐ Type ℓ
|
|
|
|
|
case t.body of
|
2023-03-13 22:22:26 -04:00
|
|
|
|
Y t' => checkTypeC (extendDim t.name ctx) t' u
|
|
|
|
|
N t' => checkTypeC ctx t' u
|
2023-03-05 07:14:25 -05:00
|
|
|
|
-- if Ψ | Γ ⊢₀ l ⇐ A‹0›
|
2023-04-01 10:02:02 -04:00
|
|
|
|
check0 ctx l t.zero
|
2023-03-05 07:14:25 -05:00
|
|
|
|
-- if Ψ | Γ ⊢₀ r ⇐ A‹1›
|
2023-04-01 10:02:02 -04:00
|
|
|
|
check0 ctx r t.one
|
2023-03-05 07:14:25 -05:00
|
|
|
|
-- then Ψ | Γ ⊢₀ Eq [i ⇒ A] l r ⇐ Type ℓ
|
|
|
|
|
|
|
|
|
|
checkType' ctx t@(DLam {}) u =
|
2023-03-31 13:23:30 -04:00
|
|
|
|
throw $ NotType ctx t
|
2023-03-05 07:14:25 -05:00
|
|
|
|
|
2023-03-26 08:40:54 -04:00
|
|
|
|
checkType' ctx Nat u = pure ()
|
2023-03-31 13:23:30 -04:00
|
|
|
|
checkType' ctx Zero u = throw $ NotType ctx Zero
|
|
|
|
|
checkType' ctx t@(Succ _) u = throw $ NotType ctx t
|
2023-03-26 08:40:54 -04:00
|
|
|
|
|
2023-03-31 13:11:35 -04:00
|
|
|
|
checkType' ctx (BOX q ty) u = checkType ctx ty u
|
2023-03-31 13:23:30 -04:00
|
|
|
|
checkType' ctx t@(Box _) u = throw $ NotType ctx t
|
2023-03-31 13:11:35 -04:00
|
|
|
|
|
2023-03-05 07:14:25 -05:00
|
|
|
|
checkType' ctx (E e) u = do
|
2023-03-31 13:11:35 -04:00
|
|
|
|
-- if Ψ | Γ ⊢₀ E ⇒ Type ℓ
|
2023-03-05 07:14:25 -05:00
|
|
|
|
infres <- inferC ctx szero e
|
2023-03-31 13:11:35 -04:00
|
|
|
|
-- if Ψ | Γ ⊢ Type ℓ <: Type 𝓀
|
2023-03-05 07:14:25 -05:00
|
|
|
|
case u of
|
|
|
|
|
Just u => subtype ctx infres.type (TYPE u)
|
2023-04-17 17:58:24 -04:00
|
|
|
|
Nothing => ignore $
|
2023-04-18 16:55:23 -04:00
|
|
|
|
expectTYPE !defs ctx infres.type
|
2023-03-31 13:11:35 -04:00
|
|
|
|
-- then Ψ | Γ ⊢₀ E ⇐ Type 𝓀
|
2023-03-05 07:14:25 -05:00
|
|
|
|
|
2023-04-03 11:46:23 -04:00
|
|
|
|
private covering
|
|
|
|
|
check0ScopeN : {s : Nat} ->
|
|
|
|
|
TyContext d n -> CtxExtension0' s d n ->
|
|
|
|
|
ScopeTermN s d n -> Term d n -> TC ()
|
|
|
|
|
check0ScopeN ctx ext (S _ (N body)) ty = check0 ctx body ty
|
|
|
|
|
check0ScopeN ctx ext (S names (Y body)) ty =
|
2023-04-15 09:13:01 -04:00
|
|
|
|
check0 (extendTyN (addNames0 ext names) ctx) body (weakT s ty)
|
2023-04-03 11:46:23 -04:00
|
|
|
|
|
|
|
|
|
private covering
|
|
|
|
|
check0Scope : TyContext d n -> Term d n ->
|
|
|
|
|
ScopeTerm d n -> Term d n -> TC ()
|
|
|
|
|
check0Scope ctx t = check0ScopeN ctx [< t]
|
|
|
|
|
|
|
|
|
|
|
|
|
|
|
private covering
|
|
|
|
|
checkTypeScopeN : TyContext d n -> CtxExtension0' s d n ->
|
|
|
|
|
ScopeTermN s d n -> Maybe Universe -> TC ()
|
|
|
|
|
checkTypeScopeN ctx ext (S _ (N body)) u = checkType ctx body u
|
|
|
|
|
checkTypeScopeN ctx ext (S names (Y body)) u =
|
|
|
|
|
checkType (extendTyN (addNames0 ext names) ctx) body u
|
|
|
|
|
|
|
|
|
|
private covering
|
|
|
|
|
checkTypeScope : TyContext d n -> Term d n ->
|
|
|
|
|
ScopeTerm d n -> Maybe Universe -> TC ()
|
|
|
|
|
checkTypeScope ctx s = checkTypeScopeN ctx [< s]
|
|
|
|
|
|
2023-03-05 07:14:25 -05:00
|
|
|
|
|
2023-02-19 11:51:44 -05:00
|
|
|
|
private covering
|
2023-04-01 13:16:43 -04:00
|
|
|
|
infer' : TyContext d n -> SQty ->
|
|
|
|
|
(subj : Elim d n) -> (0 nc : NotClo subj) =>
|
|
|
|
|
TC (InferResult' d n)
|
2022-04-23 18:21:30 -04:00
|
|
|
|
|
2023-02-20 15:42:21 -05:00
|
|
|
|
infer' ctx sg (F x) = do
|
2023-02-14 15:14:47 -05:00
|
|
|
|
-- if π·x : A {≔ s} in global context
|
2023-04-18 16:55:23 -04:00
|
|
|
|
g <- lookupFree x !defs
|
2023-01-26 13:54:46 -05:00
|
|
|
|
-- if σ ≤ π
|
2023-04-01 13:16:43 -04:00
|
|
|
|
expectCompatQ sg.fst g.qty.fst
|
2023-02-13 16:05:27 -05:00
|
|
|
|
-- then Ψ | Γ ⊢ σ · x ⇒ A ⊳ 𝟎
|
2023-04-03 11:46:23 -04:00
|
|
|
|
let Val d = ctx.dimLen; Val n = ctx.termLen
|
|
|
|
|
pure $ InfRes {type = g.type, qout = zeroFor ctx}
|
2022-04-23 18:21:30 -04:00
|
|
|
|
|
2023-02-20 15:42:21 -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-02-19 11:51:44 -05:00
|
|
|
|
pure $ lookupBound sg.fst i ctx.tctx
|
2023-02-22 01:40:19 -05:00
|
|
|
|
where
|
2023-04-01 13:16:43 -04:00
|
|
|
|
lookupBound : forall n. Qty -> Var n -> TContext d n -> InferResult' d n
|
2023-04-15 09:13:01 -04:00
|
|
|
|
lookupBound pi VZ (ctx :< type) =
|
|
|
|
|
InfRes {type = weakT 1 type, qout = zeroFor ctx :< pi}
|
2023-02-22 01:40:19 -05:00
|
|
|
|
lookupBound pi (VS i) (ctx :< _) =
|
|
|
|
|
let InfRes {type, qout} = lookupBound pi i ctx in
|
2023-04-15 09:13:01 -04:00
|
|
|
|
InfRes {type = weakT 1 type, qout = qout :< Zero}
|
2022-04-27 15:58:09 -04:00
|
|
|
|
|
2023-02-20 15:42:21 -05:00
|
|
|
|
infer' ctx sg (fun :@ arg) = do
|
2023-02-13 16:05:27 -05:00
|
|
|
|
-- if Ψ | Γ ⊢ σ · f ⇒ (π·x : A) → B ⊳ Σ₁
|
2023-02-19 11:51:44 -05:00
|
|
|
|
funres <- inferC ctx sg fun
|
2023-04-18 16:55:23 -04:00
|
|
|
|
(qty, argty, res) <- expectPi !defs ctx funres.type
|
2023-02-14 15:14:47 -05:00
|
|
|
|
-- if Ψ | Γ ⊢ σ ⨴ π · s ⇐ A ⊳ Σ₂
|
2023-02-19 11:51:44 -05:00
|
|
|
|
argout <- checkC ctx (subjMult sg qty) arg argty
|
2023-04-20 13:29:57 -04:00
|
|
|
|
-- then Ψ | Γ ⊢ σ · f s ⇒ B[s] ⊳ Σ₁ + πΣ₂
|
2023-01-20 20:34:28 -05:00
|
|
|
|
pure $ InfRes {
|
|
|
|
|
type = sub1 res $ arg :# argty,
|
2023-04-20 13:29:57 -04:00
|
|
|
|
qout = funres.qout + qty * argout
|
2023-01-20 20:34:28 -05:00
|
|
|
|
}
|
|
|
|
|
|
2023-02-22 01:40:19 -05:00
|
|
|
|
infer' ctx sg (CasePair pi pair ret body) = do
|
2023-03-26 18:08:09 -04:00
|
|
|
|
-- no check for 1 ≤ π, since pairs have a single constructor.
|
|
|
|
|
-- e.g. at 0 the components are also 0 in the body
|
|
|
|
|
--
|
2023-02-22 01:45:10 -05:00
|
|
|
|
-- if Ψ | Γ ⊢ σ · pair ⇒ (x : A) × B ⊳ Σ₁
|
|
|
|
|
pairres <- inferC ctx sg pair
|
|
|
|
|
-- if Ψ | Γ, p : (x : A) × B ⊢₀ ret ⇐ Type
|
2023-04-01 13:16:43 -04:00
|
|
|
|
checkTypeC (extendTy Zero ret.name pairres.type ctx) ret.term Nothing
|
2023-04-18 16:55:23 -04:00
|
|
|
|
(tfst, tsnd) <- expectSig !defs ctx pairres.type
|
2023-02-22 01:45:10 -05:00
|
|
|
|
-- if Ψ | Γ, x : A, y : B ⊢ σ · body ⇐
|
|
|
|
|
-- ret[(x, y) ∷ (x : A) × B/p] ⊳ Σ₂, ρ₁·x, ρ₂·y
|
2023-02-23 04:04:00 -05:00
|
|
|
|
-- with ρ₁, ρ₂ ≤ πσ
|
2023-03-16 13:18:49 -04:00
|
|
|
|
let [< x, y] = body.names
|
2023-02-23 04:04:00 -05:00
|
|
|
|
pisg = pi * sg.fst
|
2023-03-15 10:54:51 -04:00
|
|
|
|
bodyctx = extendTyN [< (pisg, x, tfst), (pisg, y, tsnd.term)] ctx
|
|
|
|
|
bodyty = substCasePairRet pairres.type ret
|
2023-03-31 13:11:35 -04:00
|
|
|
|
bodyout <- checkC bodyctx sg body.term bodyty >>= popQs [< pisg, pisg]
|
2023-04-15 09:13:01 -04:00
|
|
|
|
-- then Ψ | Γ ⊢ σ · caseπ ⋯ ⇒ ret[pair/p] ⊳ πΣ₁ + Σ₂
|
2023-01-26 13:54:46 -05:00
|
|
|
|
pure $ InfRes {
|
|
|
|
|
type = sub1 ret pair,
|
2023-03-31 13:11:35 -04:00
|
|
|
|
qout = pi * pairres.qout + bodyout
|
2023-01-26 13:54:46 -05:00
|
|
|
|
}
|
|
|
|
|
|
2023-03-13 22:22:26 -04:00
|
|
|
|
infer' ctx sg (CaseEnum pi t ret arms) {d, n} = do
|
2023-02-22 01:45:10 -05:00
|
|
|
|
-- if Ψ | Γ ⊢ σ · t ⇒ {ts} ⊳ Σ₁
|
|
|
|
|
tres <- inferC ctx sg t
|
2023-04-18 16:55:23 -04:00
|
|
|
|
ttags <- expectEnum !defs ctx tres.type
|
2023-03-26 18:08:09 -04:00
|
|
|
|
-- if 1 ≤ π, OR there is only zero or one option
|
2023-04-01 13:16:43 -04:00
|
|
|
|
unless (length (SortedSet.toList ttags) <= 1) $ expectCompatQ One pi
|
2023-02-22 01:45:10 -05:00
|
|
|
|
-- if Ψ | Γ, x : {ts} ⊢₀ A ⇐ Type
|
2023-04-01 13:16:43 -04:00
|
|
|
|
checkTypeC (extendTy Zero ret.name tres.type ctx) ret.term Nothing
|
2023-02-22 01:45:10 -05:00
|
|
|
|
-- if for each "a ⇒ s" in arms,
|
2023-03-31 13:11:35 -04:00
|
|
|
|
-- Ψ | Γ ⊢ σ · s ⇐ A[a ∷ {ts}/x] ⊳ Σᵢ
|
|
|
|
|
-- with Σ₂ = lubs Σᵢ
|
2023-03-13 22:22:26 -04:00
|
|
|
|
let arms = SortedMap.toList arms
|
2023-03-26 08:41:17 -04:00
|
|
|
|
let armTags = SortedSet.fromList $ map fst arms
|
2023-03-31 13:23:30 -04:00
|
|
|
|
unless (ttags == armTags) $ throw $ BadCaseEnum ttags armTags
|
2023-03-13 22:22:26 -04:00
|
|
|
|
armres <- for arms $ \(a, s) =>
|
2023-02-22 01:45:10 -05:00
|
|
|
|
checkC ctx sg s (sub1 ret (Tag a :# tres.type))
|
2023-03-26 18:01:32 -04:00
|
|
|
|
let Just armout = lubs ctx armres
|
2023-04-15 09:13:01 -04:00
|
|
|
|
| _ => throw $ BadQtys "case arms" ctx $
|
2023-03-26 18:01:32 -04:00
|
|
|
|
zipWith (\qs, (t, rhs) => (qs, Tag t)) armres arms
|
2023-02-22 01:45:10 -05:00
|
|
|
|
pure $ InfRes {
|
|
|
|
|
type = sub1 ret t,
|
|
|
|
|
qout = pi * tres.qout + armout
|
|
|
|
|
}
|
|
|
|
|
|
2023-03-26 08:40:54 -04:00
|
|
|
|
infer' ctx sg (CaseNat pi pi' n ret zer suc) = do
|
|
|
|
|
-- if 1 ≤ π
|
2023-04-01 13:16:43 -04:00
|
|
|
|
expectCompatQ One pi
|
2023-03-26 08:40:54 -04:00
|
|
|
|
-- if Ψ | Γ ⊢ σ · n ⇒ ℕ ⊳ Σn
|
|
|
|
|
nres <- inferC ctx sg n
|
2023-04-18 16:55:23 -04:00
|
|
|
|
expectNat !defs ctx nres.type
|
2023-03-26 08:40:54 -04:00
|
|
|
|
-- if Ψ | Γ, n : ℕ ⊢₀ A ⇐ Type
|
2023-04-01 13:16:43 -04:00
|
|
|
|
checkTypeC (extendTy Zero ret.name Nat ctx) ret.term Nothing
|
2023-03-26 08:40:54 -04:00
|
|
|
|
-- if Ψ | Γ ⊢ σ · zer ⇐ A[0 ∷ ℕ/n] ⊳ Σz
|
|
|
|
|
zerout <- checkC ctx sg zer (sub1 ret (Zero :# Nat))
|
|
|
|
|
-- if Ψ | Γ, n : ℕ, ih : A ⊢ σ · suc ⇐ A[succ p ∷ ℕ/n] ⊳ Σs, ρ₁.p, ρ₂.ih
|
2023-03-26 18:01:32 -04:00
|
|
|
|
-- with ρ₂ ≤ π'σ, (ρ₁ + ρ₂) ≤ πσ
|
2023-03-26 08:40:54 -04:00
|
|
|
|
let [< p, ih] = suc.names
|
|
|
|
|
pisg = pi * sg.fst
|
|
|
|
|
sucCtx = extendTyN [< (pisg, p, Nat), (pi', ih, ret.term)] ctx
|
2023-03-31 13:26:55 -04:00
|
|
|
|
sucType = substCaseSuccRet ret
|
2023-03-26 08:40:54 -04:00
|
|
|
|
sucout :< qp :< qih <- checkC sucCtx sg suc.term sucType
|
2023-03-26 18:01:32 -04:00
|
|
|
|
expectCompatQ qih (pi' * sg.fst)
|
2023-03-26 08:40:54 -04:00
|
|
|
|
-- [fixme] better error here
|
|
|
|
|
expectCompatQ (qp + qih) pisg
|
2023-04-15 09:13:01 -04:00
|
|
|
|
-- then Ψ | Γ ⊢ caseπ ⋯ ⇒ A[n] ⊳ πΣn + Σz + ωΣs
|
2023-03-26 08:40:54 -04:00
|
|
|
|
pure $ InfRes {
|
|
|
|
|
type = sub1 ret n,
|
2023-04-01 13:16:43 -04:00
|
|
|
|
qout = pi * nres.qout + zerout + Any * sucout
|
2023-03-26 08:40:54 -04:00
|
|
|
|
}
|
|
|
|
|
|
2023-03-31 13:11:35 -04:00
|
|
|
|
infer' ctx sg (CaseBox pi box ret body) = do
|
|
|
|
|
-- if Ψ | Γ ⊢ σ · b ⇒ [ρ.A] ⊳ Σ₁
|
|
|
|
|
boxres <- inferC ctx sg box
|
2023-04-18 16:55:23 -04:00
|
|
|
|
(q, ty) <- expectBOX !defs ctx boxres.type
|
2023-03-31 13:11:35 -04:00
|
|
|
|
-- if Ψ | Γ, x : [ρ.A] ⊢₀ R ⇐ Type
|
2023-04-01 13:16:43 -04:00
|
|
|
|
checkTypeC (extendTy Zero ret.name boxres.type ctx) ret.term Nothing
|
2023-03-31 13:11:35 -04:00
|
|
|
|
-- if Ψ | Γ, x : A ⊢ t ⇐ R[[x] ∷ [ρ.A/x]] ⊳ Σ₂, ς·x
|
|
|
|
|
-- with ς ≤ ρπσ
|
|
|
|
|
let qpisg = q * pi * sg.fst
|
2023-04-01 10:01:29 -04:00
|
|
|
|
bodyCtx = extendTy qpisg body.name ty ctx
|
2023-03-31 13:11:35 -04:00
|
|
|
|
bodyType = substCaseBoxRet ty ret
|
|
|
|
|
bodyout <- checkC bodyCtx sg body.term bodyType >>= popQ qpisg
|
2023-04-15 09:13:01 -04:00
|
|
|
|
-- then Ψ | Γ ⊢ caseπ ⋯ ⇒ R[b/x] ⊳ Σ₁ + Σ₂
|
2023-03-31 13:11:35 -04:00
|
|
|
|
pure $ InfRes {
|
|
|
|
|
type = sub1 ret box,
|
|
|
|
|
qout = boxres.qout + bodyout
|
|
|
|
|
}
|
|
|
|
|
|
2023-02-20 15:42:21 -05:00
|
|
|
|
infer' ctx sg (fun :% dim) = do
|
2023-03-26 18:01:32 -04:00
|
|
|
|
-- if Ψ | Γ ⊢ σ · f ⇒ Eq [𝑖 ⇒ A] l r ⊳ Σ
|
2023-02-19 11:51:44 -05:00
|
|
|
|
InfRes {type, qout} <- inferC ctx sg fun
|
2023-04-18 16:55:23 -04:00
|
|
|
|
ty <- fst <$> expectEq !defs ctx type
|
2023-03-26 18:01:32 -04:00
|
|
|
|
-- then Ψ | Γ ⊢ σ · f p ⇒ A‹p/𝑖› ⊳ Σ
|
2023-01-20 20:34:28 -05:00
|
|
|
|
pure $ InfRes {type = dsub1 ty dim, qout}
|
|
|
|
|
|
2023-04-15 09:13:01 -04:00
|
|
|
|
infer' ctx sg (Coe (S [< i] ty) p q val) = do
|
|
|
|
|
let ty = ty.term
|
|
|
|
|
checkType (extendDim i ctx) ty Nothing
|
|
|
|
|
qout <- checkC ctx sg val (ty // one p)
|
|
|
|
|
pure $ InfRes {type = ty // one q, qout}
|
|
|
|
|
|
|
|
|
|
infer' ctx sg (Comp ty p q val r (S [< j0] val0) (S [< j1] val1)) = do
|
|
|
|
|
checkType ctx ty Nothing
|
|
|
|
|
qout <- checkC ctx sg val ty
|
|
|
|
|
let ty' = dweakT 1 ty; val' = dweakT 1 val; p' = weakD 1 p
|
|
|
|
|
ctx0 = extendDim j0 $ eqDim r (K Zero) ctx
|
|
|
|
|
val0 = val0.term
|
|
|
|
|
qout0 <- check ctx0 sg val0 ty'
|
|
|
|
|
equal (eqDim (BV 0) p' ctx0) ty' val0 val'
|
|
|
|
|
let ctx1 = extendDim j0 $ eqDim r (K One) ctx
|
|
|
|
|
val1 = val1.term
|
|
|
|
|
qout1 <- check ctx1 sg val1 ty'
|
|
|
|
|
equal (eqDim (BV 0) p' ctx1) ty' val1 val'
|
|
|
|
|
let qout0' = toMaybe $ map (, val0 // one p) qout0
|
|
|
|
|
qout1' = toMaybe $ map (, val1 // one p) qout1
|
|
|
|
|
qouts = (qout, val) :: catMaybes [qout0', qout1']
|
|
|
|
|
let Just qout = lubs ctx $ map fst qouts
|
|
|
|
|
| Nothing => throw $ BadQtys "composition" ctx qouts
|
|
|
|
|
pure $ InfRes {type = ty, qout}
|
|
|
|
|
|
|
|
|
|
infer' ctx sg (TypeCase ty ret arms def) = do
|
2023-04-03 11:46:23 -04:00
|
|
|
|
-- if σ = 0
|
|
|
|
|
expectEqualQ Zero sg.fst
|
|
|
|
|
-- if Ψ, Γ ⊢₀ e ⇒ Type u
|
2023-04-18 16:55:23 -04:00
|
|
|
|
u <- expectTYPE !defs ctx . type =<< inferC ctx szero ty
|
2023-04-15 09:13:01 -04:00
|
|
|
|
-- if Ψ, Γ ⊢₀ C ⇐ Type (non-dependent return type)
|
2023-04-03 11:46:23 -04:00
|
|
|
|
checkTypeC ctx ret Nothing
|
2023-04-15 09:13:01 -04:00
|
|
|
|
-- if Ψ, Γ' ⊢₀ A ⇐ C for each rhs A
|
|
|
|
|
for_ allKinds $ \k =>
|
|
|
|
|
for_ (lookupPrecise k arms) $ \(S names t) =>
|
|
|
|
|
check0 (extendTyN (addNames0 (typecaseTel k u) names) ctx)
|
|
|
|
|
t.term (weakT (arity k) ret)
|
2023-04-03 11:46:23 -04:00
|
|
|
|
-- then Ψ, Γ ⊢₀ type-case ⋯ ⇒ C
|
|
|
|
|
pure $ InfRes {type = ret, qout = zeroFor ctx}
|
|
|
|
|
|
2023-02-20 15:42:21 -05:00
|
|
|
|
infer' ctx sg (term :# type) = do
|
2023-02-13 16:05:27 -05:00
|
|
|
|
-- if Ψ | Γ ⊢₀ A ⇐ Type ℓ
|
2023-03-05 07:14:25 -05:00
|
|
|
|
checkTypeC ctx type Nothing
|
2023-02-13 16:05:27 -05:00
|
|
|
|
-- if Ψ | Γ ⊢ σ · s ⇐ A ⊳ Σ
|
2023-02-19 11:51:44 -05:00
|
|
|
|
qout <- checkC 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}
|