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

357 lines
12 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-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 "" "::"
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,
ofD, dotD, zeroD, succD : Doc HL
eqD = hl Syntax "Eq"
colonD = hl Syntax ":"
commaD = hl Syntax ","
semiD = hl Syntax ";"
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"
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-04-01 13:16:43 -04:00
Maybe Qty -> BaseName -> a -> Doc HL -> b -> m (Doc HL)
prettyBindType q 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 =>
BinderSort -> SnocList BaseName -> Doc HL -> a -> m (Doc HL)
2023-02-26 08:54:18 -05:00
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
prettyLams : PrettyHL a => Pretty.HasEnv m =>
Maybe (Doc HL) -> BinderSort -> SnocList BaseName -> 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
data TypeLine a = MkTypeLine BaseName a
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) =
map bracks $ withPrec Outer $ prettyLams Nothing D [< i] ty
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 =>
List (SnocList BaseName, Doc HL, a) -> m (Doc HL)
2023-02-26 08:54:18 -05:00
prettyArms =
2023-03-26 10:13:36 -04:00
map (braces . aseparate semiD) .
traverse (\(xs, l, r) => prettyArm T 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) =>
Doc HL -> a -> BaseName -> b ->
List (SnocList BaseName, Doc HL, c) ->
m (Doc HL)
prettyCase' intro elim r ret arms = do
elim <- pretty0M elim
ret <- case r of
Unused => pretty0M ret
_ => prettyLams Nothing T [< r] ret
arms <- prettyArms arms
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) =>
Qty -> a -> BaseName -> b ->
List (SnocList BaseName, 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-03-26 08:40:54 -04:00
export
2023-04-01 13:16:43 -04:00
toNatLit : Term d n -> Maybe Nat
2023-03-26 08:40:54 -04:00
toNatLit Zero = Just 0
toNatLit (Succ n) = [|S $ toNatLit n|]
toNatLit _ = Nothing
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
2022-04-23 18:21:30 -04:00
prettyM (TYPE l) =
parensIfM App $ !typeD <+> hl Syntax !(prettyUnivSuffix l)
prettyM (Pi qty s (S _ (N t))) = do
dom <- pretty0M $ MkWithQty qty s
cod <- withPrec AnnR $ prettyM t
parensIfM AnnR $ asep [dom <++> !arrowD, cod]
prettyM (Pi qty s (S [< x] (Y t))) =
prettyBindType (Just qty) x s !arrowD t
2023-02-22 01:40:19 -05:00
prettyM (Lam (S x t)) =
let GotLams {names, body, _} = getLams' x t.term Refl in
prettyLams (Just !lamD) T (toSnocList' names) body
prettyM (Sig s (S _ (N t))) = do
s <- withPrec InTimes $ prettyM s
t <- withPrec Times $ prettyM t
parensIfM Times $ asep [s <++> !timesD, t]
prettyM (Sig s (S [< x] (Y t))) =
2023-04-01 13:16:43 -04:00
prettyBindType Nothing x s !timesD t
2023-01-26 13:54:46 -05:00
prettyM (Pair s t) =
let GotPairs {init, last, _} = getPairs' [< s] t in
prettyTuple $ toList $ init :< last
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
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
2023-04-02 09:50:56 -04:00
prettyApps Nothing (L eqD)
[epretty $ MkTypeLine i ty, epretty l, epretty r]
2023-02-22 01:40:19 -05:00
prettyM (DLam (S i t)) =
let GotDLams {names, body, _} = getDLams' i t.term Refl in
prettyLams (Just !dlamD) D (toSnocList' names) body
2023-03-26 08:40:54 -04:00
prettyM Nat = natD
prettyM Zero = pure $ hl Syntax "0"
prettyM (Succ n) =
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-03-31 13:11:35 -04:00
prettyM (BOX pi ty) = do
pi <- pretty0M pi
ty <- pretty0M ty
pure $ bracks $ hcat [pi, dotD, align ty]
prettyM (Box val) = prettyBoxVal val
2023-03-15 10:54:51 -04:00
prettyM (E e) = 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|]
2023-02-26 05:25:11 -05:00
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|]
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
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
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)]
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-03-26 08:40:54 -04:00
prettyM (CaseNat pi pi' nat (S [< r] ret) zer (S [< s, ih] suc)) =
prettyCase pi nat r ret.term
[([<], zeroD, eterm zer),
([< s, ih], !succPat, eterm suc.term)]
where
succPat : m (Doc HL)
2023-04-02 09:50:56 -04:00
succPat = case (ih, pi') of
(Unused, Zero) => pure $ succD <++> !(pretty0M s)
_ => pure $ asep [succD <++> !(pretty0M s) <+> comma,
!(pretty0M $ MkWithQty pi' ih)]
2023-03-31 13:11:35 -04:00
prettyM (CaseBox pi box (S [< r] ret) (S [< u] body)) =
prettyCase pi box r ret.term
[([< u], !(prettyBoxVal $ TV u), body.term)]
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 AnnR $ prettyM a
parensIfM AnnR $ hang 2 $ s <++> !annD <%%> a
2023-04-03 11:46:23 -04:00
prettyM (TypeCase ty ret univ
(S [< piA, piB] pi) (S [< sigA, sigB] sig) enum
(S [< eqA0, eqA1, eqA, eqL, eqR] eq)
nat (S [< boxA] box)) = do
let v : BaseName -> Doc HL := pretty0 True . TV
pipat <- pure $ parens $ hsep [v piA, !arrowD, v piB]
sigpat <- pure $ parens $ hsep [v sigA, !timesD, v sigB]
eqpat <- prettyApps Nothing (L eqD)
[TV eqA0, TV eqA1, TV eqA, TV eqL, TV eqR]
boxpat <- pure $ bracks $ v boxA
prettyCase' typecaseD ty Unused ret
[([<], !typeD, eterm univ),
([< piA, piB], pipat, eterm pi.term),
([< sigA, sigB], sigpat, eterm sig.term),
([<], hl Syntax "{}", eterm enum),
([< eqA0, eqA1, eqA, eqL, eqR], eqpat, eterm eq.term),
([<], !natD, eterm nat),
([< boxA], boxpat, eterm box.term)]
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|]
2023-02-26 05:25:11 -05:00
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|]
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-03-15 10:54:51 -04:00
(dnames : NContext d) -> (tnames : NContext 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-03-25 15:48:26 -04:00
pretty0With unicode (toSnocList' dnames) (toSnocList' tnames) term