quox/lib/Quox/Typing.idr

200 lines
6.0 KiB
Idris

module Quox.Typing
import public Quox.Typing.Context as Typing
import public Quox.Typing.EqMode as Typing
import public Quox.Typing.Error as Typing
import public Quox.Syntax
import public Quox.Definition
import public Quox.Reduce
import Control.Eff
%default total
public export
CheckResult' : Type -> Nat -> Type
CheckResult' = QOutput
public export
CheckResult : DimEq d -> Type -> Nat -> Type
CheckResult eqs q n = IfConsistent eqs $ CheckResult' q n
public export
record InferResult' q d n where
constructor InfRes
type : Term q d n
qout : QOutput q n
public export
InferResult : DimEq d -> TermLike
InferResult eqs q d n = IfConsistent eqs $ InferResult' q d n
export
lookupFree' : Has (ErrorEff q) fs =>
Definitions' q g -> Name -> Eff fs (Definition' q g)
lookupFree' defs x =
case lookup x defs of
Just d => pure d
Nothing => throw $ NotInScope x
public export
substCasePairRet : Term q d n -> ScopeTerm q d n -> Term q d (2 + n)
substCasePairRet dty retty =
let arg = Pair (BVT 1) (BVT 0) :# (dty // fromNat 2) in
retty.term // (arg ::: shift 2)
public export
substCaseSuccRet : ScopeTerm q d n -> Term q d (2 + n)
substCaseSuccRet retty = retty.term // (Succ (BVT 1) :# Nat ::: shift 2)
public export
substCaseBoxRet : Term q d n -> ScopeTerm q d n -> Term q d (S n)
substCaseBoxRet dty retty =
retty.term // (Box (BVT 0) :# weakT dty ::: shift 1)
parameters (defs : Definitions' q _) {auto _ : Has (ErrorEff q) fs}
export covering %inline
whnfT : {0 isRedex : RedexTest tm} -> Whnf tm isRedex WhnfError =>
{d, n : Nat} -> tm q d n -> Eff fs (NonRedex tm q d n defs)
whnfT = either (throw . WhnfError) pure . whnf defs
parameters {d2, n : Nat} (ctx : Lazy (TyContext q d1 n))
(th : Lazy (DSubst d2 d1))
export covering %inline
expectTYPE_ : Term q d2 n -> Eff fs Universe
expectTYPE_ s = case fst !(whnfT s) of
TYPE l => pure l
_ => throw $ ExpectedTYPE ctx (s // th)
export covering %inline
expectPi_ : Term q d2 n -> Eff fs (q, Term q d2 n, ScopeTerm q d2 n)
expectPi_ s = case fst !(whnfT s) of
Pi {qty, arg, res, _} => pure (qty, arg, res)
_ => throw $ ExpectedPi ctx (s // th)
export covering %inline
expectSig_ : Term q d2 n -> Eff fs (Term q d2 n, ScopeTerm q d2 n)
expectSig_ s = case fst !(whnfT s) of
Sig {fst, snd, _} => pure (fst, snd)
_ => throw $ ExpectedSig ctx (s // th)
export covering %inline
expectEnum_ : Term q d2 n -> Eff fs (SortedSet TagVal)
expectEnum_ s = case fst !(whnfT s) of
Enum tags => pure tags
_ => throw $ ExpectedEnum ctx (s // th)
export covering %inline
expectEq_ : Term q d2 n ->
Eff fs (DScopeTerm q d2 n, Term q d2 n, Term q d2 n)
expectEq_ s = case fst !(whnfT s) of
Eq {ty, l, r, _} => pure (ty, l, r)
_ => throw $ ExpectedEq ctx (s // th)
export covering %inline
expectNat_ : Term q d2 n -> Eff fs ()
expectNat_ s = case fst !(whnfT s) of
Nat => pure ()
_ => throw $ ExpectedNat ctx (s // th)
export covering %inline
expectBOX_ : Term q d2 n -> Eff fs (q, Term q d2 n)
expectBOX_ s = case fst !(whnfT s) of
BOX q a => pure (q, a)
_ => throw $ ExpectedBOX ctx (s // th)
-- [fixme] refactor this stuff
parameters (ctx : TyContext q d n)
export covering %inline
expectTYPE : Term q d n -> Eff fs Universe
expectTYPE =
let Val d = ctx.dimLen; Val n = ctx.termLen in
expectTYPE_ ctx id
export covering %inline
expectPi : Term q d n -> Eff fs (q, Term q d n, ScopeTerm q d n)
expectPi =
let Val d = ctx.dimLen; Val n = ctx.termLen in
expectPi_ ctx id
export covering %inline
expectSig : Term q d n -> Eff fs (Term q d n, ScopeTerm q d n)
expectSig =
let Val d = ctx.dimLen; Val n = ctx.termLen in
expectSig_ ctx id
export covering %inline
expectEnum : Term q d n -> Eff fs (SortedSet TagVal)
expectEnum =
let Val d = ctx.dimLen; Val n = ctx.termLen in
expectEnum_ ctx id
export covering %inline
expectEq : Term q d n -> Eff fs (DScopeTerm q d n, Term q d n, Term q d n)
expectEq =
let Val d = ctx.dimLen; Val n = ctx.termLen in
expectEq_ ctx id
export covering %inline
expectNat : Term q d n -> Eff fs ()
expectNat =
let Val d = ctx.dimLen; Val n = ctx.termLen in
expectNat_ ctx id
export covering %inline
expectBOX : Term q d n -> Eff fs (q, Term q d n)
expectBOX =
let Val d = ctx.dimLen; Val n = ctx.termLen in
expectBOX_ ctx id
parameters (ctx : EqContext q n)
export covering %inline
expectTYPEE : Term q 0 n -> Eff fs Universe
expectTYPEE t =
let Val n = ctx.termLen in
expectTYPE_ (toTyContext ctx) (shift0 ctx.dimLen) t
export covering %inline
expectPiE : Term q 0 n -> Eff fs (q, Term q 0 n, ScopeTerm q 0 n)
expectPiE t =
let Val n = ctx.termLen in
expectPi_ (toTyContext ctx) (shift0 ctx.dimLen) t
export covering %inline
expectSigE : Term q 0 n -> Eff fs (Term q 0 n, ScopeTerm q 0 n)
expectSigE t =
let Val n = ctx.termLen in
expectSig_ (toTyContext ctx) (shift0 ctx.dimLen) t
export covering %inline
expectEnumE : Term q 0 n -> Eff fs (SortedSet TagVal)
expectEnumE t =
let Val n = ctx.termLen in
expectEnum_ (toTyContext ctx) (shift0 ctx.dimLen) t
export covering %inline
expectEqE : Term q 0 n -> Eff fs (DScopeTerm q 0 n, Term q 0 n, Term q 0 n)
expectEqE t =
let Val n = ctx.termLen in
expectEq_ (toTyContext ctx) (shift0 ctx.dimLen) t
export covering %inline
expectNatE : Term q 0 n -> Eff fs ()
expectNatE t =
let Val n = ctx.termLen in
expectNat_ (toTyContext ctx) (shift0 ctx.dimLen) t
export covering %inline
expectBOXE : Term q 0 n -> Eff fs (q, Term q 0 n)
expectBOXE t =
let Val n = ctx.termLen in
expectBOX_ (toTyContext ctx) (shift0 ctx.dimLen) t