put names into contexts, and contexts into errors
This commit is contained in:
parent
f4af1a5a78
commit
86d21caf24
13 changed files with 520 additions and 324 deletions
|
@ -174,6 +174,12 @@ zipWithLazy : forall tm1, tm2, tm3.
|
|||
Telescope (\n => Lazy (tm3 n)) from to
|
||||
zipWithLazy f = zipWith $ delay .: f
|
||||
|
||||
export
|
||||
unzip : Telescope (\n => (tm1 n, tm2 n)) from to ->
|
||||
(Telescope tm1 from to, Telescope tm2 from to)
|
||||
unzip [<] = ([<], [<])
|
||||
unzip (tel :< (x, y)) = let (xs, ys) = unzip tel in (xs :< x, ys :< y)
|
||||
|
||||
|
||||
export
|
||||
lengthPrf : Telescope _ from to -> Subset Nat (\len => len + from = to)
|
||||
|
|
|
@ -27,17 +27,18 @@ private %inline
|
|||
mode : HasCmpContext m => m EqMode
|
||||
mode = asks mode
|
||||
|
||||
private %inline
|
||||
clashT : CanEqual q m => Term q d n -> Term q d n -> Term q d n -> m a
|
||||
clashT ty s t = throwError $ ClashT !mode ty s t
|
||||
parameters {auto _ : CanEqual q m} (ctx : EqContext q n)
|
||||
private %inline
|
||||
clashT : Term q 0 n -> Term q 0 n -> Term q 0 n -> m a
|
||||
clashT ty s t = throwError $ ClashT ctx !mode ty s t
|
||||
|
||||
private %inline
|
||||
clashTy : CanEqual q m => Term q d n -> Term q d n -> m a
|
||||
clashTy s t = throwError $ ClashTy !mode s t
|
||||
private %inline
|
||||
clashTy : Term q 0 n -> Term q 0 n -> m a
|
||||
clashTy s t = throwError $ ClashTy ctx !mode s t
|
||||
|
||||
private %inline
|
||||
clashE : CanEqual q m => Elim q d n -> Elim q d n -> m a
|
||||
clashE e f = throwError $ ClashE !mode e f
|
||||
private %inline
|
||||
clashE : Elim q 0 n -> Elim q 0 n -> m a
|
||||
clashE e f = throwError $ ClashE ctx !mode e f
|
||||
|
||||
|
||||
||| true if a term is syntactically a type, or is neutral.
|
||||
|
@ -114,8 +115,12 @@ parameters {auto _ : HasErr q m}
|
|||
Right _ => throwError $ e t
|
||||
|
||||
export %inline
|
||||
ensureTyCon : HasErr q m => (t : Term q d n) -> m (So (isTyCon t))
|
||||
ensureTyCon = ensure NotType isTyCon
|
||||
ensureTyCon : (ctx : EqContext q n) -> (t : Term q 0 n) -> m (So (isTyCon t))
|
||||
ensureTyCon ctx t = case nchoose $ isTyCon t of
|
||||
Left y => pure y
|
||||
Right n =>
|
||||
let (d ** ctx) = toTyContext ctx in
|
||||
throwError $ NotType ctx (t // shift0 d)
|
||||
|
||||
parameters (defs : Definitions' q _) {auto _ : (CanEqual q m, Eq q)}
|
||||
mutual
|
||||
|
@ -125,13 +130,13 @@ parameters (defs : Definitions' q _) {auto _ : (CanEqual q m, Eq q)}
|
|||
|||
|
||||
||| ⚠ **assumes that `s`, `t` have already been checked against `ty`**. ⚠
|
||||
export covering %inline
|
||||
compare0 : TContext q 0 n -> (ty, s, t : Term q 0 n) -> m ()
|
||||
compare0 : EqContext q n -> (ty, s, t : Term q 0 n) -> m ()
|
||||
compare0 ctx ty s t =
|
||||
wrapErr (WhileComparingT (MkTyContext new ctx) !mode ty s t) $ do
|
||||
wrapErr (WhileComparingT ctx !mode ty s t) $ do
|
||||
Element ty nty <- whnfT defs ty
|
||||
Element s ns <- whnfT defs s
|
||||
Element t nt <- whnfT defs t
|
||||
tty <- ensureTyCon ty
|
||||
tty <- ensureTyCon ctx ty
|
||||
compare0' ctx ty s t
|
||||
|
||||
||| converts an elim "Γ ⊢ e" to "Γ, x ⊢ e x", for comparing with
|
||||
|
@ -141,7 +146,7 @@ parameters (defs : Definitions' q _) {auto _ : (CanEqual q m, Eq q)}
|
|||
toLamBody e = E $ weakE e :@ BVT 0
|
||||
|
||||
private covering
|
||||
compare0' : TContext q 0 n ->
|
||||
compare0' : EqContext q n ->
|
||||
(ty, s, t : Term q 0 n) ->
|
||||
(0 nty : NotRedex defs ty) => (0 tty : So (isTyCon ty)) =>
|
||||
(0 ns : NotRedex defs s) => (0 nt : NotRedex defs t) =>
|
||||
|
@ -162,14 +167,14 @@ parameters (defs : Definitions' q _) {auto _ : (CanEqual q m, Eq q)}
|
|||
(Lam b, E e) => eta e b
|
||||
|
||||
(E e, E f) => Elim.compare0 ctx e f
|
||||
_ => throwError $ WrongType ty s t
|
||||
_ => throwError $ WrongType ctx ty s t
|
||||
where
|
||||
ctx' : TContext q 0 (S n)
|
||||
ctx' = ctx :< arg
|
||||
ctx' : EqContext q (S n)
|
||||
ctx' = extendTy res.name arg ctx
|
||||
|
||||
eta : Elim q 0 n -> ScopeTerm q 0 n -> m ()
|
||||
eta e (S _ (Y b)) = compare0 ctx' res.term (toLamBody e) b
|
||||
eta e (S _ (N _)) = clashT ty s t
|
||||
eta e (S _ (N _)) = clashT ctx ty s t
|
||||
|
||||
compare0' ctx ty@(Sig {fst, snd, _}) s t = local {mode := Equal} $
|
||||
case (s, t) of
|
||||
|
@ -183,7 +188,7 @@ parameters (defs : Definitions' q _) {auto _ : (CanEqual q m, Eq q)}
|
|||
compare0 ctx (sub1 snd (sFst :# fst)) sSnd tSnd
|
||||
|
||||
(E e, E f) => compare0 ctx e f
|
||||
_ => throwError $ WrongType ty s t
|
||||
_ => throwError $ WrongType ctx ty s t
|
||||
|
||||
compare0' ctx ty@(Enum tags) s t = local {mode := Equal} $
|
||||
case (s, t) of
|
||||
|
@ -191,9 +196,9 @@ parameters (defs : Definitions' q _) {auto _ : (CanEqual q m, Eq q)}
|
|||
-- Γ ⊢ `t = `t : {ts}
|
||||
--
|
||||
-- t ∈ ts is in the typechecker, not here, ofc
|
||||
(Tag t1, Tag t2) => unless (t1 == t2) $ clashT ty s t
|
||||
(Tag t1, Tag t2) => unless (t1 == t2) $ clashT ctx ty s t
|
||||
(E e, E f) => compare0 ctx e f
|
||||
_ => throwError $ WrongType ty s t
|
||||
_ => throwError $ WrongType ctx ty s t
|
||||
|
||||
compare0' _ (Eq {}) _ _ =
|
||||
-- ✨ uip ✨
|
||||
|
@ -204,25 +209,25 @@ parameters (defs : Definitions' q _) {auto _ : (CanEqual q m, Eq q)}
|
|||
compare0' ctx ty@(E _) s t = do
|
||||
-- a neutral type can only be inhabited by neutral values
|
||||
-- e.g. an abstract value in an abstract type, bound variables, …
|
||||
E e <- pure s | _ => throwError $ WrongType ty s t
|
||||
E f <- pure t | _ => throwError $ WrongType ty s t
|
||||
E e <- pure s | _ => throwError $ WrongType ctx ty s t
|
||||
E f <- pure t | _ => throwError $ WrongType ctx ty s t
|
||||
Elim.compare0 ctx e f
|
||||
|
||||
||| compares two types, using the current variance `mode` for universes.
|
||||
||| fails if they are not types, even if they would happen to be equal.
|
||||
export covering %inline
|
||||
compareType : TContext q 0 n -> (s, t : Term q 0 n) -> m ()
|
||||
compareType : EqContext q n -> (s, t : Term q 0 n) -> m ()
|
||||
compareType ctx s t = do
|
||||
Element s ns <- whnfT defs s
|
||||
Element t nt <- whnfT defs t
|
||||
ts <- ensureTyCon s
|
||||
tt <- ensureTyCon t
|
||||
st <- either pure (const $ clashTy s t) $
|
||||
ts <- ensureTyCon ctx s
|
||||
tt <- ensureTyCon ctx t
|
||||
st <- either pure (const $ clashTy ctx s t) $
|
||||
nchoose $ sameTyCon s t
|
||||
compareType' ctx s t
|
||||
|
||||
private covering
|
||||
compareType' : TContext q 0 n -> (s, t : Term q 0 n) ->
|
||||
compareType' : EqContext q n -> (s, t : Term q 0 n) ->
|
||||
(0 ns : NotRedex defs s) => (0 ts : So (isTyCon s)) =>
|
||||
(0 nt : NotRedex defs t) => (0 tt : So (isTyCon t)) =>
|
||||
(0 st : So (sameTyCon s t)) =>
|
||||
|
@ -242,7 +247,7 @@ parameters (defs : Definitions' q _) {auto _ : (CanEqual q m, Eq q)}
|
|||
-- Γ ⊢ (π·x : A₁) → B₁ <: (π·x : A₂) → B₂
|
||||
expectEqualQ sQty tQty
|
||||
local {mode $= flip} $ compareType ctx sArg tArg -- contra
|
||||
compareType (ctx :< sArg) sRes.term tRes.term
|
||||
compareType (extendTy sRes.name sArg ctx) sRes.term tRes.term
|
||||
|
||||
compareType' ctx (Sig {fst = sFst, snd = sSnd, _})
|
||||
(Sig {fst = tFst, snd = tSnd, _}) = do
|
||||
|
@ -250,7 +255,7 @@ parameters (defs : Definitions' q _) {auto _ : (CanEqual q m, Eq q)}
|
|||
-- --------------------------------------
|
||||
-- Γ ⊢ (x : A₁) × B₁ <: (x : A₂) × B₂
|
||||
compareType ctx sFst tFst
|
||||
compareType (ctx :< sFst) sSnd.term tSnd.term
|
||||
compareType (extendTy sSnd.name sFst ctx) sSnd.term tSnd.term
|
||||
|
||||
compareType' ctx (Eq {ty = sTy, l = sl, r = sr, _})
|
||||
(Eq {ty = tTy, l = tl, r = tr, _}) = do
|
||||
|
@ -258,8 +263,8 @@ parameters (defs : Definitions' q _) {auto _ : (CanEqual q m, Eq q)}
|
|||
-- Γ ⊢ l₁ = l₂ : A₁‹𝟎/i› Γ ⊢ r₁ = r₂ : A₁‹𝟏/i›
|
||||
-- ------------------------------------------------
|
||||
-- Γ ⊢ Eq [i ⇒ A₁] l₁ r₂ <: Eq [i ⇒ A₂] l₂ r₂
|
||||
compareType ctx sTy.zero tTy.zero
|
||||
compareType ctx sTy.one tTy.one
|
||||
compareType (extendDim sTy.name Zero ctx) sTy.zero tTy.zero
|
||||
compareType (extendDim sTy.name One ctx) sTy.one tTy.one
|
||||
local {mode := Equal} $ do
|
||||
Term.compare0 ctx sTy.zero sl tl
|
||||
Term.compare0 ctx sTy.one sr tr
|
||||
|
@ -270,7 +275,7 @@ parameters (defs : Definitions' q _) {auto _ : (CanEqual q m, Eq q)}
|
|||
--
|
||||
-- no subtyping based on tag subsets, since that would need
|
||||
-- a runtime coercion
|
||||
unless (tags1 == tags2) $ clashTy s t
|
||||
unless (tags1 == tags2) $ clashTy ctx s t
|
||||
|
||||
compareType' ctx (E e) (E f) = do
|
||||
-- no fanciness needed here cos anything other than a neutral
|
||||
|
@ -281,33 +286,33 @@ parameters (defs : Definitions' q _) {auto _ : (CanEqual q m, Eq q)}
|
|||
|||
|
||||
||| ⚠ **assumes the elim is already typechecked.** ⚠
|
||||
private covering
|
||||
computeElimType : TContext q 0 n -> (e : Elim q 0 n) ->
|
||||
computeElimType : EqContext q n -> (e : Elim q 0 n) ->
|
||||
(0 ne : NotRedex defs e) ->
|
||||
m (Term q 0 n)
|
||||
computeElimType ctx (F x) _ = do
|
||||
defs <- lookupFree' defs x
|
||||
pure $ defs.type.get
|
||||
computeElimType ctx (B i) _ = do
|
||||
pure $ ctx !! i
|
||||
pure $ ctx.tctx !! i
|
||||
computeElimType ctx (f :@ s) ne = do
|
||||
(_, arg, res) <- expectPi defs !(computeElimType ctx f (noOr1 ne))
|
||||
(_, arg, res) <- expectPiE defs ctx !(computeElimType ctx f (noOr1 ne))
|
||||
pure $ sub1 res (s :# arg)
|
||||
computeElimType ctx (CasePair {pair, ret, _}) _ = do
|
||||
pure $ sub1 ret pair
|
||||
computeElimType ctx (CaseEnum {tag, ret, _}) _ = do
|
||||
pure $ sub1 ret tag
|
||||
computeElimType ctx (f :% p) ne = do
|
||||
(ty, _, _) <- expectEq defs !(computeElimType ctx f (noOr1 ne))
|
||||
(ty, _, _) <- expectEqE defs ctx !(computeElimType ctx f (noOr1 ne))
|
||||
pure $ dsub1 ty p
|
||||
computeElimType ctx (_ :# ty) _ = do
|
||||
pure ty
|
||||
|
||||
private covering
|
||||
replaceEnd : TContext q 0 n ->
|
||||
replaceEnd : EqContext q n ->
|
||||
(e : Elim q 0 n) -> DimConst -> (0 ne : NotRedex defs e) ->
|
||||
m (Elim q 0 n)
|
||||
replaceEnd ctx e p ne = do
|
||||
(ty, l, r) <- expectEq defs !(computeElimType ctx e ne)
|
||||
(ty, l, r) <- expectEqE defs ctx !(computeElimType ctx e ne)
|
||||
pure $ ends l r p :# dsub1 ty (K p)
|
||||
|
||||
namespace Elim
|
||||
|
@ -319,9 +324,9 @@ parameters (defs : Definitions' q _) {auto _ : (CanEqual q m, Eq q)}
|
|||
||| ⚠ **assumes that they have both been typechecked, and have
|
||||
||| equal types.** ⚠
|
||||
export covering %inline
|
||||
compare0 : TContext q 0 n -> (e, f : Elim q 0 n) -> m ()
|
||||
compare0 : EqContext q n -> (e, f : Elim q 0 n) -> m ()
|
||||
compare0 ctx e f =
|
||||
wrapErr (WhileComparingE (MkTyContext new ctx) !mode e f) $ do
|
||||
wrapErr (WhileComparingE ctx !mode e f) $ do
|
||||
Element e ne <- whnfT defs e
|
||||
Element f nf <- whnfT defs f
|
||||
-- [fixme] there is a better way to do this "isSubSing" stuff for sure
|
||||
|
@ -329,7 +334,7 @@ parameters (defs : Definitions' q _) {auto _ : (CanEqual q m, Eq q)}
|
|||
compare0' ctx e f ne nf
|
||||
|
||||
private covering
|
||||
compare0' : TContext q 0 n ->
|
||||
compare0' : EqContext q n ->
|
||||
(e, f : Elim q 0 n) ->
|
||||
(0 ne : NotRedex defs e) -> (0 nf : NotRedex defs f) ->
|
||||
m ()
|
||||
|
@ -342,38 +347,40 @@ parameters (defs : Definitions' q _) {auto _ : (CanEqual q m, Eq q)}
|
|||
compare0' ctx e (f :% K q) ne nf =
|
||||
compare0 ctx e !(replaceEnd ctx f q $ noOr1 nf)
|
||||
|
||||
compare0' _ e@(F x) f@(F y) _ _ = unless (x == y) $ clashE e f
|
||||
compare0' _ e@(F _) f _ _ = clashE e f
|
||||
compare0' ctx e@(F x) f@(F y) _ _ = unless (x == y) $ clashE ctx e f
|
||||
compare0' ctx e@(F _) f _ _ = clashE ctx e f
|
||||
|
||||
compare0' ctx e@(B i) f@(B j) _ _ = unless (i == j) $ clashE e f
|
||||
compare0' _ e@(B _) f _ _ = clashE e f
|
||||
compare0' ctx e@(B i) f@(B j) _ _ = unless (i == j) $ clashE ctx e f
|
||||
compare0' ctx e@(B _) f _ _ = clashE ctx e f
|
||||
|
||||
compare0' ctx (e :@ s) (f :@ t) ne nf =
|
||||
local {mode := Equal} $ do
|
||||
compare0 ctx e f
|
||||
(_, arg, _) <- expectPi defs !(computeElimType ctx e (noOr1 ne))
|
||||
(_, arg, _) <- expectPiE defs ctx !(computeElimType ctx e (noOr1 ne))
|
||||
Term.compare0 ctx arg s t
|
||||
compare0' _ e@(_ :@ _) f _ _ = clashE e f
|
||||
compare0' ctx e@(_ :@ _) f _ _ = clashE ctx e f
|
||||
|
||||
compare0' ctx (CasePair epi e eret ebody)
|
||||
(CasePair fpi f fret fbody) ne nf =
|
||||
local {mode := Equal} $ do
|
||||
compare0 ctx e f
|
||||
ety <- computeElimType ctx e (noOr1 ne)
|
||||
compareType (ctx :< ety) eret.term fret.term
|
||||
(fst, snd) <- expectSig defs ety
|
||||
Term.compare0 (ctx :< fst :< snd.term) (substCasePairRet ety eret)
|
||||
ebody.term fbody.term
|
||||
compareType (extendTy eret.name ety ctx) eret.term fret.term
|
||||
(fst, snd) <- expectSigE defs ctx ety
|
||||
let [x, y] = ebody.names
|
||||
Term.compare0 (extendTyN [< (x, fst), (y, snd.term)] ctx)
|
||||
(substCasePairRet ety eret)
|
||||
ebody.term fbody.term
|
||||
expectEqualQ epi fpi
|
||||
compare0' _ e@(CasePair {}) f _ _ = clashE e f
|
||||
compare0' ctx e@(CasePair {}) f _ _ = clashE ctx e f
|
||||
|
||||
compare0' ctx (CaseEnum epi e eret earms)
|
||||
(CaseEnum fpi f fret farms) ne nf =
|
||||
local {mode := Equal} $ do
|
||||
compare0 ctx e f
|
||||
ety <- computeElimType ctx e (noOr1 ne)
|
||||
compareType (ctx :< ety) eret.term fret.term
|
||||
for_ !(expectEnum defs ety) $ \t =>
|
||||
compareType (extendTy eret.name ety ctx) eret.term fret.term
|
||||
for_ !(expectEnumE defs ctx ety) $ \t =>
|
||||
compare0 ctx (sub1 eret $ Tag t :# ety)
|
||||
!(lookupArm t earms) !(lookupArm t farms)
|
||||
expectEqualQ epi fpi
|
||||
|
@ -382,7 +389,7 @@ parameters (defs : Definitions' q _) {auto _ : (CanEqual q m, Eq q)}
|
|||
lookupArm t arms = case lookup t arms of
|
||||
Just arm => pure arm
|
||||
Nothing => throwError $ TagNotIn t (fromList $ keys arms)
|
||||
compare0' _ e@(CaseEnum {}) f _ _ = clashE e f
|
||||
compare0' ctx e@(CaseEnum {}) f _ _ = clashE ctx e f
|
||||
|
||||
compare0' ctx (s :# a) (t :# b) _ _ =
|
||||
Term.compare0 ctx !(bigger a b) s t
|
||||
|
@ -392,10 +399,13 @@ parameters (defs : Definitions' q _) {auto _ : (CanEqual q m, Eq q)}
|
|||
|
||||
compare0' ctx (s :# a) f _ _ = Term.compare0 ctx a s (E f)
|
||||
compare0' ctx e (t :# b) _ _ = Term.compare0 ctx b (E e) t
|
||||
compare0' _ e@(_ :# _) f _ _ = clashE e f
|
||||
compare0' ctx e@(_ :# _) f _ _ = clashE ctx e f
|
||||
|
||||
|
||||
parameters {auto _ : (HasDefs' q _ m, HasErr q m, Eq q)} (ctx : TyContext q d n)
|
||||
-- [todo] only split on the dvars that are actually used anywhere in
|
||||
-- the calls to `splits`
|
||||
|
||||
parameters (mode : EqMode)
|
||||
namespace Term
|
||||
export covering
|
||||
|
@ -404,8 +414,8 @@ parameters {auto _ : (HasDefs' q _ m, HasErr q m, Eq q)} (ctx : TyContext q d n)
|
|||
defs <- ask
|
||||
runReaderT {m} (MkCmpContext {mode}) $
|
||||
for_ (splits ctx.dctx) $ \th =>
|
||||
compare0 defs (map (// th) ctx.tctx)
|
||||
(ty // th) (s // th) (t // th)
|
||||
let ectx = makeEqContext ctx th in
|
||||
compare0 defs ectx (ty // th) (s // th) (t // th)
|
||||
|
||||
export covering
|
||||
compareType : (s, t : Term q d n) -> m ()
|
||||
|
@ -413,7 +423,8 @@ parameters {auto _ : (HasDefs' q _ m, HasErr q m, Eq q)} (ctx : TyContext q d n)
|
|||
defs <- ask
|
||||
runReaderT {m} (MkCmpContext {mode}) $
|
||||
for_ (splits ctx.dctx) $ \th =>
|
||||
compareType defs (map (// th) ctx.tctx) (s // th) (t // th)
|
||||
let ectx = makeEqContext ctx th in
|
||||
compareType defs ectx (s // th) (t // th)
|
||||
|
||||
namespace Elim
|
||||
||| you don't have to pass the type in but the arguments must still be
|
||||
|
@ -424,7 +435,8 @@ parameters {auto _ : (HasDefs' q _ m, HasErr q m, Eq q)} (ctx : TyContext q d n)
|
|||
defs <- ask
|
||||
runReaderT {m} (MkCmpContext {mode}) $
|
||||
for_ (splits ctx.dctx) $ \th =>
|
||||
compare0 defs (map (// th) ctx.tctx) (e // th) (f // th)
|
||||
let ectx = makeEqContext ctx th in
|
||||
compare0 defs ectx (e // th) (f // th)
|
||||
|
||||
namespace Term
|
||||
export covering %inline
|
||||
|
|
|
@ -60,6 +60,16 @@ ifConsistent ZeroIsOne act = pure Nothing
|
|||
ifConsistent (C _) act = Just <$> act
|
||||
|
||||
|
||||
export
|
||||
fromGround' : Context' DimConst d -> DimEq' d
|
||||
fromGround' [<] = [<]
|
||||
fromGround' (ctx :< e) = fromGround' ctx :< Just (K e)
|
||||
|
||||
export
|
||||
fromGround : Context' DimConst d -> DimEq d
|
||||
fromGround = C . fromGround'
|
||||
|
||||
|
||||
export %inline
|
||||
zeroEq : DimEq 0
|
||||
zeroEq = C [<]
|
||||
|
|
|
@ -58,6 +58,10 @@ public export %inline
|
|||
shift : (by : Nat) -> Subst env from (by + from)
|
||||
shift by = Shift $ fromNat by
|
||||
|
||||
public export %inline
|
||||
shift0 : (by : Nat) -> Subst env 0 by
|
||||
shift0 by = rewrite sym $ plusZeroRightNeutral by in Shift $ fromNat by
|
||||
|
||||
|
||||
public export
|
||||
(.) : CanSubstSelf f => Subst f from mid -> Subst f mid to -> Subst f from to
|
||||
|
|
|
@ -165,6 +165,14 @@ public export %inline
|
|||
SY : Vect s BaseName -> f (s + n) -> Scoped s f n
|
||||
SY ns = S ns . Y
|
||||
|
||||
public export %inline
|
||||
name : Scoped 1 f n -> BaseName
|
||||
name (S [x] _) = x
|
||||
|
||||
public export %inline
|
||||
(.name) : Scoped 1 f n -> BaseName
|
||||
s.name = name s
|
||||
|
||||
||| more convenient Pi
|
||||
public export %inline
|
||||
Pi_ : (qty : q) -> (x : BaseName) ->
|
||||
|
|
|
@ -3,6 +3,8 @@ module Quox.Typechecker
|
|||
import public Quox.Typing
|
||||
import public Quox.Equal
|
||||
|
||||
import Data.List
|
||||
|
||||
%default total
|
||||
|
||||
|
||||
|
@ -107,7 +109,7 @@ parameters {auto _ : IsQty q} {auto _ : CanTC q m}
|
|||
(subj : Term q d n) -> (0 nc : NotClo subj) => Term q d n ->
|
||||
m (CheckResult' q n)
|
||||
toCheckType ctx sg t ty = do
|
||||
u <- expectTYPE !ask ty
|
||||
u <- expectTYPE !ask ctx ty
|
||||
expectEqualQ zero sg.fst
|
||||
checkTypeNoWrap ctx t (Just u)
|
||||
pure $ zeroFor ctx
|
||||
|
@ -122,18 +124,18 @@ parameters {auto _ : IsQty q} {auto _ : CanTC q m}
|
|||
check' ctx sg t@(Pi {}) ty = toCheckType ctx sg t ty
|
||||
|
||||
check' ctx sg (Lam body) ty = do
|
||||
(qty, arg, res) <- expectPi !ask ty
|
||||
(qty, arg, res) <- expectPi !ask ctx ty
|
||||
-- if Ψ | Γ, x : A ⊢ σ · t ⇐ B ⊳ Σ, ρ·x
|
||||
-- with ρ ≤ σπ
|
||||
let qty' = sg.fst * qty
|
||||
qout <- checkC (extendTy arg ctx) sg body.term res.term
|
||||
qout <- checkC (extendTy body.name arg ctx) sg body.term res.term
|
||||
-- then Ψ | Γ ⊢ σ · (λx ⇒ t) ⇐ (π·x : A) → B ⊳ Σ
|
||||
popQ qty' qout
|
||||
|
||||
check' ctx sg t@(Sig {}) ty = toCheckType ctx sg t ty
|
||||
|
||||
check' ctx sg (Pair fst snd) ty = do
|
||||
(tfst, tsnd) <- expectSig !ask ty
|
||||
(tfst, tsnd) <- expectSig !ask ctx ty
|
||||
-- if Ψ | Γ ⊢ σ · s ⇐ A ⊳ Σ₁
|
||||
qfst <- checkC ctx sg fst tfst
|
||||
let tsnd = sub1 tsnd (fst :# tfst)
|
||||
|
@ -145,7 +147,7 @@ parameters {auto _ : IsQty q} {auto _ : CanTC q m}
|
|||
check' ctx sg t@(Enum _) ty = toCheckType ctx sg t ty
|
||||
|
||||
check' ctx sg (Tag t) ty = do
|
||||
tags <- expectEnum !ask ty
|
||||
tags <- expectEnum !ask ctx ty
|
||||
-- if t ∈ ts
|
||||
unless (t `elem` tags) $ throwError $ TagNotIn t tags
|
||||
-- then Ψ | Γ ⊢ σ · t ⇐ {ts} ⊳ 𝟎
|
||||
|
@ -154,9 +156,9 @@ parameters {auto _ : IsQty q} {auto _ : CanTC q m}
|
|||
check' ctx sg t@(Eq {}) ty = toCheckType ctx sg t ty
|
||||
|
||||
check' ctx sg (DLam body) ty = do
|
||||
(ty, l, r) <- expectEq !ask ty
|
||||
(ty, l, r) <- expectEq !ask ctx ty
|
||||
-- if Ψ, i | Γ ⊢ σ · t ⇐ A ⊳ Σ
|
||||
qout <- checkC (extendDim ctx) sg body.term ty.term
|
||||
qout <- checkC (extendDim body.name ctx) sg body.term ty.term
|
||||
-- if Ψ | Γ ⊢ t‹0› = l : A‹0›
|
||||
equal ctx ty.zero body.zero l
|
||||
-- if Ψ | Γ ⊢ t‹1› = r : A‹1›
|
||||
|
@ -188,36 +190,36 @@ parameters {auto _ : IsQty q} {auto _ : CanTC q m}
|
|||
checkTypeC ctx arg u
|
||||
-- if Ψ | Γ, x : A ⊢₀ B ⇐ Type ℓ
|
||||
case res.body of
|
||||
Y res => checkTypeC (extendTy arg ctx) res u
|
||||
N res => checkTypeC ctx res u
|
||||
Y res' => checkTypeC (extendTy res.name arg ctx) res' u
|
||||
N res' => checkTypeC ctx res' u
|
||||
-- then Ψ | Γ ⊢₀ (π·x : A) → B ⇐ Type ℓ
|
||||
|
||||
checkType' ctx t@(Lam {}) u =
|
||||
throwError $ NotType t
|
||||
throwError $ NotType ctx t
|
||||
|
||||
checkType' ctx (Sig fst snd) u = do
|
||||
-- if Ψ | Γ ⊢₀ A ⇐ Type ℓ
|
||||
checkTypeC ctx fst u
|
||||
-- if Ψ | Γ, x : A ⊢₀ B ⇐ Type ℓ
|
||||
case snd.body of
|
||||
Y snd => checkTypeC (extendTy fst ctx) snd u
|
||||
N snd => checkTypeC ctx snd u
|
||||
Y snd' => checkTypeC (extendTy snd.name fst ctx) snd' u
|
||||
N snd' => checkTypeC ctx snd' u
|
||||
-- then Ψ | Γ ⊢₀ (x : A) × B ⇐ Type ℓ
|
||||
|
||||
checkType' ctx t@(Pair {}) u =
|
||||
throwError $ NotType t
|
||||
throwError $ NotType ctx t
|
||||
|
||||
checkType' ctx (Enum _) u = pure ()
|
||||
-- Ψ | Γ ⊢₀ {ts} ⇐ Type ℓ
|
||||
|
||||
checkType' ctx t@(Tag {}) u =
|
||||
throwError $ NotType t
|
||||
throwError $ NotType ctx t
|
||||
|
||||
checkType' ctx (Eq t l r) u = do
|
||||
-- if Ψ, i | Γ ⊢₀ A ⇐ Type ℓ
|
||||
case t.body of
|
||||
Y t => checkTypeC (extendDim ctx) t u
|
||||
N t => checkTypeC ctx t u
|
||||
Y t' => checkTypeC (extendDim t.name ctx) t' u
|
||||
N t' => checkTypeC ctx t' u
|
||||
-- if Ψ | Γ ⊢₀ l ⇐ A‹0›
|
||||
check0 ctx t.zero l
|
||||
-- if Ψ | Γ ⊢₀ r ⇐ A‹1›
|
||||
|
@ -225,7 +227,7 @@ parameters {auto _ : IsQty q} {auto _ : CanTC q m}
|
|||
-- then Ψ | Γ ⊢₀ Eq [i ⇒ A] l r ⇐ Type ℓ
|
||||
|
||||
checkType' ctx t@(DLam {}) u =
|
||||
throwError $ NotType t
|
||||
throwError $ NotType ctx t
|
||||
|
||||
checkType' ctx (E e) u = do
|
||||
-- if Ψ | Γ ⊢ σ · e ⇒ A' ⊳ Σ
|
||||
|
@ -233,7 +235,7 @@ parameters {auto _ : IsQty q} {auto _ : CanTC q m}
|
|||
-- if Ψ | Γ ⊢ A' <: A
|
||||
case u of
|
||||
Just u => subtype ctx infres.type (TYPE u)
|
||||
Nothing => ignore $ expectTYPE !ask infres.type
|
||||
Nothing => ignore $ expectTYPE !ask ctx infres.type
|
||||
-- then Ψ | Γ ⊢ σ · e ⇐ A ⊳ Σ
|
||||
|
||||
|
||||
|
@ -268,7 +270,7 @@ parameters {auto _ : IsQty q} {auto _ : CanTC q m}
|
|||
infer' ctx sg (fun :@ arg) = do
|
||||
-- if Ψ | Γ ⊢ σ · f ⇒ (π·x : A) → B ⊳ Σ₁
|
||||
funres <- inferC ctx sg fun
|
||||
(qty, argty, res) <- expectPi !ask funres.type
|
||||
(qty, argty, res) <- expectPi !ask ctx funres.type
|
||||
-- if Ψ | Γ ⊢ σ ⨴ π · s ⇐ A ⊳ Σ₂
|
||||
argout <- checkC ctx (subjMult sg qty) arg argty
|
||||
-- then Ψ | Γ ⊢ σ · f s ⇒ B[s] ⊳ Σ₁ + Σ₂
|
||||
|
@ -283,12 +285,13 @@ parameters {auto _ : IsQty q} {auto _ : CanTC q m}
|
|||
-- if Ψ | Γ ⊢ σ · pair ⇒ (x : A) × B ⊳ Σ₁
|
||||
pairres <- inferC ctx sg pair
|
||||
-- if Ψ | Γ, p : (x : A) × B ⊢₀ ret ⇐ Type
|
||||
checkTypeC (extendTy pairres.type ctx) ret.term Nothing
|
||||
(tfst, tsnd) <- expectSig !ask pairres.type
|
||||
checkTypeC (extendTy ret.name pairres.type ctx) ret.term Nothing
|
||||
(tfst, tsnd) <- expectSig !ask ctx pairres.type
|
||||
-- if Ψ | Γ, x : A, y : B ⊢ σ · body ⇐
|
||||
-- ret[(x, y) ∷ (x : A) × B/p] ⊳ Σ₂, ρ₁·x, ρ₂·y
|
||||
-- with ρ₁, ρ₂ ≤ πσ
|
||||
let bodyctx = extendTyN [< tfst, tsnd.term] ctx
|
||||
let [x, y] = body.names
|
||||
bodyctx = extendTyN [< (x, tfst), (y, tsnd.term)] ctx
|
||||
bodyty = substCasePairRet pairres.type ret
|
||||
pisg = pi * sg.fst
|
||||
bodyout <- checkC bodyctx sg body.term bodyty
|
||||
|
@ -298,35 +301,36 @@ parameters {auto _ : IsQty q} {auto _ : CanTC q m}
|
|||
qout = pi * pairres.qout + !(popQs [< pisg, pisg] bodyout)
|
||||
}
|
||||
|
||||
infer' ctx sg (CaseEnum pi t ret arms) {n} = do
|
||||
infer' ctx sg (CaseEnum pi t ret arms) {d, n} = do
|
||||
-- if 1 ≤ π
|
||||
expectCompatQ one pi
|
||||
-- if Ψ | Γ ⊢ σ · t ⇒ {ts} ⊳ Σ₁
|
||||
tres <- inferC ctx sg t
|
||||
-- if Ψ | Γ, x : {ts} ⊢₀ A ⇐ Type
|
||||
checkTypeC (extendTy tres.type ctx) ret.term Nothing
|
||||
checkTypeC (extendTy ret.name tres.type ctx) ret.term Nothing
|
||||
-- if for each "a ⇒ s" in arms,
|
||||
-- Ψ | Γ ⊢ σ · s ⇐ A[a ∷ {ts}/x] ⊳ Σ₂
|
||||
-- for fixed Σ₂
|
||||
armres <- for (SortedMap.toList arms) $ \(a, s) =>
|
||||
let arms = SortedMap.toList arms
|
||||
armres <- for arms $ \(a, s) =>
|
||||
checkC ctx sg s (sub1 ret (Tag a :# tres.type))
|
||||
armout <- allEqual armres
|
||||
armout <- allEqual (zip armres arms)
|
||||
-- then Ψ | Γ ⊢ σ · case ⋯ ⇒ ret[t/x] ⊳ πΣ₁ + Σ₂
|
||||
pure $ InfRes {
|
||||
type = sub1 ret t,
|
||||
qout = pi * tres.qout + armout
|
||||
}
|
||||
where
|
||||
allEqual : List (QOutput q n) -> m (QOutput q n)
|
||||
allEqual : List (QOutput q n, TagVal, Term q d n) -> m (QOutput q n)
|
||||
allEqual [] = pure $ zeroFor ctx
|
||||
allEqual lst@(x :: xs) =
|
||||
if all (== x) xs then pure x
|
||||
else throwError $ BadCaseQtys lst
|
||||
allEqual lst@((x, _) :: xs) =
|
||||
if all (\y => x == fst y) xs then pure x
|
||||
else throwError $ BadCaseQtys ctx lst
|
||||
|
||||
infer' ctx sg (fun :% dim) = do
|
||||
-- if Ψ | Γ ⊢ σ · f ⇒ Eq [i ⇒ A] l r ⊳ Σ
|
||||
InfRes {type, qout} <- inferC ctx sg fun
|
||||
ty <- fst <$> expectEq !ask type
|
||||
ty <- fst <$> expectEq !ask ctx type
|
||||
-- then Ψ | Γ ⊢ σ · f p ⇒ A‹p› ⊳ Σ
|
||||
pure $ InfRes {type = dsub1 ty dim, qout}
|
||||
|
||||
|
|
|
@ -11,53 +11,6 @@ import public Quox.Reduce
|
|||
%default total
|
||||
|
||||
|
||||
namespace TContext
|
||||
export %inline
|
||||
pushD : TContext q d n -> TContext q (S d) n
|
||||
pushD tel = map (// shift 1) tel
|
||||
|
||||
export %inline
|
||||
zeroFor : IsQty q => Context tm n -> QOutput q n
|
||||
zeroFor ctx = zero <$ ctx
|
||||
|
||||
namespace TyContext
|
||||
public export %inline
|
||||
empty : {d : Nat} -> TyContext q d 0
|
||||
empty = MkTyContext {dctx = new, tctx = [<]}
|
||||
|
||||
export %inline
|
||||
extendTyN : Telescope (Term q d) from to ->
|
||||
TyContext q d from -> TyContext q d to
|
||||
extendTyN ss = {tctx $= (. ss)}
|
||||
|
||||
export %inline
|
||||
extendTy : Term q d n -> TyContext q d n -> TyContext q d (S n)
|
||||
extendTy s = extendTyN [< s]
|
||||
|
||||
export %inline
|
||||
extendDim : TyContext q d n -> TyContext q (S d) n
|
||||
extendDim = {dctx $= (:<? Nothing), tctx $= pushD}
|
||||
|
||||
export %inline
|
||||
eqDim : Dim d -> Dim d -> TyContext q d n -> TyContext q d n
|
||||
eqDim p q = {dctx $= set p q}
|
||||
|
||||
|
||||
namespace QOutput
|
||||
parameters {auto _ : IsQty q}
|
||||
export %inline
|
||||
(+) : QOutput q n -> QOutput q n -> QOutput q n
|
||||
(+) = zipWith (+)
|
||||
|
||||
export %inline
|
||||
(*) : q -> QOutput q n -> QOutput q n
|
||||
(*) pi = map (pi *)
|
||||
|
||||
export %inline
|
||||
zeroFor : TyContext q _ n -> QOutput q n
|
||||
zeroFor ctx = zeroFor ctx.tctx
|
||||
|
||||
|
||||
public export
|
||||
CheckResult' : Type -> Nat -> Type
|
||||
CheckResult' = QOutput
|
||||
|
@ -98,37 +51,89 @@ parameters {auto _ : HasErr q m} (defs : Definitions' q _)
|
|||
tm q d n -> m (NonRedex tm q d n defs)
|
||||
whnfT = either (throwError . WhnfError) pure . whnf defs
|
||||
|
||||
export covering %inline
|
||||
expectTYPE : Term q d n -> m Universe
|
||||
expectTYPE s =
|
||||
case fst !(whnfT s) of
|
||||
parameters (ctx : Lazy (TyContext q d1 n)) (th : Lazy (DSubst d2 d1))
|
||||
export covering %inline
|
||||
expectTYPE_ : Term q d2 n -> m Universe
|
||||
expectTYPE_ s = case fst !(whnfT s) of
|
||||
TYPE l => pure l
|
||||
_ => throwError $ ExpectedTYPE s
|
||||
_ => throwError $ ExpectedTYPE ctx (s // th)
|
||||
|
||||
export covering %inline
|
||||
expectPi : Term q d n -> m (q, Term q d n, ScopeTerm q d n)
|
||||
expectPi s =
|
||||
case fst !(whnfT s) of
|
||||
export covering %inline
|
||||
expectPi_ : Term q d2 n -> m (q, Term q d2 n, ScopeTerm q d2 n)
|
||||
expectPi_ s = case fst !(whnfT s) of
|
||||
Pi {qty, arg, res, _} => pure (qty, arg, res)
|
||||
_ => throwError $ ExpectedPi s
|
||||
_ => throwError $ ExpectedPi ctx (s // th)
|
||||
|
||||
export covering %inline
|
||||
expectSig : Term q d n -> m (Term q d n, ScopeTerm q d n)
|
||||
expectSig s =
|
||||
case fst !(whnfT s) of
|
||||
export covering %inline
|
||||
expectSig_ : Term q d2 n -> m (Term q d2 n, ScopeTerm q d2 n)
|
||||
expectSig_ s = case fst !(whnfT s) of
|
||||
Sig {fst, snd, _} => pure (fst, snd)
|
||||
_ => throwError $ ExpectedSig s
|
||||
_ => throwError $ ExpectedSig ctx (s // th)
|
||||
|
||||
export covering %inline
|
||||
expectEnum : Term q d n -> m (SortedSet TagVal)
|
||||
expectEnum s =
|
||||
case fst !(whnfT s) of
|
||||
export covering %inline
|
||||
expectEnum_ : Term q d2 n -> m (SortedSet TagVal)
|
||||
expectEnum_ s = case fst !(whnfT s) of
|
||||
Enum tags => pure tags
|
||||
_ => throwError $ ExpectedEnum s
|
||||
_ => throwError $ ExpectedEnum ctx (s // th)
|
||||
|
||||
export covering %inline
|
||||
expectEq : Term q d n -> m (DScopeTerm q d n, Term q d n, Term q d n)
|
||||
expectEq s =
|
||||
case fst !(whnfT s) of
|
||||
export covering %inline
|
||||
expectEq_ : Term q d2 n -> m (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)
|
||||
_ => throwError $ ExpectedEq s
|
||||
_ => throwError $ ExpectedEq ctx (s // th)
|
||||
|
||||
|
||||
-- [fixme] refactor this stuff
|
||||
|
||||
parameters (ctx : TyContext q d n)
|
||||
export covering %inline
|
||||
expectTYPE : Term q d n -> m Universe
|
||||
expectTYPE = expectTYPE_ ctx id
|
||||
|
||||
export covering %inline
|
||||
expectPi : Term q d n -> m (q, Term q d n, ScopeTerm q d n)
|
||||
expectPi = expectPi_ ctx id
|
||||
|
||||
export covering %inline
|
||||
expectSig : Term q d n -> m (Term q d n, ScopeTerm q d n)
|
||||
expectSig = expectSig_ ctx id
|
||||
|
||||
export covering %inline
|
||||
expectEnum : Term q d n -> m (SortedSet TagVal)
|
||||
expectEnum = expectEnum_ ctx id
|
||||
|
||||
export covering %inline
|
||||
expectEq : Term q d n -> m (DScopeTerm q d n, Term q d n, Term q d n)
|
||||
expectEq = expectEq_ ctx id
|
||||
|
||||
|
||||
parameters (ctx : EqContext q n)
|
||||
export covering %inline
|
||||
expectTYPEE : Term q 0 n -> m Universe
|
||||
expectTYPEE t =
|
||||
let ctx = delay (toTyContext ctx) in
|
||||
expectTYPE_ (DPair.snd ctx) (shift0 $ DPair.fst ctx) t
|
||||
|
||||
export covering %inline
|
||||
expectPiE : Term q 0 n -> m (q, Term q 0 n, ScopeTerm q 0 n)
|
||||
expectPiE t =
|
||||
let ctx = delay (toTyContext ctx) in
|
||||
expectPi_ (DPair.snd ctx) (shift0 $ DPair.fst ctx) t
|
||||
|
||||
export covering %inline
|
||||
expectSigE : Term q 0 n -> m (Term q 0 n, ScopeTerm q 0 n)
|
||||
expectSigE t =
|
||||
let ctx = delay (toTyContext ctx) in
|
||||
expectSig_ (DPair.snd ctx) (shift0 $ DPair.fst ctx) t
|
||||
|
||||
export covering %inline
|
||||
expectEnumE : Term q 0 n -> m (SortedSet TagVal)
|
||||
expectEnumE t =
|
||||
let ctx = delay (toTyContext ctx) in
|
||||
expectEnum_ (DPair.snd ctx) (shift0 $ DPair.fst ctx) t
|
||||
|
||||
export covering %inline
|
||||
expectEqE : Term q 0 n -> m (DScopeTerm q 0 n, Term q 0 n, Term q 0 n)
|
||||
expectEqE t =
|
||||
let ctx = delay (toTyContext ctx) in
|
||||
expectEq_ (DPair.snd ctx) (shift0 $ DPair.fst ctx) t
|
||||
|
|
|
@ -3,6 +3,8 @@ module Quox.Typing.Context
|
|||
import Quox.Syntax
|
||||
import Quox.Context
|
||||
|
||||
%default total
|
||||
|
||||
|
||||
public export
|
||||
TContext : Type -> Nat -> Nat -> Type
|
||||
|
@ -12,11 +14,125 @@ public export
|
|||
QOutput : Type -> Nat -> Type
|
||||
QOutput = Context'
|
||||
|
||||
public export
|
||||
NContext : Nat -> Type
|
||||
NContext = Context' BaseName
|
||||
|
||||
public export
|
||||
DimAssign : Nat -> Type
|
||||
DimAssign = Context' DimConst
|
||||
|
||||
|
||||
public export
|
||||
record TyContext q d n where
|
||||
constructor MkTyContext
|
||||
dctx : DimEq d
|
||||
tctx : TContext q d n
|
||||
|
||||
dctx : DimEq d
|
||||
dnames : NContext d
|
||||
tctx : TContext q d n
|
||||
tnames : NContext n
|
||||
%name TyContext ctx
|
||||
|
||||
|
||||
public export
|
||||
record EqContext q n where
|
||||
constructor MkEqContext
|
||||
-- (only used for errors; not needed by the actual equality test)
|
||||
dassign : DimAssign dimLen
|
||||
dnames : NContext dimLen
|
||||
tctx : TContext q 0 n
|
||||
tnames : NContext n
|
||||
%name EqContext ctx
|
||||
|
||||
|
||||
namespace TContext
|
||||
export %inline
|
||||
pushD : TContext q d n -> TContext q (S d) n
|
||||
pushD tel = map (// shift 1) tel
|
||||
|
||||
export %inline
|
||||
zeroFor : IsQty q => Context tm n -> QOutput q n
|
||||
zeroFor ctx = zero <$ ctx
|
||||
|
||||
namespace TyContext
|
||||
public export %inline
|
||||
empty : TyContext q 0 0
|
||||
empty = MkTyContext {dctx = new, dnames = [<], tctx = [<], tnames = [<]}
|
||||
|
||||
export %inline
|
||||
extendTyN : Telescope (\n => (BaseName, Term q d n)) from to ->
|
||||
TyContext q d from -> TyContext q d to
|
||||
extendTyN xss ctx =
|
||||
let (xs, ss) = unzip xss in {tctx $= (. ss), tnames $= (. xs)} ctx
|
||||
|
||||
export %inline
|
||||
extendTy : BaseName -> Term q d n -> TyContext q d n -> TyContext q d (S n)
|
||||
extendTy x s = extendTyN [< (x, s)]
|
||||
|
||||
export %inline
|
||||
extendDim : BaseName -> TyContext q d n -> TyContext q (S d) n
|
||||
extendDim x = {dctx $= (:<? Nothing), dnames $= (:< x), tctx $= pushD}
|
||||
|
||||
export %inline
|
||||
eqDim : Dim d -> Dim d -> TyContext q d n -> TyContext q d n
|
||||
eqDim p q = {dctx $= set p q}
|
||||
|
||||
|
||||
namespace QOutput
|
||||
parameters {auto _ : IsQty q}
|
||||
export %inline
|
||||
(+) : QOutput q n -> QOutput q n -> QOutput q n
|
||||
(+) = zipWith (+)
|
||||
|
||||
export %inline
|
||||
(*) : q -> QOutput q n -> QOutput q n
|
||||
(*) pi = map (pi *)
|
||||
|
||||
export %inline
|
||||
zeroFor : TyContext q _ n -> QOutput q n
|
||||
zeroFor ctx = zeroFor ctx.tctx
|
||||
|
||||
|
||||
export
|
||||
makeDAssign : DSubst d 0 -> DimAssign d
|
||||
makeDAssign (Shift SZ) = [<]
|
||||
makeDAssign (K e ::: th) = makeDAssign th :< e
|
||||
|
||||
export
|
||||
makeEqContext : TyContext q d n -> DSubst d 0 -> EqContext q n
|
||||
makeEqContext ctx th = MkEqContext {
|
||||
dassign = makeDAssign th,
|
||||
dnames = ctx.dnames,
|
||||
tctx = map (// th) ctx.tctx,
|
||||
tnames = ctx.tnames
|
||||
}
|
||||
|
||||
namespace EqContext
|
||||
public export %inline
|
||||
empty : EqContext q 0
|
||||
empty = MkEqContext [<] [<] [<] [<]
|
||||
|
||||
export %inline
|
||||
extendTyN : Telescope (\n => (BaseName, Term q 0 n)) from to ->
|
||||
EqContext q from -> EqContext q to
|
||||
extendTyN tel ctx =
|
||||
let (xs, ss) = unzip tel in {tctx $= (. ss), tnames $= (. xs)} ctx
|
||||
|
||||
export %inline
|
||||
extendTy : BaseName -> Term q 0 n -> EqContext q n -> EqContext q (S n)
|
||||
extendTy x s = extendTyN [< (x, s)]
|
||||
|
||||
export %inline
|
||||
extendDim : BaseName -> DimConst -> EqContext q n -> EqContext q n
|
||||
extendDim x e ctx = {dassign $= (:< e), dnames $= (:< x)} ctx
|
||||
|
||||
export
|
||||
toTyContext : (e : EqContext q n) -> (d ** TyContext q d n)
|
||||
toTyContext (MkEqContext {dassign, dnames, tctx, tnames})
|
||||
with (lengthPrf0 dnames)
|
||||
_ | Element d eq =
|
||||
(d ** MkTyContext {
|
||||
dctx = rewrite eq in fromGround dassign,
|
||||
dnames = rewrite eq in dnames,
|
||||
tctx = map (// shift0 d) tctx,
|
||||
tnames
|
||||
})
|
||||
|
|
|
@ -12,26 +12,25 @@ import public Control.Monad.Either
|
|||
|
||||
public export
|
||||
data Error q
|
||||
= ExpectedTYPE (Term q d n)
|
||||
| ExpectedPi (Term q d n)
|
||||
| ExpectedSig (Term q d n)
|
||||
| ExpectedEnum (Term q d n)
|
||||
| ExpectedEq (Term q d n)
|
||||
= ExpectedTYPE (TyContext q d n) (Term q d n)
|
||||
| ExpectedPi (TyContext q d n) (Term q d n)
|
||||
| ExpectedSig (TyContext q d n) (Term q d n)
|
||||
| ExpectedEnum (TyContext q d n) (Term q d n)
|
||||
| ExpectedEq (TyContext q d n) (Term q d n)
|
||||
| BadUniverse Universe Universe
|
||||
| TagNotIn TagVal (SortedSet TagVal)
|
||||
| BadCaseQtys (List (QOutput q n))
|
||||
| BadCaseQtys (TyContext q d n) (List (QOutput q n, TagVal, Term q d n))
|
||||
|
||||
-- first arg of ClashT is the type
|
||||
| ClashT EqMode (Term q d n) (Term q d n) (Term q d n)
|
||||
| ClashTy EqMode (Term q d n) (Term q d n)
|
||||
| ClashE EqMode (Elim q d n) (Elim q d n)
|
||||
| ClashU EqMode Universe Universe
|
||||
| ClashQ q q
|
||||
| ClashD (Dim d) (Dim d)
|
||||
| NotInScope Name
|
||||
-- first term arg of ClashT is the type
|
||||
| ClashT (EqContext q n) EqMode (Term q 0 n) (Term q 0 n) (Term q 0 n)
|
||||
| ClashTy (EqContext q n) EqMode (Term q 0 n) (Term q 0 n)
|
||||
| ClashE (EqContext q n) EqMode (Elim q 0 n) (Elim q 0 n)
|
||||
| ClashU EqMode Universe Universe
|
||||
| ClashQ q q
|
||||
| NotInScope Name
|
||||
|
||||
| NotType (Term q d n)
|
||||
| WrongType (Term q d n) (Term q d n) (Term q d n)
|
||||
| NotType (TyContext q d n) (Term q d n)
|
||||
| WrongType (EqContext q n) (Term q 0 n) (Term q 0 n) (Term q 0 n)
|
||||
|
||||
-- extra context
|
||||
| WhileChecking
|
||||
|
@ -49,13 +48,13 @@ data Error q
|
|||
(Elim q d n)
|
||||
(Error q)
|
||||
| WhileComparingT
|
||||
(TyContext q d n) EqMode
|
||||
(Term q d n) -- type
|
||||
(Term q d n) (Term q d n) -- lhs/rhs
|
||||
(EqContext q n) EqMode
|
||||
(Term q 0 n) -- type
|
||||
(Term q 0 n) (Term q 0 n) -- lhs/rhs
|
||||
(Error q)
|
||||
| WhileComparingE
|
||||
(TyContext q d n) EqMode
|
||||
(Elim q d n) (Elim q d n)
|
||||
(EqContext q n) EqMode
|
||||
(Elim q 0 n) (Elim q 0 n)
|
||||
(Error q)
|
||||
|
||||
| WhnfError WhnfError
|
||||
|
@ -137,6 +136,3 @@ parameters {auto _ : HasErr q m}
|
|||
expectModeU : EqMode -> Universe -> Universe -> m ()
|
||||
expectModeU mode = expect (ClashU mode) $ ucmp mode
|
||||
|
||||
export %inline
|
||||
expectEqualD : Dim d -> Dim d -> m ()
|
||||
expectEqualD = expect ClashD (==)
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue