quox/lib/Quox/Typing.idr

150 lines
4.6 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 Language.Reflection
import Control.Eff
%default total
%language ElabReflection
private TTName : Type
TTName = TT.Name
%hide TT.Name
public export
CheckResult' : Nat -> Type
CheckResult' = QOutput
public export
CheckResult : DimEq d -> Nat -> Type
CheckResult eqs n = IfConsistent eqs $ CheckResult' n
public export
record InferResult' d n where
constructor InfRes
type : Term d n
qout : QOutput n
public export
InferResult : DimEq d -> TermLike
InferResult eqs d n = IfConsistent eqs $ InferResult' d n
export
lookupFree : Has ErrorEff fs => Name -> Definitions -> Eff fs Definition
lookupFree x defs = maybe (throw $ NotInScope x) pure $ lookup x defs
public export
substCasePairRet : Term d n -> ScopeTerm d n -> Term 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 d n -> Term d (2 + n)
substCaseSuccRet retty = retty.term // (Succ (BVT 1) :# Nat ::: shift 2)
public export
substCaseBoxRet : Term d n -> ScopeTerm d n -> Term d (S n)
substCaseBoxRet dty retty =
retty.term // (Box (BVT 0) :# weakT 1 dty ::: shift 1)
parameters (defs : Definitions) {auto _ : Has ErrorEff fs}
namespace TyContext
parameters (ctx : TyContext d n)
export covering
whnf : {0 isRedex : RedexTest tm} -> Whnf tm isRedex =>
tm d n -> Eff fs (NonRedex tm d n defs)
whnf = let Val n = ctx.termLen; Val d = ctx.dimLen in
rethrow . whnf defs (toWhnfContext ctx)
private covering %macro
expect : (forall d, n. NameContexts d n -> Term d n -> Error) ->
TTImp -> TTImp -> Elab (Term d n -> Eff fs a)
expect k l r = do
f <- check `(\case ~(l) => Just ~(r); _ => Nothing)
pure $ \t => maybe (throw $ k ctx.names t) pure . f . fst =<< whnf t
export covering %inline
expectTYPE : Term d n -> Eff fs Universe
expectTYPE = expect ExpectedTYPE `(TYPE l) `(l)
export covering %inline
expectPi : Term d n -> Eff fs (Qty, Term d n, ScopeTerm d n)
expectPi = expect ExpectedPi `(Pi {qty, arg, res}) `((qty, arg, res))
export covering %inline
expectSig : Term d n -> Eff fs (Term d n, ScopeTerm d n)
expectSig = expect ExpectedSig `(Sig {fst, snd}) `((fst, snd))
export covering %inline
expectEnum : Term d n -> Eff fs (SortedSet TagVal)
expectEnum = expect ExpectedEnum `(Enum ts) `(ts)
export covering %inline
expectEq : Term d n -> Eff fs (DScopeTerm d n, Term d n, Term d n)
expectEq = expect ExpectedEq `(Eq {ty, l, r}) `((ty, l, r))
export covering %inline
expectNat : Term d n -> Eff fs ()
expectNat = expect ExpectedNat `(Nat) `(())
export covering %inline
expectBOX : Term d n -> Eff fs (Qty, Term d n)
expectBOX = expect ExpectedBOX `(BOX {qty, ty}) `((qty, ty))
namespace EqContext
parameters (ctx : EqContext n)
export covering
whnf : {0 isRedex : RedexTest tm} -> Whnf tm isRedex =>
tm 0 n -> Eff fs (NonRedex tm 0 n defs)
whnf = let Val n = ctx.termLen in rethrow . whnf defs (toWhnfContext ctx)
private covering %macro
expect : (forall d, n. NameContexts d n -> Term d n -> Error) ->
TTImp -> TTImp -> Elab (Term 0 n -> Eff fs a)
expect k l r = do
f <- check `(\case ~(l) => Just ~(r); _ => Nothing)
pure $ \t =>
let err = throw $ k ctx.names (t // shift0 ctx.dimLen) in
maybe err pure . f . fst =<< whnf t
export covering %inline
expectTYPE : Term 0 n -> Eff fs Universe
expectTYPE = expect ExpectedTYPE `(TYPE l) `(l)
export covering %inline
expectPi : Term 0 n -> Eff fs (Qty, Term 0 n, ScopeTerm 0 n)
expectPi = expect ExpectedPi `(Pi {qty, arg, res}) `((qty, arg, res))
export covering %inline
expectSig : Term 0 n -> Eff fs (Term 0 n, ScopeTerm 0 n)
expectSig = expect ExpectedSig `(Sig {fst, snd}) `((fst, snd))
export covering %inline
expectEnum : Term 0 n -> Eff fs (SortedSet TagVal)
expectEnum = expect ExpectedEnum `(Enum ts) `(ts)
export covering %inline
expectEq : Term 0 n -> Eff fs (DScopeTerm 0 n, Term 0 n, Term 0 n)
expectEq = expect ExpectedEq `(Eq {ty, l, r}) `((ty, l, r))
export covering %inline
expectNat : Term 0 n -> Eff fs ()
expectNat = expect ExpectedNat `(Nat) `(())
export covering %inline
expectBOX : Term 0 n -> Eff fs (Qty, Term 0 n)
expectBOX = expect ExpectedBOX `(BOX {qty, ty}) `((qty, ty))