wrap type errors in extra context
This commit is contained in:
parent
858b5db530
commit
85a55f8123
4 changed files with 58 additions and 27 deletions
|
@ -116,12 +116,13 @@ parameters (defs : Definitions' q _) {auto _ : (CanEqual q m, Eq q)}
|
||||||
||| ⚠ **assumes that `s`, `t` have already been checked against `ty`**. ⚠
|
||| ⚠ **assumes that `s`, `t` have already been checked against `ty`**. ⚠
|
||||||
export covering %inline
|
export covering %inline
|
||||||
compare0 : TContext q 0 n -> (ty, s, t : Term q 0 n) -> m ()
|
compare0 : TContext q 0 n -> (ty, s, t : Term q 0 n) -> m ()
|
||||||
compare0 ctx ty s t = do
|
compare0 ctx ty s t =
|
||||||
let Element ty nty = whnfD defs ty
|
wrapErr (WhileComparingT (MkTyContext new ctx) !mode ty s t) $ do
|
||||||
Element s ns = whnfD defs s
|
let Element ty nty = whnfD defs ty
|
||||||
Element t nt = whnfD defs t
|
Element s ns = whnfD defs s
|
||||||
tty <- ensureTyCon ty
|
Element t nt = whnfD defs t
|
||||||
compare0' ctx ty s t
|
tty <- ensureTyCon ty
|
||||||
|
compare0' ctx ty s t
|
||||||
|
|
||||||
||| converts an elim "Γ ⊢ e" to "Γ, x ⊢ e x", for comparing with
|
||| converts an elim "Γ ⊢ e" to "Γ, x ⊢ e x", for comparing with
|
||||||
||| a lambda "Γ ⊢ λx ⇒ t" that has been converted to "Γ, x ⊢ t".
|
||| a lambda "Γ ⊢ λx ⇒ t" that has been converted to "Γ, x ⊢ t".
|
||||||
|
@ -289,12 +290,12 @@ parameters (defs : Definitions' q _) {auto _ : (CanEqual q m, Eq q)}
|
||||||
export covering %inline
|
export covering %inline
|
||||||
compare0 : TContext q 0 n -> (e, f : Elim q 0 n) -> m ()
|
compare0 : TContext q 0 n -> (e, f : Elim q 0 n) -> m ()
|
||||||
compare0 ctx e f =
|
compare0 ctx e f =
|
||||||
let Element e ne = whnfD defs e
|
wrapErr (WhileComparingE (MkTyContext new ctx) !mode e f) $ do
|
||||||
Element f nf = whnfD defs f
|
let Element e ne = whnfD defs e
|
||||||
in
|
Element f nf = whnfD defs f
|
||||||
-- [fixme] there is a better way to do this "isSubSing" stuff for sure
|
-- [fixme] there is a better way to do this "isSubSing" stuff for sure
|
||||||
unless (isSubSing defs !(computeElimType ctx e ne)) $
|
unless (isSubSing defs !(computeElimType ctx e ne)) $
|
||||||
compare0' ctx e f ne nf
|
compare0' ctx e f ne nf
|
||||||
|
|
||||||
private covering
|
private covering
|
||||||
compare0' : TContext q 0 n ->
|
compare0' : TContext q 0 n ->
|
||||||
|
|
|
@ -80,8 +80,9 @@ parameters {auto _ : IsQty q} {auto _ : CanTC q m}
|
||||||
checkC : (ctx : TyContext q d n) -> SQty q -> Term q d n -> Term q d n ->
|
checkC : (ctx : TyContext q d n) -> SQty q -> Term q d n -> Term q d n ->
|
||||||
m (CheckResult' q n)
|
m (CheckResult' q n)
|
||||||
checkC ctx sg subj ty =
|
checkC ctx sg subj ty =
|
||||||
let Element subj nc = pushSubsts subj in
|
wrapErr (WhileChecking ctx sg.fst subj ty) $
|
||||||
check' ctx sg subj nc ty
|
let Element subj nc = pushSubsts subj in
|
||||||
|
check' ctx sg subj nc ty
|
||||||
|
|
||||||
|
|
||||||
||| "Ψ | Γ ⊢ σ · e ⇒ A ⊳ Σ"
|
||| "Ψ | Γ ⊢ σ · e ⇒ A ⊳ Σ"
|
||||||
|
@ -104,8 +105,9 @@ parameters {auto _ : IsQty q} {auto _ : CanTC q m}
|
||||||
inferC : (ctx : TyContext q d n) -> SQty q -> Elim q d n ->
|
inferC : (ctx : TyContext q d n) -> SQty q -> Elim q d n ->
|
||||||
m (InferResult' q d n)
|
m (InferResult' q d n)
|
||||||
inferC ctx sg subj =
|
inferC ctx sg subj =
|
||||||
let Element subj nc = pushSubsts subj in
|
wrapErr (WhileInferring ctx sg.fst subj) $
|
||||||
infer' ctx sg subj nc
|
let Element subj nc = pushSubsts subj in
|
||||||
|
infer' ctx sg subj nc
|
||||||
|
|
||||||
|
|
||||||
private covering
|
private covering
|
||||||
|
|
|
@ -124,10 +124,34 @@ data Error q
|
||||||
| NotType (Term q d n)
|
| NotType (Term q d n)
|
||||||
| WrongType (Term q d n) (Term q d n) (Term q d n)
|
| WrongType (Term q d n) (Term q d n) (Term q d n)
|
||||||
|
|
||||||
|
-- extra context
|
||||||
|
| WhileChecking
|
||||||
|
(TyContext q d n) q
|
||||||
|
(Term q d n) -- term
|
||||||
|
(Term q d n) -- type
|
||||||
|
(Error q)
|
||||||
|
| WhileInferring
|
||||||
|
(TyContext q d n) 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
|
||||||
|
(Error q)
|
||||||
|
| WhileComparingE
|
||||||
|
(TyContext q d n) EqMode
|
||||||
|
(Elim q d n) (Elim q d n)
|
||||||
|
(Error q)
|
||||||
|
|
||||||
public export
|
public export
|
||||||
0 HasErr : Type -> (Type -> Type) -> Type
|
0 HasErr : Type -> (Type -> Type) -> Type
|
||||||
HasErr q = MonadError (Error q)
|
HasErr q = MonadError (Error q)
|
||||||
|
|
||||||
|
export %inline
|
||||||
|
wrapErr : HasErr q m => (Error q -> Error q) -> m a -> m a
|
||||||
|
wrapErr f act = catchError act $ throwError . f
|
||||||
|
|
||||||
export %inline
|
export %inline
|
||||||
ucmp : EqMode -> Universe -> Universe -> Bool
|
ucmp : EqMode -> Universe -> Universe -> Bool
|
||||||
ucmp Equal = (==)
|
ucmp Equal = (==)
|
||||||
|
|
|
@ -11,51 +11,55 @@ PrettyHL q => ToInfo (Error q) where
|
||||||
("name", show x)]
|
("name", show x)]
|
||||||
toInfo (ExpectedTYPE t) =
|
toInfo (ExpectedTYPE t) =
|
||||||
[("type", "ExpectedTYPE"),
|
[("type", "ExpectedTYPE"),
|
||||||
("got", prettyStr True t)]
|
("got", prettyStr True t)]
|
||||||
toInfo (ExpectedPi t) =
|
toInfo (ExpectedPi t) =
|
||||||
[("type", "ExpectedPi"),
|
[("type", "ExpectedPi"),
|
||||||
("got", prettyStr True t)]
|
("got", prettyStr True t)]
|
||||||
toInfo (ExpectedSig t) =
|
toInfo (ExpectedSig t) =
|
||||||
[("type", "ExpectedSig"),
|
[("type", "ExpectedSig"),
|
||||||
("got", prettyStr True t)]
|
("got", prettyStr True t)]
|
||||||
toInfo (ExpectedEq t) =
|
toInfo (ExpectedEq t) =
|
||||||
[("type", "ExpectedEq"),
|
[("type", "ExpectedEq"),
|
||||||
("got", prettyStr True t)]
|
("got", prettyStr True t)]
|
||||||
toInfo (BadUniverse k l) =
|
toInfo (BadUniverse k l) =
|
||||||
[("type", "BadUniverse"),
|
[("type", "BadUniverse"),
|
||||||
("low", show k),
|
("low", show k),
|
||||||
("high", show l)]
|
("high", show l)]
|
||||||
toInfo (ClashT mode ty s t) =
|
toInfo (ClashT mode ty s t) =
|
||||||
[("type", "ClashT"),
|
[("type", "ClashT"),
|
||||||
("mode", show mode),
|
("mode", show mode),
|
||||||
("ty", prettyStr True ty),
|
("ty", prettyStr True ty),
|
||||||
("left", prettyStr True s),
|
("left", prettyStr True s),
|
||||||
("right", prettyStr True t)]
|
("right", prettyStr True t)]
|
||||||
toInfo (ClashE mode e f) =
|
toInfo (ClashE mode e f) =
|
||||||
[("type", "ClashE"),
|
[("type", "ClashE"),
|
||||||
("mode", show mode),
|
("mode", show mode),
|
||||||
("left", prettyStr True e),
|
("left", prettyStr True e),
|
||||||
("right", prettyStr True f)]
|
("right", prettyStr True f)]
|
||||||
toInfo (ClashU mode k l) =
|
toInfo (ClashU mode k l) =
|
||||||
[("type", "ClashU"),
|
[("type", "ClashU"),
|
||||||
("mode", show mode),
|
("mode", show mode),
|
||||||
("left", prettyStr True k),
|
("left", prettyStr True k),
|
||||||
("right", prettyStr True l)]
|
("right", prettyStr True l)]
|
||||||
toInfo (ClashQ pi rh) =
|
toInfo (ClashQ pi rh) =
|
||||||
[("type", "ClashQ"),
|
[("type", "ClashQ"),
|
||||||
("left", prettyStr True pi),
|
("left", prettyStr True pi),
|
||||||
("right", prettyStr True rh)]
|
("right", prettyStr True rh)]
|
||||||
toInfo (ClashD p q) =
|
toInfo (ClashD p q) =
|
||||||
[("type", "ClashD"),
|
[("type", "ClashD"),
|
||||||
("left", prettyStr True p),
|
("left", prettyStr True p),
|
||||||
("right", prettyStr True q)]
|
("right", prettyStr True q)]
|
||||||
toInfo (NotType ty) =
|
toInfo (NotType ty) =
|
||||||
[("type", "NotType"),
|
[("type", "NotType"),
|
||||||
("actual", prettyStr True ty)]
|
("actual", prettyStr True ty)]
|
||||||
toInfo (WrongType ty s t) =
|
toInfo (WrongType ty s t) =
|
||||||
[("type", "WrongType"),
|
[("type", "WrongType"),
|
||||||
("ty", prettyStr True ty),
|
("ty", prettyStr True ty),
|
||||||
("left", prettyStr True s),
|
("left", prettyStr True s),
|
||||||
("right", prettyStr True t)]
|
("right", prettyStr True t)]
|
||||||
|
|
||||||
|
-- [todo] add nested yamls to TAP and include context here
|
||||||
|
toInfo (WhileChecking _ _ _ _ err) = toInfo err
|
||||||
|
toInfo (WhileInferring _ _ _ err) = toInfo err
|
||||||
|
toInfo (WhileComparingT _ _ _ _ _ err) = toInfo err
|
||||||
|
toInfo (WhileComparingE _ _ _ _ err) = toInfo err
|
||||||
|
|
Loading…
Reference in a new issue