pretty printing errors

This commit is contained in:
rhiannon morris 2023-03-15 15:54:51 +01:00
parent 54ba4e237f
commit 32f38238ef
14 changed files with 424 additions and 217 deletions

View file

@ -107,22 +107,14 @@ parameters (defs : Definitions' q g)
E _ => pure False
parameters {auto _ : HasErr q m}
export %inline
ensure : (a -> Error q) -> (p : a -> Bool) -> (t : a) -> m (So (p t))
ensure e p t = case nchoose $ p t of
Left y => pure y
Right _ => throwError $ e t
export
ensureTyCon : HasErr q m =>
(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 => throwError $ NotType (toTyContext ctx) (t // shift0 ctx.dimLen)
export %inline
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)}
parameters (defs : Definitions' q _) {auto _ : (CanEqual q m, IsQty q)}
mutual
namespace Term
||| `compare0 ctx ty s t` compares `s` and `t` at type `ty`, according to
@ -153,24 +145,27 @@ parameters (defs : Definitions' q _) {auto _ : (CanEqual q m, Eq q)}
m ()
compare0' ctx (TYPE _) s t = compareType ctx s t
compare0' ctx ty@(Pi {arg, res, _}) s t {n} = local {mode := Equal} $
compare0' ctx ty@(Pi {qty, arg, res}) s t {n} = local {mode := Equal} $
case (s, t) of
-- Γ, x : A ⊢ s = t : B
-- -----------------------------------------
-- Γ ⊢ (λx ⇒ s) = (λx ⇒ t) : (π·x : A) → B
-- Γ, x : A ⊢ s = t : B
-- -------------------------------------------
-- Γ ⊢ (λ x ⇒ s) = (λ x ⇒ t) : (π·x : A) → B
(Lam b1, Lam b2) => compare0 ctx' res.term b1.term b2.term
-- Γ, x : A ⊢ s = e x : B
-- ----------------------------------
-- Γ ⊢ (λx ⇒ s) = e : (π·x : A) → B
-- -----------------------------------
-- Γ ⊢ (λ x ⇒ s) = e : (π·x : A) → B
(E e, Lam b) => eta e b
(Lam b, E e) => eta e b
(E e, E f) => Elim.compare0 ctx e f
_ => throwError $ WrongType ctx ty s t
(Lam _, t) => throwError $ WrongType ctx ty t
(E _, t) => throwError $ WrongType ctx ty t
(s, _) => throwError $ WrongType ctx ty s
where
ctx' : EqContext q (S n)
ctx' = extendTy res.name arg ctx
ctx' = extendTy qty 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
@ -179,16 +174,19 @@ parameters (defs : Definitions' q _) {auto _ : (CanEqual q m, Eq q)}
compare0' ctx ty@(Sig {fst, snd, _}) s t = local {mode := Equal} $
case (s, t) of
-- Γ ⊢ s₁ = t₁ : A Γ ⊢ s₂ = t₂ : B{s₁/x}
-- -------------------------------------------
-- Γ ⊢ (s₁,t₁) = (s₂,t₂) : (x : A) × B
-- --------------------------------------------
-- Γ ⊢ (s₁, t₁) = (s₂,t₂) : (x : A) × B
--
-- [todo] η for π ≥ 0 maybe
(Pair sFst sSnd, Pair tFst tSnd) => do
compare0 ctx fst sFst tFst
compare0 ctx (sub1 snd (sFst :# fst)) sSnd tSnd
(E e, E f) => compare0 ctx e f
_ => throwError $ WrongType ctx ty s t
(E e, E f) => Elim.compare0 ctx e f
(Pair {}, t) => throwError $ WrongType ctx ty t
(E _, t) => throwError $ WrongType ctx ty t
(s, _) => throwError $ WrongType ctx ty s
compare0' ctx ty@(Enum tags) s t = local {mode := Equal} $
case (s, t) of
@ -197,8 +195,11 @@ parameters (defs : Definitions' q _) {auto _ : (CanEqual q m, Eq q)}
--
-- t ∈ ts is in the typechecker, not here, ofc
(Tag t1, Tag t2) => unless (t1 == t2) $ clashT ctx ty s t
(E e, E f) => compare0 ctx e f
_ => throwError $ WrongType ctx ty s t
(E e, E f) => Elim.compare0 ctx e f
(Tag _, t) => throwError $ WrongType ctx ty t
(E _, t) => throwError $ WrongType ctx ty t
(s, _) => throwError $ WrongType ctx ty s
compare0' _ (Eq {}) _ _ =
-- ✨ uip ✨
@ -209,8 +210,8 @@ 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 ctx ty s t
E f <- pure t | _ => throwError $ WrongType ctx ty s t
E e <- pure s | _ => throwError $ WrongType ctx ty s
E f <- pure t | _ => throwError $ WrongType ctx ty t
Elim.compare0 ctx e f
||| compares two types, using the current variance `mode` for universes.
@ -247,7 +248,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 (extendTy sRes.name sArg ctx) sRes.term tRes.term
compareType (extendTy zero sRes.name sArg ctx) sRes.term tRes.term
compareType' ctx (Sig {fst = sFst, snd = sSnd, _})
(Sig {fst = tFst, snd = tSnd, _}) = do
@ -255,7 +256,7 @@ parameters (defs : Definitions' q _) {auto _ : (CanEqual q m, Eq q)}
-- --------------------------------------
-- Γ ⊢ (x : A₁) × B₁ <: (x : A₂) × B₂
compareType ctx sFst tFst
compareType (extendTy sSnd.name sFst ctx) sSnd.term tSnd.term
compareType (extendTy zero 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
@ -365,10 +366,10 @@ parameters (defs : Definitions' q _) {auto _ : (CanEqual q m, Eq q)}
local {mode := Equal} $ do
compare0 ctx e f
ety <- computeElimType ctx e (noOr1 ne)
compareType (extendTy eret.name ety ctx) eret.term fret.term
compareType (extendTy zero 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)
Term.compare0 (extendTyN [< (epi, x, fst), (epi, y, snd.term)] ctx)
(substCasePairRet ety eret)
ebody.term fbody.term
expectEqualQ epi fpi
@ -379,7 +380,7 @@ parameters (defs : Definitions' q _) {auto _ : (CanEqual q m, Eq q)}
local {mode := Equal} $ do
compare0 ctx e f
ety <- computeElimType ctx e (noOr1 ne)
compareType (extendTy eret.name ety ctx) eret.term fret.term
compareType (extendTy zero 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)
@ -402,7 +403,8 @@ parameters (defs : Definitions' q _) {auto _ : (CanEqual q m, Eq q)}
compare0' ctx e@(_ :# _) f _ _ = clashE ctx e f
parameters {auto _ : (HasDefs' q _ m, HasErr q m, Eq q)} (ctx : TyContext q d n)
parameters {auto _ : (HasDefs' q _ m, HasErr q m, IsQty q)}
(ctx : TyContext q d n)
-- [todo] only split on the dvars that are actually used anywhere in
-- the calls to `splits`