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

@ -1,8 +1,10 @@
module Quox.Syntax.Dim
import Quox.Name
import Quox.Syntax.Var
import Quox.Syntax.Subst
import Quox.Pretty
import Quox.Context
import Decidable.Equality
import Control.Function
@ -31,29 +33,26 @@ data Dim : Nat -> Type where
B : Var d -> Dim d
%name Dim.Dim p, q
private %inline
drepr : Dim n -> Either DimConst (Var n)
drepr (K k) = Left k
drepr (B x) = Right x
export Eq (Dim n) where (==) = (==) `on` drepr
export Ord (Dim n) where compare = compare `on` drepr
export
Show (Dim n) where
show (K k) = showCon App "K" $ show k
show (B i) = showCon App "B" $ show i
%runElab deriveIndexed "Dim" [Eq, Ord, Show]
export
PrettyHL DimConst where
prettyM Zero = hl Dim <$> ifUnicode "𝟬" "0"
prettyM One = hl Dim <$> ifUnicode "𝟭" "1"
prettyM = pure . hl Dim . ends "0" "1"
export
PrettyHL (Dim n) where
prettyM (K e) = prettyM e
prettyM (B i) = prettyVar DVar DVarErr (!ask).dnames i
export
prettyDim : (dnames : NContext d) -> Dim d -> Doc HL
prettyDim dnames p =
let env = MakePrettyEnv {
dnames = toSnocList' dnames, tnames = [<],
unicode = True, prec = Outer
} in
runReader env $ prettyM p
||| `endsOr l r x e` returns:
||| - `l` if `p` is `K Zero`;

View file

@ -19,9 +19,9 @@ export
PrettyHL Three where
prettyM pi = hl Qty <$>
case pi of
Zero => ifUnicode "𝟬" "0"
One => ifUnicode "𝟭" "1"
Any => ifUnicode "𝛚" "*"
Zero => pure "0"
One => pure "1"
Any => ifUnicode "ω" "#"
public export
DecEq Three where
@ -88,6 +88,7 @@ data IsGlobal3 : Pred Three where
GZero : IsGlobal3 Zero
GAny : IsGlobal3 Any
public export
isGlobal3 : Dec1 IsGlobal3
isGlobal3 Zero = Yes GZero
isGlobal3 One = No $ \case _ impossible

View file

@ -3,6 +3,7 @@ module Quox.Syntax.Term.Pretty
import Quox.Syntax.Term.Base
import Quox.Syntax.Term.Split
import Quox.Syntax.Term.Subst
import Quox.Context
import Quox.Pretty
import Data.Vect
@ -42,16 +43,26 @@ prettyUnivSuffix l =
'0' => ''; '1' => ''; '2' => ''; '3' => ''; '4' => ''
'5' => ''; '6' => ''; '7' => ''; '8' => ''; '9' => ''; _ => c
export
prettyUniverse : Universe -> Doc HL
prettyUniverse = hl Syntax . pretty
export
prettyBind : PrettyHL a => PrettyHL q => Pretty.HasEnv m =>
List q -> BaseName -> a -> m (Doc HL)
prettyBind qtys x s = do
var <- prettyQtyBinds qtys $ TV x
s <- withPrec Outer $ prettyM s
pure $ var <++> colonD <%%> hang 2 s
export
prettyBindType : PrettyHL a => PrettyHL b => PrettyHL q =>
Pretty.HasEnv m =>
List q -> BaseName -> a -> Doc HL -> b -> m (Doc HL)
prettyBindType qtys x s arr t = do
var <- prettyQtyBinds qtys $ TV x
s <- withPrec Outer $ prettyM s
t <- withPrec Outer $ under T x $ prettyM t
let bind = parens (var <++> colonD <%%> hang 2 s)
parensIfM Outer $ hang 2 $ bind <%%> t
bind <- prettyBind qtys x s
t <- withPrec Outer $ under T x $ prettyM t
parensIfM Outer $ hang 2 $ parens bind <++> arr <%%> t
export
prettyArm : PrettyHL a => Pretty.HasEnv m =>
@ -144,7 +155,7 @@ parameters (showSubsts : Bool)
prettyM (DLam (S i t)) =
let GotDLams {names, body, _} = getDLams' i t.term Refl in
prettyLams (Just !dlamD) D (toList names) body
prettyM (E e) = bracks <$> prettyM e
prettyM (E e) = prettyM e
prettyM (CloT s th) =
if showSubsts then
parensIfM SApp . hang 2 =<<
@ -207,3 +218,16 @@ PrettyHL q => PrettyHL (Term q d n) where
export covering %inline
PrettyHL q => PrettyHL (Elim q d n) where
prettyM = prettyM @{ElimSubst False}
export covering
prettyTerm : PrettyHL q => (unicode : Bool) ->
(dnames : NContext d) -> (tnames : NContext n) ->
Term q d n -> Doc HL
prettyTerm unicode dnames tnames term =
let env = MakePrettyEnv {
dnames = toSnocList' dnames,
tnames = toSnocList' tnames,
unicode, prec = Outer
} in
runReader env $ prettyM term