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
|
|
|
|
|
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-01-20 20:34:28 -05:00
|
|
|
|
private %inline
|
2023-03-04 15:02:51 -05:00
|
|
|
|
typeD, arrowD, darrowD, timesD, lamD, eqndD, dlamD, annD :
|
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-02-25 13:14:11 -05:00
|
|
|
|
arrowD = hlF Syntax $ ifUnicode "→" "->"
|
2023-03-04 15:02:51 -05:00
|
|
|
|
darrowD = hlF Syntax $ ifUnicode "⇒" "=>"
|
2023-02-25 13:14:11 -05:00
|
|
|
|
timesD = hlF Syntax $ ifUnicode "×" "**"
|
|
|
|
|
lamD = hlF Syntax $ ifUnicode "λ" "fun"
|
|
|
|
|
eqndD = hlF Syntax $ ifUnicode "≡" "=="
|
|
|
|
|
dlamD = hlF Syntax $ ifUnicode "δ" "dfun"
|
|
|
|
|
annD = hlF Syntax $ ifUnicode "∷" "::"
|
2022-04-23 18:21:30 -04:00
|
|
|
|
|
2023-01-20 20:34:28 -05:00
|
|
|
|
private %inline
|
2023-03-04 15:02:51 -05:00
|
|
|
|
eqD, colonD, commaD, caseD, returnD, ofD : Doc HL
|
2023-01-26 13:54:46 -05:00
|
|
|
|
eqD = hl Syntax "Eq"
|
|
|
|
|
colonD = hl Syntax ":"
|
|
|
|
|
commaD = hl Syntax ","
|
|
|
|
|
caseD = hl Syntax "case"
|
|
|
|
|
ofD = hl Syntax "of"
|
|
|
|
|
returnD = hl Syntax "return"
|
2022-04-23 18:21:30 -04:00
|
|
|
|
|
2023-02-13 16:05:27 -05: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-02-26 08:54:18 -05:00
|
|
|
|
export
|
2023-02-26 05:21:25 -05:00
|
|
|
|
prettyBindType : PrettyHL a => PrettyHL b => PrettyHL q =>
|
|
|
|
|
Pretty.HasEnv m =>
|
2023-02-22 01:40:19 -05:00
|
|
|
|
List q -> BaseName -> a -> Doc HL -> b -> m (Doc HL)
|
2023-02-26 08:54:18 -05:00
|
|
|
|
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
|
|
|
|
|
|
|
|
|
|
export
|
|
|
|
|
prettyArm : PrettyHL a => Pretty.HasEnv m =>
|
|
|
|
|
BinderSort -> List BaseName -> Doc HL -> a -> m (Doc HL)
|
|
|
|
|
prettyArm sort xs pat body = do
|
|
|
|
|
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
|
2023-02-26 05:21:25 -05:00
|
|
|
|
prettyLams : PrettyHL a => Pretty.HasEnv m =>
|
2023-02-26 08:54:18 -05:00
|
|
|
|
Maybe (Doc HL) -> BinderSort -> List BaseName -> a -> m (Doc HL)
|
|
|
|
|
prettyLams lam sort names body = do
|
|
|
|
|
let var = case sort of T => TVar; D => DVar
|
|
|
|
|
header <- sequence $ [hlF var $ prettyM x | x <- names]
|
|
|
|
|
let header = sep $ maybe header (:: header) lam
|
|
|
|
|
parensIfM Outer =<< prettyArm sort names header body
|
|
|
|
|
|
|
|
|
|
export
|
2023-02-26 05:21:25 -05:00
|
|
|
|
prettyApps : PrettyHL f => PrettyHL a => Pretty.HasEnv m =>
|
2023-02-26 05:22:44 -05:00
|
|
|
|
Maybe (Doc HL) -> f -> List a -> m (Doc HL)
|
2023-02-26 08:54:18 -05:00
|
|
|
|
prettyApps pfx fun args = do
|
|
|
|
|
fun <- withPrec Arg $ prettyM fun
|
|
|
|
|
args <- traverse (withPrec Arg . prettyArg) args
|
|
|
|
|
parensIfM App $ hang 2 $ sep $ fun :: args
|
2023-02-26 05:22:44 -05:00
|
|
|
|
where
|
|
|
|
|
prettyArg : a -> m (Doc HL)
|
2023-02-26 08:54:18 -05:00
|
|
|
|
prettyArg = case pfx of
|
|
|
|
|
Nothing => prettyM
|
|
|
|
|
Just pfx => \x => pure $ hl Delim pfx <+> !(prettyM x)
|
2023-02-22 01:40:19 -05:00
|
|
|
|
|
2023-02-26 08:54:18 -05:00
|
|
|
|
export
|
2023-02-26 05:21:25 -05:00
|
|
|
|
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
|
2023-02-26 05:21:25 -05:00
|
|
|
|
prettyArms : PrettyHL a => Pretty.HasEnv m =>
|
2023-02-22 01:40:19 -05:00
|
|
|
|
List (List BaseName, Doc HL, a) -> m (Doc HL)
|
2023-02-26 08:54:18 -05:00
|
|
|
|
prettyArms =
|
|
|
|
|
map (braces . asep) . traverse (\(xs, l, r) => prettyArm T xs l r)
|
2023-02-22 01:40:19 -05:00
|
|
|
|
|
2023-02-26 08:54:18 -05:00
|
|
|
|
export
|
2023-02-26 05:21:25 -05:00
|
|
|
|
prettyCase : PrettyHL a => PrettyHL b => PrettyHL c => PrettyHL q =>
|
|
|
|
|
Pretty.HasEnv m =>
|
2023-02-22 01:40:19 -05:00
|
|
|
|
q -> a -> BaseName -> b -> List (List BaseName, Doc HL, c) ->
|
|
|
|
|
m (Doc HL)
|
2023-02-26 08:54:18 -05:00
|
|
|
|
prettyCase pi elim r ret arms = do
|
|
|
|
|
elim <- prettyQtyBinds [pi] elim
|
|
|
|
|
ret <- prettyLams Nothing T [r] ret
|
|
|
|
|
arms <- prettyArms arms
|
|
|
|
|
pure $ asep [caseD <++> elim, returnD <++> ret, ofD <++> arms]
|
2023-02-22 01:40:19 -05:00
|
|
|
|
|
2023-02-22 01:45:10 -05:00
|
|
|
|
-- [fixme] put delimiters around tags that aren't simple names
|
|
|
|
|
export
|
|
|
|
|
prettyTag : TagVal -> Doc HL
|
2023-03-04 15:02:51 -05:00
|
|
|
|
prettyTag t = hl Tag $ "'" <+> fromString t
|
2023-02-22 01:45:10 -05:00
|
|
|
|
|
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-02-26 05:25:11 -05:00
|
|
|
|
[TermSubst] PrettyHL q => PrettyHL (Term q d n) using TermSubst ElimSubst
|
|
|
|
|
where
|
2022-04-23 18:21:30 -04:00
|
|
|
|
prettyM (TYPE l) =
|
2023-02-26 05:20:06 -05:00
|
|
|
|
parensIfM App $ !typeD <+> hl Syntax !(prettyUnivSuffix l)
|
2023-02-22 01:40:19 -05:00
|
|
|
|
prettyM (Pi qty s (S [x] t)) =
|
2023-02-26 05:25:11 -05:00
|
|
|
|
prettyBindType [qty] x s !arrowD t.term
|
2023-02-22 01:40:19 -05:00
|
|
|
|
prettyM (Lam (S x t)) =
|
|
|
|
|
let GotLams {names, body, _} = getLams' x t.term Refl in
|
2023-02-26 08:54:18 -05:00
|
|
|
|
prettyLams (Just !lamD) T (toList names) body
|
2023-02-22 01:40:19 -05:00
|
|
|
|
prettyM (Sig s (S [x] t)) =
|
2023-02-26 05:25:11 -05:00
|
|
|
|
prettyBindType {q} [] x s !timesD t.term
|
2023-01-26 13:54:46 -05:00
|
|
|
|
prettyM (Pair s t) =
|
2023-02-26 05:24:28 -05:00
|
|
|
|
let GotPairs {init, last, _} = getPairs' [< s] t in
|
|
|
|
|
prettyTuple $ toList $ init :< last
|
2023-02-22 01:45:10 -05:00
|
|
|
|
prettyM (Enum tags) =
|
2023-03-04 15:02:51 -05:00
|
|
|
|
pure $ delims "{" "}" . aseparate comma $ map prettyTag $
|
2023-02-26 05:23:30 -05:00
|
|
|
|
Prelude.toList tags
|
2023-02-22 01:45:10 -05:00
|
|
|
|
prettyM (Tag t) =
|
|
|
|
|
pure $ prettyTag t
|
2023-02-26 08:54:18 -05:00
|
|
|
|
prettyM (Eq (S _ (N ty)) l r) = do
|
|
|
|
|
l <- withPrec InEq $ prettyM l
|
|
|
|
|
r <- withPrec InEq $ prettyM r
|
|
|
|
|
ty <- withPrec InEq $ prettyM ty
|
|
|
|
|
parensIfM Eq $ asep [l <++> !eqndD, r <++> colonD, ty]
|
|
|
|
|
prettyM (Eq (S [i] (Y ty)) l r) = do
|
|
|
|
|
ty <- bracks <$> withPrec Outer (prettyLams Nothing D [i] ty)
|
|
|
|
|
l <- withPrec Arg $ prettyM l
|
|
|
|
|
r <- withPrec Arg $ prettyM r
|
|
|
|
|
parensIfM App $ eqD <++> asep [ty, l, r]
|
2023-02-22 01:40:19 -05:00
|
|
|
|
prettyM (DLam (S i t)) =
|
|
|
|
|
let GotDLams {names, body, _} = getDLams' i t.term Refl in
|
2023-02-26 08:54:18 -05:00
|
|
|
|
prettyLams (Just !dlamD) D (toList names) body
|
2023-01-20 20:34:28 -05:00
|
|
|
|
prettyM (E e) = bracks <$> prettyM e
|
2022-04-23 18:21:30 -04:00
|
|
|
|
prettyM (CloT s th) =
|
2023-02-26 05:25:11 -05:00
|
|
|
|
if showSubsts then
|
|
|
|
|
parensIfM SApp . hang 2 =<<
|
|
|
|
|
[|withPrec SApp (prettyM s) </> prettyTSubst th|]
|
|
|
|
|
else
|
|
|
|
|
prettyM $ pushSubstsWith' id th s
|
2022-04-23 18:21:30 -04:00
|
|
|
|
prettyM (DCloT s th) =
|
2023-02-26 05:25:11 -05:00
|
|
|
|
if showSubsts then
|
|
|
|
|
parensIfM SApp . hang 2 =<<
|
|
|
|
|
[|withPrec SApp (prettyM s) </> prettyDSubst th|]
|
|
|
|
|
else
|
|
|
|
|
prettyM $ pushSubstsWith' th id s
|
2022-04-23 18:21:30 -04:00
|
|
|
|
|
|
|
|
|
export covering
|
2023-02-26 05:25:11 -05:00
|
|
|
|
[ElimSubst] PrettyHL q => PrettyHL (Elim q d n) using TermSubst ElimSubst
|
|
|
|
|
where
|
2022-04-23 18:21:30 -04:00
|
|
|
|
prettyM (F x) =
|
|
|
|
|
hl' Free <$> prettyM x
|
|
|
|
|
prettyM (B i) =
|
|
|
|
|
prettyVar TVar TVarErr (!ask).tnames i
|
|
|
|
|
prettyM (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-02-22 01:40:19 -05: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]
|
2023-02-26 05:25:11 -05:00
|
|
|
|
prettyCase pi p r ret.term [([x, y], pat, body.term)]
|
2023-02-22 01:45:10 -05:00
|
|
|
|
prettyM (CaseEnum pi t (S [r] ret) arms) =
|
2023-02-26 05:25:11 -05:00
|
|
|
|
prettyCase pi t r ret.term
|
2023-02-22 01:45:10 -05:00
|
|
|
|
[([], prettyTag t, b) | (t, b) <- SortedMap.toList arms]
|
2023-01-20 20:34:28 -05:00
|
|
|
|
prettyM (e :% d) =
|
|
|
|
|
let GotDArgs {fun, args, _} = getDArgs' e [d] in
|
2023-02-26 05:22:44 -05:00
|
|
|
|
prettyApps (Just "@") fun args
|
2023-02-26 08:54:18 -05:00
|
|
|
|
prettyM (s :# a) = do
|
|
|
|
|
s <- withPrec AnnL $ prettyM s
|
|
|
|
|
a <- withPrec Ann $ prettyM a
|
|
|
|
|
parensIfM Ann $ hang 2 $ s <++> !annD <//> a
|
2022-04-23 18:21:30 -04:00
|
|
|
|
prettyM (CloE e th) =
|
2023-02-26 05:25:11 -05:00
|
|
|
|
if showSubsts then
|
|
|
|
|
parensIfM SApp . hang 2 =<<
|
|
|
|
|
[|withPrec SApp (prettyM e) </> prettyTSubst th|]
|
|
|
|
|
else
|
|
|
|
|
prettyM $ pushSubstsWith' id th e
|
2022-04-23 18:21:30 -04:00
|
|
|
|
prettyM (DCloE e th) =
|
2023-02-26 05:25:11 -05:00
|
|
|
|
if showSubsts then
|
|
|
|
|
parensIfM SApp . hang 2 =<<
|
|
|
|
|
[|withPrec SApp (prettyM e) </> prettyDSubst th|]
|
|
|
|
|
else
|
|
|
|
|
prettyM $ pushSubstsWith' th id e
|
2022-04-23 18:21:30 -04:00
|
|
|
|
|
|
|
|
|
export covering
|
2023-01-08 14:44:25 -05:00
|
|
|
|
prettyTSubst : Pretty.HasEnv m => PrettyHL q =>
|
|
|
|
|
TSubst q 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
|
|
|
|
|
PrettyHL q => PrettyHL (Term q d n) where
|
|
|
|
|
prettyM = prettyM @{TermSubst False}
|
|
|
|
|
|
|
|
|
|
export covering %inline
|
|
|
|
|
PrettyHL q => PrettyHL (Elim q d n) where
|
|
|
|
|
prettyM = prettyM @{ElimSubst False}
|