quox/lib/Quox/Syntax/Term/Pretty.idr

424 lines
13 KiB
Idris
Raw Normal View History

2022-04-23 18:21:30 -04:00
module Quox.Syntax.Term.Pretty
import Quox.Syntax.Term.Base
import Quox.Syntax.Term.Split
import Quox.Syntax.Term.Subst
2023-03-15 10:54:51 -04:00
import Quox.Context
2022-04-23 18:21:30 -04:00
import Quox.Pretty
import Data.Vect
2022-05-02 16:38:37 -04:00
%default total
2022-04-23 18:21:30 -04:00
2023-03-26 08:40:54 -04:00
export %inline
typeD, arrowD, darrowD, timesD, lamD, eqndD, dlamD, annD, natD :
2023-01-26 13:54:46 -05:00
Pretty.HasEnv m => m (Doc HL)
2023-02-25 13:14:26 -05:00
typeD = hlF Syntax $ ifUnicode "" "Type"
2023-04-15 09:13:01 -04:00
arrowD = hlF Delim $ ifUnicode "" "->"
darrowD = hlF Delim $ ifUnicode "" "=>"
timesD = hlF Delim $ ifUnicode "×" "**"
2023-02-25 13:14:11 -05:00
lamD = hlF Syntax $ ifUnicode "λ" "fun"
2023-04-15 09:13:01 -04:00
eqndD = hlF Delim $ ifUnicode "" "=="
2023-02-25 13:14:11 -05:00
dlamD = hlF Syntax $ ifUnicode "δ" "dfun"
2023-04-15 09:13:01 -04:00
annD = hlF Delim $ ifUnicode "" "::"
2023-03-26 08:40:54 -04:00
natD = hlF Syntax $ ifUnicode "" "Nat"
2022-04-23 18:21:30 -04:00
2023-03-26 08:40:54 -04:00
export %inline
2023-04-03 11:46:23 -04:00
eqD, colonD, commaD, semiD, caseD, typecaseD, returnD,
2023-04-15 09:13:01 -04:00
ofD, dotD, zeroD, succD, coeD, compD : Doc HL
2023-04-03 11:46:23 -04:00
eqD = hl Syntax "Eq"
2023-04-15 09:13:01 -04:00
colonD = hl Delim ":"
commaD = hl Delim ","
semiD = hl Delim ";"
2023-04-03 11:46:23 -04:00
caseD = hl Syntax "case"
typecaseD = hl Syntax "type-case"
ofD = hl Syntax "of"
returnD = hl Syntax "return"
dotD = hl Delim "."
zeroD = hl Syntax "zero"
succD = hl Syntax "succ"
2023-04-15 09:13:01 -04:00
coeD = hl Syntax "coe"
compD = hl Syntax "compD"
2022-04-23 18:21:30 -04:00
2023-03-05 10:48:29 -05:00
export
prettyUnivSuffix : Pretty.HasEnv m => Universe -> m (Doc HL)
prettyUnivSuffix l =
ifUnicode (pretty $ pack $ map sub $ unpack $ show l) (pretty l)
where
sub : Char -> Char
sub c = case c of
'0' => ''; '1' => ''; '2' => ''; '3' => ''; '4' => ''
'5' => ''; '6' => ''; '7' => ''; '8' => ''; '9' => ''; _ => c
2023-03-15 10:54:51 -04:00
export
prettyUniverse : Universe -> Doc HL
prettyUniverse = hl Syntax . pretty
public export
2023-04-01 13:16:43 -04:00
data WithQty a = MkWithQty Qty a
2023-03-15 10:54:51 -04:00
export
2023-04-01 13:16:43 -04:00
PrettyHL a => PrettyHL (WithQty a) where
prettyM (MkWithQty q x) = do
q <- pretty0M q
x <- withPrec Arg $ prettyM x
pure $ hcat [q, dotD, x]
public export
data Binder a = MkBinder BaseName a
export
PrettyHL a => PrettyHL (Binder a) where
prettyM (MkBinder x ty) = do
x <- pretty0M $ TV x
ty <- align <$> pretty0M ty
pure $ parens $ sep [hsep [x, colonD], ty]
2023-03-15 10:54:51 -04:00
2023-02-26 08:54:18 -05:00
export
2023-04-01 13:16:43 -04:00
prettyBindType : PrettyHL a => PrettyHL b =>
Pretty.HasEnv m =>
2023-05-01 21:06:25 -04:00
Maybe Qty -> BindName -> a -> Doc HL -> b -> m (Doc HL)
prettyBindType q (BN x _) s arr t = do
bind <- case q of
Nothing => pretty0M $ MkBinder x s
Just q => pretty0M $ MkWithQty q $ MkBinder x s
t <- withPrec AnnR $ under T x $ prettyM t
parensIfM AnnR $ hang 2 $ bind <++> arr <%%> t
2023-02-26 08:54:18 -05:00
export
prettyArm : PrettyHL a => Pretty.HasEnv m =>
2023-05-01 21:06:25 -04:00
BinderSort -> SnocList BindName -> Doc HL -> a -> m (Doc HL)
2023-02-26 08:54:18 -05:00
prettyArm sort xs pat body = do
2023-05-01 21:06:25 -04:00
let xs = map name xs
2023-02-26 08:54:18 -05:00
body <- withPrec Outer $ unders sort xs $ prettyM body
2023-03-04 15:02:51 -05:00
pure $ hang 2 $ sep [pat <++> !darrowD, body]
2023-02-22 01:40:19 -05:00
2023-02-26 08:54:18 -05:00
export
prettyLams : PrettyHL a => Pretty.HasEnv m =>
2023-05-01 21:06:25 -04:00
Maybe (Doc HL) -> BinderSort -> SnocList BindName -> a ->
m (Doc HL)
2023-02-26 08:54:18 -05:00
prettyLams lam sort names body = do
let var = case sort of T => TVar; D => DVar
header <- sequence $ [hlF var $ prettyM x | x <- toList names]
2023-02-26 08:54:18 -05:00
let header = sep $ maybe header (:: header) lam
parensIfM Outer =<< prettyArm sort names header body
2023-02-26 08:54:18 -05:00
2023-04-02 09:50:56 -04:00
public export
2023-05-01 21:06:25 -04:00
data TypeLine a = MkTypeLine BindName a
2023-04-02 09:50:56 -04:00
2023-02-26 08:54:18 -05:00
export
2023-04-02 09:50:56 -04:00
PrettyHL a => PrettyHL (TypeLine a) where
prettyM (MkTypeLine i ty) =
2023-05-01 21:06:25 -04:00
if i.name == Unused then
bracks <$> pretty0M ty
else
map bracks $ withPrec Outer $ prettyLams Nothing D [< i] ty
2023-04-02 09:50:56 -04:00
export
prettyApps' : PrettyHL f => PrettyHL a => Pretty.HasEnv m =>
f -> List (Maybe (Doc HL), a) -> m (Doc HL)
prettyApps' fun args = do
2023-03-16 13:19:17 -04:00
fun <- withPrec App $ prettyM fun
2023-04-02 09:50:56 -04:00
args <- traverse prettyArg args
2023-02-26 08:54:18 -05:00
parensIfM App $ hang 2 $ sep $ fun :: args
2023-02-26 05:22:44 -05:00
where
2023-04-02 09:50:56 -04:00
prettyArg : (Maybe (Doc HL), a) -> m (Doc HL)
prettyArg (Nothing, arg) = withPrec Arg (prettyM arg)
prettyArg (Just pfx, arg) = (hl Delim pfx <+>) <$> withPrec Arg (prettyM arg)
export
prettyApps : PrettyHL f => PrettyHL a => Pretty.HasEnv m =>
Maybe (Doc HL) -> f -> List a -> m (Doc HL)
prettyApps pfx f args = prettyApps' f (map (pfx,) args)
2023-02-22 01:40:19 -05:00
2023-02-26 08:54:18 -05:00
export
prettyTuple : PrettyHL a => Pretty.HasEnv m => List a -> m (Doc HL)
2023-02-22 01:40:19 -05:00
prettyTuple = map (parens . align . separate commaD) . traverse prettyM
2023-02-26 08:54:18 -05:00
export
prettyArms : PrettyHL a => Pretty.HasEnv m =>
2023-05-01 21:06:25 -04:00
BinderSort -> List (SnocList BindName, Doc HL, a) -> m (Doc HL)
2023-04-15 09:13:01 -04:00
prettyArms s =
2023-03-26 10:13:36 -04:00
map (braces . aseparate semiD) .
2023-04-15 09:13:01 -04:00
traverse (\(xs, l, r) => prettyArm s xs l r)
2023-02-22 01:40:19 -05:00
2023-04-03 11:46:23 -04:00
export
prettyCase' : (PrettyHL a, PrettyHL b, PrettyHL c, Pretty.HasEnv m) =>
2023-05-01 21:06:25 -04:00
Doc HL -> a -> BindName -> b ->
List (SnocList BindName, Doc HL, c) ->
2023-04-03 11:46:23 -04:00
m (Doc HL)
prettyCase' intro elim r ret arms = do
elim <- pretty0M elim
2023-05-01 21:06:25 -04:00
ret <- case r.name of
Unused => under T r.name $ pretty0M ret
2023-04-03 11:46:23 -04:00
_ => prettyLams Nothing T [< r] ret
2023-04-15 09:13:01 -04:00
arms <- prettyArms T arms
2023-04-03 11:46:23 -04:00
pure $ asep [intro <++> elim, returnD <++> ret, ofD <++> arms]
2023-02-26 08:54:18 -05:00
export
2023-04-01 13:16:43 -04:00
prettyCase : (PrettyHL a, PrettyHL b, PrettyHL c, Pretty.HasEnv m) =>
2023-05-01 21:06:25 -04:00
Qty -> a -> BindName -> b ->
List (SnocList BindName, Doc HL, c) ->
2023-02-22 01:40:19 -05:00
m (Doc HL)
2023-02-26 08:54:18 -05:00
prettyCase pi elim r ret arms = do
caseq <- (caseD <+>) <$> prettySuffix pi
2023-04-03 11:46:23 -04:00
prettyCase' caseq elim r ret arms
2023-02-22 01:40:19 -05:00
export
escapeString : String -> String
escapeString = concatMap esc1 . unpack where
esc1 : Char -> String
esc1 '"' = #"\""#
esc1 '\\' = #"\\"#
esc1 '\n' = #"\n"#
esc1 c = singleton c
2023-03-16 13:39:24 -04:00
export
quoteTag : TagVal -> Doc HL
quoteTag tag =
if isName tag then fromString tag else
hcat ["\"", fromString $ escapeString tag, "\""]
2023-03-16 13:39:24 -04:00
export
prettyTag : TagVal -> Doc HL
2023-03-16 13:39:24 -04:00
prettyTag t = hl Tag $ "'" <+> quoteTag t
2023-03-16 13:19:17 -04:00
export
prettyTagBare : TagVal -> Doc HL
2023-03-16 13:39:24 -04:00
prettyTagBare t = hl Tag $ quoteTag t
2023-03-16 13:19:17 -04:00
2023-03-31 13:11:35 -04:00
export
prettyBoxVal : PrettyHL a => Pretty.HasEnv m => a -> m (Doc HL)
prettyBoxVal val = bracks <$> pretty0M val
2023-04-15 09:13:01 -04:00
export
2023-05-01 21:06:25 -04:00
prettyCompPat : Pretty.HasEnv m => DimConst -> BindName -> m (Doc HL)
prettyCompPat e j = hsep <$> sequence [pretty0M e, pretty0M $ DV j.name]
2023-04-15 09:13:01 -04:00
2023-03-26 08:40:54 -04:00
export
2023-04-01 13:16:43 -04:00
toNatLit : Term d n -> Maybe Nat
2023-05-01 21:06:25 -04:00
toNatLit (Zero _) = Just 0
toNatLit (Succ n _) = [|S $ toNatLit n|]
toNatLit _ = Nothing
2023-03-26 08:40:54 -04:00
private
2023-04-01 13:16:43 -04:00
eterm : Term d n -> Exists (Term d)
2023-03-26 08:40:54 -04:00
eterm = Evidence n
2023-02-22 01:40:19 -05:00
2023-02-26 05:25:11 -05:00
parameters (showSubsts : Bool)
mutual
2022-04-23 18:21:30 -04:00
export covering
2023-04-01 13:16:43 -04:00
[TermSubst] PrettyHL (Term d n) using ElimSubst
2023-02-26 05:25:11 -05:00
where
2023-05-01 21:06:25 -04:00
prettyM (TYPE l _) =
2023-04-17 15:45:05 -04:00
pure $ !typeD <+> hl Syntax !(prettyUnivSuffix l)
2023-04-15 09:13:01 -04:00
2023-05-01 21:06:25 -04:00
prettyM (Pi qty s (S _ (N t)) _) = do
dom <- pretty0M $ MkWithQty qty s
cod <- withPrec AnnR $ prettyM t
parensIfM AnnR $ asep [dom <++> !arrowD, cod]
2023-04-15 09:13:01 -04:00
2023-05-01 21:06:25 -04:00
prettyM (Pi qty s (S [< x] (Y t)) _) =
prettyBindType (Just qty) x s !arrowD t
2023-04-15 09:13:01 -04:00
2023-05-01 21:06:25 -04:00
prettyM (Lam (S x t) _) =
2023-02-22 01:40:19 -05:00
let GotLams {names, body, _} = getLams' x t.term Refl in
prettyLams (Just !lamD) T (toSnocList' names) body
2023-04-15 09:13:01 -04:00
2023-05-01 21:06:25 -04:00
prettyM (Sig s (S _ (N t)) _) = do
s <- withPrec InTimes $ prettyM s
t <- withPrec Times $ prettyM t
parensIfM Times $ asep [s <++> !timesD, t]
2023-04-15 09:13:01 -04:00
2023-05-01 21:06:25 -04:00
prettyM (Sig s (S [< x] (Y t)) _) =
2023-04-01 13:16:43 -04:00
prettyBindType Nothing x s !timesD t
2023-04-15 09:13:01 -04:00
2023-05-01 21:06:25 -04:00
prettyM (Pair s t _) =
let GotPairs {init, last, _} = getPairs' [< s] t in
prettyTuple $ toList $ init :< last
2023-04-15 09:13:01 -04:00
2023-05-01 21:06:25 -04:00
prettyM (Enum tags _) =
2023-03-16 13:19:17 -04:00
pure $ delims "{" "}" . aseparate comma $ map prettyTagBare $
2023-02-26 05:23:30 -05:00
Prelude.toList tags
2023-04-15 09:13:01 -04:00
2023-05-01 21:06:25 -04:00
prettyM (Tag t _) =
pure $ prettyTag t
2023-04-15 09:13:01 -04:00
2023-05-01 21:06:25 -04:00
prettyM (Eq (S _ (N ty)) l r _) = do
2023-02-26 08:54:18 -05:00
l <- withPrec InEq $ prettyM l
r <- withPrec InEq $ prettyM r
ty <- withPrec InEq $ prettyM ty
parensIfM Eq $ asep [l <++> !eqndD, r <++> colonD, ty]
2023-04-15 09:13:01 -04:00
2023-05-01 21:06:25 -04:00
prettyM (Eq (S [< i] (Y ty)) l r _) = do
2023-04-02 09:50:56 -04:00
prettyApps Nothing (L eqD)
[epretty $ MkTypeLine i ty, epretty l, epretty r]
2023-04-15 09:13:01 -04:00
2023-05-01 21:06:25 -04:00
prettyM (DLam (S i t) _) =
2023-02-22 01:40:19 -05:00
let GotDLams {names, body, _} = getDLams' i t.term Refl in
prettyLams (Just !dlamD) D (toSnocList' names) body
2023-04-15 09:13:01 -04:00
2023-05-01 21:06:25 -04:00
prettyM (Nat _) = natD
2023-04-15 09:13:01 -04:00
2023-05-01 21:06:25 -04:00
prettyM (Zero _) = pure $ hl Syntax "0"
2023-04-15 09:13:01 -04:00
2023-05-01 21:06:25 -04:00
prettyM (Succ n _) =
2023-03-26 08:40:54 -04:00
case toNatLit n of
Just n => pure $ hl Syntax $ pretty $ S n
2023-04-02 09:50:56 -04:00
Nothing => prettyApps Nothing (L succD) [n]
2023-04-15 09:13:01 -04:00
2023-05-01 21:06:25 -04:00
prettyM (BOX pi ty _) = do
2023-03-31 13:11:35 -04:00
pi <- pretty0M pi
ty <- pretty0M ty
pure $ bracks $ hcat [pi, dotD, align ty]
2023-04-15 09:13:01 -04:00
2023-05-01 21:06:25 -04:00
prettyM (Box val _) = prettyBoxVal val
2023-04-15 09:13:01 -04:00
2023-03-15 10:54:51 -04:00
prettyM (E e) = prettyM e
2023-04-15 09:13:01 -04:00
prettyM (CloT (Sub s th)) =
2023-02-26 05:25:11 -05:00
if showSubsts then
parensIfM SApp . hang 2 =<<
[|withPrec SApp (prettyM s) <%> prettyTSubst th|]
2023-02-26 05:25:11 -05:00
else
prettyM $ pushSubstsWith' id th s
2023-04-15 09:13:01 -04:00
prettyM (DCloT (Sub s th)) =
2023-02-26 05:25:11 -05:00
if showSubsts then
parensIfM SApp . hang 2 =<<
[|withPrec SApp (prettyM s) <%> prettyDSubst th|]
2023-02-26 05:25:11 -05:00
else
prettyM $ pushSubstsWith' th id s
2022-04-23 18:21:30 -04:00
export covering
2023-04-01 13:16:43 -04:00
[ElimSubst] PrettyHL (Elim d n) using TermSubst
2023-02-26 05:25:11 -05:00
where
2023-05-01 21:06:25 -04:00
prettyM (F x _) =
2022-04-23 18:21:30 -04:00
hl' Free <$> prettyM x
2023-04-15 09:13:01 -04:00
2023-05-01 21:06:25 -04:00
prettyM (B i _) =
2022-04-23 18:21:30 -04:00
prettyVar TVar TVarErr (!ask).tnames i
2023-04-15 09:13:01 -04:00
2023-05-01 21:06:25 -04:00
prettyM (App e s _) =
2023-01-20 20:34:28 -05:00
let GotArgs {fun, args, _} = getArgs' e [s] in
2023-02-26 05:22:44 -05:00
prettyApps Nothing fun args
2023-04-15 09:13:01 -04:00
2023-05-01 21:06:25 -04:00
prettyM (CasePair pi p (S [< r] ret) (S [< x, y] body) _) = do
2023-02-26 08:54:18 -05:00
pat <- parens . separate commaD <$> traverse (hlF TVar . prettyM) [x, y]
prettyCase pi p r ret.term [([< x, y], pat, body.term)]
2023-04-15 09:13:01 -04:00
2023-05-01 21:06:25 -04:00
prettyM (CaseEnum pi t (S [< r] ret) arms _) =
2023-02-26 05:25:11 -05:00
prettyCase pi t r ret.term
[([<], prettyTag t, b) | (t, b) <- SortedMap.toList arms]
2023-04-15 09:13:01 -04:00
2023-05-01 21:06:25 -04:00
prettyM (CaseNat pi pi' nat (S [< r] ret) zer (S [< s, ih] suc) _) =
2023-03-26 08:40:54 -04:00
prettyCase pi nat r ret.term
[([<], zeroD, eterm zer),
([< s, ih], !succPat, eterm suc.term)]
2023-04-15 09:13:01 -04:00
where
succPat : m (Doc HL)
succPat = case (ih, pi') of
2023-05-01 21:06:25 -04:00
(BN Unused _, Zero) => pure $ succD <++> !(pretty0M s)
2023-04-15 09:13:01 -04:00
_ => pure $ asep [succD <++> !(pretty0M s) <+> comma,
!(pretty0M $ MkWithQty pi' ih)]
2023-05-01 21:06:25 -04:00
prettyM (CaseBox pi box (S [< r] ret) (S [< u] body) _) =
2023-03-31 13:11:35 -04:00
prettyCase pi box r ret.term
2023-05-01 21:06:25 -04:00
[([< u], !(prettyBoxVal $ TV u.name), body.term)]
2023-04-15 09:13:01 -04:00
2023-05-01 21:06:25 -04:00
prettyM (DApp e d _) =
2023-01-20 20:34:28 -05:00
let GotDArgs {fun, args, _} = getDArgs' e [d] in
2023-02-26 05:22:44 -05:00
prettyApps (Just "@") fun args
2023-04-15 09:13:01 -04:00
2023-05-01 21:06:25 -04:00
prettyM (Ann s a _) = do
2023-02-26 08:54:18 -05:00
s <- withPrec AnnL $ prettyM s
a <- withPrec AnnR $ prettyM a
parensIfM AnnR $ hang 2 $ s <++> !annD <%%> a
2023-04-15 09:13:01 -04:00
2023-05-01 21:06:25 -04:00
prettyM (Coe (S [< i] ty) p q val _) =
2023-04-15 09:13:01 -04:00
let ty = case ty of
Y ty => epretty $ MkTypeLine i ty
N ty => epretty ty
in
prettyApps' (L coeD)
[(Nothing, ty),
(Just "@", epretty p),
(Just "@", epretty q),
(Nothing, epretty val)]
2023-05-01 21:06:25 -04:00
prettyM (Comp ty p q val r (S [< z] zero) (S [< o] one) _) = do
2023-04-15 09:13:01 -04:00
apps <- prettyApps' (L compD)
2023-05-01 21:06:25 -04:00
[(Nothing, epretty $ MkTypeLine (BN Unused noLoc) ty),
2023-04-15 09:13:01 -04:00
(Just "@", epretty p),
(Just "@", epretty q),
2023-04-17 15:45:05 -04:00
(Nothing, epretty val),
(Just "@", epretty r)]
2023-04-15 09:13:01 -04:00
arms <- prettyArms D
2023-04-17 15:45:05 -04:00
[([< z], !(prettyCompPat Zero z), zero.term),
([< o], !(prettyCompPat One o), one.term)]
2023-04-15 09:13:01 -04:00
pure $ apps <++> arms
2023-05-01 21:06:25 -04:00
prettyM (TypeCase ty ret arms def _) = do
2023-04-15 09:13:01 -04:00
arms <- traverse fromArm (toList arms)
2023-05-01 21:06:25 -04:00
prettyCase' typecaseD ty (BN Unused noLoc) ret $
2023-04-15 09:13:01 -04:00
arms ++ [([<], hl Syntax "_", eterm def)]
where
2023-05-01 21:06:25 -04:00
v : BindName -> Doc HL
v = pretty0 True . TV . name
2023-04-15 09:13:01 -04:00
2023-05-01 21:06:25 -04:00
tyCasePat : (k : TyConKind) -> BContext (arity k) -> m (Doc HL)
2023-04-15 09:13:01 -04:00
tyCasePat KTYPE [<] = typeD
tyCasePat KPi [< a, b] = pure $ parens $ hsep [v a, !arrowD, v b]
tyCasePat KSig [< a, b] = pure $ parens $ hsep [v a, !arrowD, v b]
tyCasePat KEnum [<] = pure $ hl Syntax "{}"
tyCasePat KNat [<] = natD
tyCasePat KBOX [< a] = pure $ bracks $ v a
2023-05-01 21:06:25 -04:00
tyCasePat KEq vars =
prettyApps Nothing (L eqD) $ map (TV . name) $ toList' vars
2023-04-15 09:13:01 -04:00
fromArm : TypeCaseArm d n ->
2023-05-01 21:06:25 -04:00
m (SnocList BindName, Doc HL, Exists (Term d))
2023-04-15 09:13:01 -04:00
fromArm (k ** S ns t) =
pure (toSnocList' ns, !(tyCasePat k ns), eterm t.term)
2023-05-01 21:06:25 -04:00
prettyM (CloE (Sub e th)) =
2023-02-26 05:25:11 -05:00
if showSubsts then
parensIfM SApp . hang 2 =<<
[|withPrec SApp (prettyM e) <%> prettyTSubst th|]
2023-02-26 05:25:11 -05:00
else
prettyM $ pushSubstsWith' id th e
2023-04-15 09:13:01 -04:00
prettyM (DCloE (Sub e th)) =
2023-02-26 05:25:11 -05:00
if showSubsts then
parensIfM SApp . hang 2 =<<
[|withPrec SApp (prettyM e) <%> prettyDSubst th|]
2023-02-26 05:25:11 -05:00
else
prettyM $ pushSubstsWith' th id e
2022-04-23 18:21:30 -04:00
export covering
2023-04-01 13:16:43 -04:00
prettyTSubst : Pretty.HasEnv m => TSubst d from to -> m (Doc HL)
2023-02-26 05:25:11 -05:00
prettyTSubst s =
prettySubstM (prettyM @{ElimSubst}) (!ask).tnames TVar "[" "]" s
export covering %inline
2023-04-01 13:16:43 -04:00
PrettyHL (Term d n) where prettyM = prettyM @{TermSubst False}
2023-02-26 05:25:11 -05:00
export covering %inline
2023-04-01 13:16:43 -04:00
PrettyHL (Elim d n) where prettyM = prettyM @{ElimSubst False}
2023-03-15 10:54:51 -04:00
export covering
2023-04-01 13:16:43 -04:00
prettyTerm : (unicode : Bool) ->
2023-05-01 21:06:25 -04:00
(dnames : BContext d) -> (tnames : BContext n) ->
2023-04-01 13:16:43 -04:00
Term d n -> Doc HL
2023-03-15 10:54:51 -04:00
prettyTerm unicode dnames tnames term =
2023-05-01 21:06:25 -04:00
pretty0With unicode (toNames dnames) (toNames tnames) term