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

651 lines
21 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.Subst
2023-03-15 10:54:51 -04:00
import Quox.Context
2022-04-23 18:21:30 -04:00
import Quox.Pretty
2024-05-27 15:28:22 -04:00
import Quox.SingletonExtra
2022-04-23 18:21:30 -04:00
import Data.Vect
2023-05-14 13:58:46 -04:00
import Derive.Prelude
2022-04-23 18:21:30 -04:00
2022-05-02 16:38:37 -04:00
%default total
2023-05-14 13:58:46 -04:00
%language ElabReflection
2022-04-23 18:21:30 -04:00
2023-05-14 13:58:46 -04:00
export
2024-05-27 15:28:22 -04:00
prettyUniverse : {opts : LayoutOpts} -> Universe -> Eff Pretty (Doc opts)
2023-05-21 14:09:34 -04:00
prettyUniverse = hl Universe . text . show
2023-05-14 13:58:46 -04:00
2022-04-23 18:21:30 -04:00
2023-05-14 13:58:46 -04:00
export
2024-05-27 15:28:22 -04:00
prettyTerm : {opts : LayoutOpts} ->
NameContexts q d n -> Term q d n -> Eff Pretty (Doc opts)
2023-03-05 10:48:29 -05:00
export
2024-05-27 15:28:22 -04:00
prettyElim : {opts : LayoutOpts} ->
NameContexts q d n -> Elim q d n -> Eff Pretty (Doc opts)
2023-05-14 13:58:46 -04:00
2024-05-27 15:28:22 -04:00
export
prettyQty : {opts : LayoutOpts} ->
NameContexts q d n -> Qty q -> Eff Pretty (Doc opts)
prettyQty names pi = let Val q = names.qtyLen in prettyQty names.qnames pi
2023-05-14 13:58:46 -04:00
2023-05-21 14:09:34 -04:00
private
superscript : String -> String
superscript = pack . map sup . unpack where
sup : Char -> Char
sup c = case c of
'0' => ''; '1' => '¹'; '2' => '²'; '3' => '³'; '4' => ''
'5' => ''; '6' => ''; '7' => ''; '8' => ''; '9' => ''; _ => c
2023-03-15 10:54:51 -04:00
2023-05-14 13:58:46 -04:00
private
2024-05-27 15:28:22 -04:00
PiBind : TermLike
PiBind q d n = (Qty q, BindName, Term q d n)
2023-05-14 13:58:46 -04:00
private
2024-05-27 15:28:22 -04:00
pbname : PiBind q d n -> BindName
2023-05-14 13:58:46 -04:00
pbname (_, x, _) = x
2023-05-14 13:58:46 -04:00
private
2024-05-27 15:28:22 -04:00
record SplitPi q d n where
2023-05-14 13:58:46 -04:00
constructor MkSplitPi
2023-06-23 12:32:05 -04:00
{0 inner : Nat}
2024-05-27 15:28:22 -04:00
binds : Telescope (PiBind q d) n inner
cod : Term q d inner
2023-05-14 13:58:46 -04:00
private
2024-05-27 15:28:22 -04:00
splitPi : {q : Nat} ->
Telescope (PiBind q d) n n' -> Term q d n' -> SplitPi q d n
splitPi binds cod@(Pi qty arg res _) =
2023-05-14 13:58:46 -04:00
splitPi (binds :< (qty, res.name, arg)) $
2024-05-27 15:28:22 -04:00
assert_smaller cod $ pushSubsts' res.term
2023-05-14 13:58:46 -04:00
splitPi binds cod = MkSplitPi {binds, cod}
2023-05-14 13:58:46 -04:00
private
2024-05-27 15:28:22 -04:00
prettyPiBind1 : {opts : LayoutOpts} ->
NameContexts q d n -> Qty q -> BindName -> Term q d n ->
2023-05-14 13:58:46 -04:00
Eff Pretty (Doc opts)
2024-05-27 15:28:22 -04:00
prettyPiBind1 names pi (BN Unused _) s =
2023-05-14 13:58:46 -04:00
hcat <$> sequence
2024-05-27 15:28:22 -04:00
[prettyQty names pi, dotD,
withPrec Arg $ assert_total prettyTerm names s]
prettyPiBind1 names pi x s =
hcat <$> sequence
[prettyQty names pi, dotD,
hl Delim $ text "(",
hsep <$> sequence
[prettyTBind x, hl Delim $ text ":",
withPrec Outer $ assert_total prettyTerm names s],
hl Delim $ text ")"]
2023-05-14 13:58:46 -04:00
private
2024-05-27 15:28:22 -04:00
prettyPiBinds : {opts : LayoutOpts} ->
NameContexts q d n ->
Telescope (PiBind q d) n n' ->
2023-05-14 13:58:46 -04:00
Eff Pretty (SnocList (Doc opts))
2024-05-27 15:28:22 -04:00
prettyPiBinds _ [<] = pure [<]
prettyPiBinds names (binds :< (q, x, t)) =
let names' = extendTermN' (map pbname binds) names in
[|prettyPiBinds names binds :<
prettyPiBind1 names' q x t|]
2023-03-15 10:54:51 -04:00
2023-05-14 13:58:46 -04:00
private
2024-05-27 15:28:22 -04:00
SigBind : TermLike
SigBind q d n = (BindName, Term q d n)
2023-02-26 08:54:18 -05:00
2023-05-14 13:58:46 -04:00
private
2024-05-27 15:28:22 -04:00
record SplitSig q d n where
2023-05-14 13:58:46 -04:00
constructor MkSplitSig
2023-06-23 12:32:05 -04:00
{0 inner : Nat}
2024-05-27 15:28:22 -04:00
binds : Telescope (SigBind q d) n inner
last : Term q d inner
2023-02-22 01:40:19 -05:00
2023-05-14 13:58:46 -04:00
private
2024-05-27 15:28:22 -04:00
splitSig : {q : Nat} ->
Telescope (SigBind q d) n n' -> Term q d n' -> SplitSig q d n
2023-05-14 13:58:46 -04:00
splitSig binds (Sig fst snd _) =
splitSig (binds :< (snd.name, fst)) $
assert_smaller snd $ pushSubsts' snd.term
splitSig binds last = MkSplitSig {binds, last}
2023-02-26 08:54:18 -05:00
2023-05-14 13:58:46 -04:00
private
2024-05-27 15:28:22 -04:00
prettySigBind1 : {opts : LayoutOpts} ->
NameContexts q d n -> BindName -> Term q d n ->
2023-05-14 13:58:46 -04:00
Eff Pretty (Doc opts)
2024-05-27 15:28:22 -04:00
prettySigBind1 names (BN Unused _) s =
withPrec InTimes $ assert_total prettyTerm names s
prettySigBind1 names x s = hcat <$> sequence
2023-05-14 13:58:46 -04:00
[hl Delim $ text "(",
hsep <$> sequence
[prettyTBind x, hl Delim $ text ":",
2024-05-27 15:28:22 -04:00
withPrec Outer $ assert_total prettyTerm names s],
2023-05-14 13:58:46 -04:00
hl Delim $ text ")"]
2023-04-02 09:50:56 -04:00
2023-05-14 13:58:46 -04:00
private
2024-05-27 15:28:22 -04:00
prettySigBinds : {opts : LayoutOpts} ->
NameContexts q d n -> Telescope (SigBind q d) n n' ->
2023-05-14 13:58:46 -04:00
Eff Pretty (SnocList (Doc opts))
2024-05-27 15:28:22 -04:00
prettySigBinds _ [<] = pure [<]
prettySigBinds names (binds :< (x, t)) =
let names' = extendTermN' (map fst binds) names in
[|prettySigBinds names binds :<
prettySigBind1 names' x t|]
2023-04-02 09:50:56 -04:00
2023-05-14 13:58:46 -04:00
private
2024-05-27 15:28:22 -04:00
prettyTypeLine : {opts : LayoutOpts} ->
NameContexts q d n -> DScopeTerm q d n ->
2023-05-14 13:58:46 -04:00
Eff Pretty (Doc opts)
2024-05-27 15:28:22 -04:00
prettyTypeLine names (S _ (N body)) =
withPrec Arg (prettyTerm names body)
prettyTypeLine names (S [< i] (Y body)) =
parens =<< do
2024-05-27 15:28:22 -04:00
let names' = extendDim i names
2023-05-14 13:58:46 -04:00
i' <- prettyDBind i
2024-05-27 15:28:22 -04:00
ty' <- withPrec Outer $ prettyTerm names' body
2023-05-14 13:58:46 -04:00
pure $ sep [hsep [i', !darrowD], ty']
2023-04-02 09:50:56 -04:00
2023-05-14 13:58:46 -04:00
private
data NameSort = T | D
%runElab derive "NameSort" [Eq]
2023-02-22 01:40:19 -05:00
2023-05-14 13:58:46 -04:00
private
NameChunks : Type
NameChunks = SnocList (NameSort, SnocList BindName)
2023-02-22 01:40:19 -05:00
2023-05-14 13:58:46 -04:00
private
2024-05-27 15:28:22 -04:00
record SplitLams q d n where
2023-05-14 13:58:46 -04:00
constructor MkSplitLams
2023-06-23 12:32:05 -04:00
{0 dinner, ninner : Nat}
2023-05-14 13:58:46 -04:00
dnames : BTelescope d dinner
tnames : BTelescope n ninner
chunks : NameChunks
2024-05-27 15:28:22 -04:00
body : Term q dinner ninner
2023-02-22 01:40:19 -05:00
2023-05-14 13:58:46 -04:00
private
2024-05-27 15:28:22 -04:00
splitLams : {q : Nat} -> Term q d n -> SplitLams q d n
2023-05-14 13:58:46 -04:00
splitLams s = go [<] [<] [<] (pushSubsts' s)
where
push : NameSort -> BindName -> NameChunks -> NameChunks
push s y [<] = [< (s, [< y])]
push s y (xss :< (s', xs)) =
if s == s' then xss :< (s', xs :< y)
else xss :< (s', xs) :< (s, [< y])
2023-06-23 12:32:05 -04:00
go : BTelescope d d' -> BTelescope n n' ->
2023-05-14 13:58:46 -04:00
SnocList (NameSort, SnocList BindName) ->
2024-05-27 15:28:22 -04:00
Term q d' n' -> SplitLams q d n
2023-05-14 13:58:46 -04:00
go dnames tnames chunks (Lam b _) =
go dnames (tnames :< b.name) (push T b.name chunks) $
assert_smaller b $ pushSubsts' b.term
go dnames tnames chunks (DLam b _) =
go (dnames :< b.name) tnames (push D b.name chunks) $
assert_smaller b $ pushSubsts' b.term
go dnames tnames chunks s =
MkSplitLams dnames tnames chunks s
2023-04-03 11:46:23 -04:00
2023-02-22 01:40:19 -05:00
2023-05-14 13:58:46 -04:00
private
2024-05-27 15:28:22 -04:00
splitTuple : {q : Nat} ->
SnocList (Term q d n) -> Term q d n -> SnocList (Term q d n)
2023-05-14 13:58:46 -04:00
splitTuple ss p@(Pair t1 t2 _) =
splitTuple (ss :< t1) $ assert_smaller p $ pushSubsts' t2
splitTuple ss t = ss :< t
2023-03-16 13:39:24 -04:00
2023-05-14 13:58:46 -04:00
private
2024-05-27 15:28:22 -04:00
prettyTArg : {opts : LayoutOpts} ->
NameContexts q d n -> Term q d n -> Eff Pretty (Doc opts)
prettyTArg names s =
withPrec Arg $ assert_total prettyTerm names s
2023-05-14 13:58:46 -04:00
private
2024-05-27 15:28:22 -04:00
prettyDArg : {opts : LayoutOpts} ->
NameContexts _ d _ -> Dim d -> Eff Pretty (Doc opts)
prettyDArg names p = [|atD <+> withPrec Arg (prettyDim names.dnames p)|]
2023-05-14 13:58:46 -04:00
private
2024-05-27 15:28:22 -04:00
splitApps : {q : Nat} ->
Elim q d n -> (Elim q d n, List (Either (Dim d) (Term q d n)))
2023-05-14 13:58:46 -04:00
splitApps e = go [] (pushSubsts' e)
where
2024-05-27 15:28:22 -04:00
go : List (Either (Dim d) (Term q d n)) -> Elim q d n ->
(Elim q d n, List (Either (Dim d) (Term q d n)))
2023-05-14 13:58:46 -04:00
go xs e@(App f s _) =
go (Right s :: xs) $ assert_smaller e $ pushSubsts' f
go xs e@(DApp f p _) =
go (Left p :: xs) $ assert_smaller e $ pushSubsts' f
go xs s = (s, xs)
private
2024-05-27 15:28:22 -04:00
prettyDTApps : {opts : LayoutOpts} ->
NameContexts q d n ->
Elim q d n -> List (Either (Dim d) (Term q d n)) ->
2023-05-14 13:58:46 -04:00
Eff Pretty (Doc opts)
2024-05-27 15:28:22 -04:00
prettyDTApps names f xs = do
f <- withPrec Arg $ assert_total prettyElim names f
xs <- for xs $ either (prettyDArg names) (prettyTArg names)
prettyAppD f xs
2023-05-14 13:58:46 -04:00
private
2024-05-27 15:28:22 -04:00
record CaseArm opts q d n where
2023-05-14 13:58:46 -04:00
constructor MkCaseArm
pat : Doc opts
dbinds : BTelescope d dinner -- 🍴
tbinds : BTelescope n ninner
2024-05-27 15:28:22 -04:00
body : Term q dinner ninner
2023-05-14 13:58:46 -04:00
2024-05-27 15:28:22 -04:00
private
prettyCaseArm : {opts : LayoutOpts} ->
NameContexts q d n -> CaseArm opts q d n ->
Eff Pretty (Doc opts)
prettyCaseArm names (MkCaseArm pat dbinds tbinds body) = do
let names' = extendDimN' dbinds $ extendTermN' tbinds names
body <- withPrec Outer $ assert_total prettyTerm names' body
header <- (pat <++>) <$> darrowD
pure $ ifMultiline (header <++> body) (vsep [header, !(indentD body)])
2023-05-14 13:58:46 -04:00
2024-05-27 15:28:22 -04:00
private
prettyCaseBody : {opts : LayoutOpts} ->
NameContexts q d n -> List (CaseArm opts q d n) ->
Eff Pretty (List (Doc opts))
prettyCaseBody names xs = traverse (prettyCaseArm names) xs
2023-05-14 13:58:46 -04:00
private
2024-05-27 15:28:22 -04:00
prettyCompPat : {opts : LayoutOpts} ->
DimConst -> BindName -> Eff Pretty (Doc opts)
2023-05-14 13:58:46 -04:00
prettyCompPat e x = [|prettyDimConst e <++> prettyDBind x|]
2023-03-16 13:19:17 -04:00
private
2024-05-27 15:28:22 -04:00
prettyCompArm : {opts : LayoutOpts} -> NameContexts q d n ->
DimConst -> DScopeTerm q d n -> Eff Pretty (Doc opts)
prettyCompArm names e s = prettyCaseArm names $
MkCaseArm !(prettyCompPat e s.name) [< s.name] [<] s.term
private
2024-05-27 15:28:22 -04:00
layoutComp : {opts : LayoutOpts} ->
(typq : List (Doc opts)) -> (val, r : Doc opts) ->
(arms : List (Doc opts)) -> Eff Pretty (Doc opts)
layoutComp typq val r arms = do
comp <- compD; lb <- hl Delim "{"; rb <- hl Delim "}"
ind <- askAt INDENT
pure $ ifMultiline
(hsep $ concat {t = List} [[comp], typq, [val, r, lb], arms, [rb]]) $
(comp <++>
vsep [sep typq, val, r <++> lb, indent ind $ vsep arms, rb]) <|>
(vsep $ (comp ::) $ map (indent ind) $ concat {t = List}
[typq, [val, r <++> lb], map (indent ind) arms, [rb]])
2023-03-31 13:11:35 -04:00
2023-03-26 08:40:54 -04:00
export
2024-05-27 15:28:22 -04:00
prettyEnum : {opts : LayoutOpts} -> List String -> Eff Pretty (Doc opts)
2023-05-14 13:58:46 -04:00
prettyEnum cases =
tightBraces =<<
2023-05-14 13:58:46 -04:00
fillSeparateTight !commaD <$>
2023-11-05 08:30:40 -05:00
traverse (hl Constant . Doc.text . quoteTag) cases
2023-03-26 08:40:54 -04:00
private
2024-05-27 15:28:22 -04:00
prettyCaseRet : {opts : LayoutOpts} ->
NameContexts q d n -> ScopeTerm q d n -> Eff Pretty (Doc opts)
prettyCaseRet names body = withPrec Outer $ case body of
S _ (N tm) => assert_total prettyTerm names tm
2023-05-14 13:58:46 -04:00
S [< x] (Y tm) => do
2024-05-27 15:28:22 -04:00
let names' = extendTerm x names
2023-05-14 13:58:46 -04:00
header <- [|prettyTBind x <++> darrowD|]
2024-05-27 15:28:22 -04:00
body <- assert_total prettyTerm names' tm
2023-11-03 12:47:01 -04:00
hangDSingle header body
2023-03-26 08:40:54 -04:00
2023-05-14 13:58:46 -04:00
private
2024-05-27 15:28:22 -04:00
prettyCase_ : {opts : LayoutOpts} ->
NameContexts q d n ->
Doc opts -> Elim q d n -> ScopeTerm q d n ->
List (CaseArm opts q d n) ->
2023-05-14 13:58:46 -04:00
Eff Pretty (Doc opts)
2024-05-27 15:28:22 -04:00
prettyCase_ names intro head ret body = do
head <- withPrec Outer $ assert_total prettyElim names head
ret <- prettyCaseRet names ret
bodys <- prettyCaseBody names body
2023-11-03 12:47:01 -04:00
return <- returnD; of_ <- ofD
lb <- hl Delim "{"; rb <- hl Delim "}"; semi <- semiD
ind <- askAt INDENT
parensIfM Outer $ ifMultiline
(hsep [intro, head, return, ret, of_, lb, hseparateTight semi bodys, rb])
(vsep [intro <++> head, return <++> ret, of_ <++> lb,
indent ind $ vseparateTight semi bodys, rb])
2023-05-14 13:58:46 -04:00
private
2024-05-27 15:28:22 -04:00
prettyCase : {opts : LayoutOpts} ->
NameContexts q d n ->
Qty q -> Elim q d n -> ScopeTerm q d n ->
List (CaseArm opts q d n) ->
2023-05-14 13:58:46 -04:00
Eff Pretty (Doc opts)
2024-05-27 15:28:22 -04:00
prettyCase names qty head ret body =
prettyCase_ names ![|caseD <+> prettyQty names qty|] head ret body
2023-05-14 13:58:46 -04:00
2023-12-04 16:47:52 -05:00
private
2024-05-27 15:28:22 -04:00
LetBinder : TermLike
LetBinder q d n = (Qty q, BindName, Elim q d n)
2023-12-04 16:47:52 -05:00
private
2024-05-27 15:28:22 -04:00
LetExpr : (q, d, n, n' : Nat) -> Type
LetExpr q d n n' = (Telescope (LetBinder q d) n n', Term q d n')
2023-12-04 16:47:52 -05:00
-- [todo] factor out this and the untyped version somehow
export
2024-05-27 15:28:22 -04:00
splitLet : Telescope (LetBinder q d) n n' -> Term q d n' ->
Exists (LetExpr q d n)
2023-12-04 16:47:52 -05:00
splitLet ys t@(Let qty rhs body _) =
splitLet (ys :< (qty, body.name, rhs)) (assert_smaller t body.term)
splitLet ys t =
Evidence _ (ys, t)
private covering
prettyLets : {opts : LayoutOpts} ->
2024-05-27 15:28:22 -04:00
NameContexts q d a -> Telescope (LetBinder q d) a b ->
2023-12-04 16:47:52 -05:00
Eff Pretty (SnocList (Doc opts))
2024-05-27 15:28:22 -04:00
prettyLets (MkNameContexts qnames dnames tnames) lets = snd <$> go lets where
peelAnn : forall d, n. Elim q d n -> Maybe (Term q d n, Term q d n)
2023-12-21 12:03:57 -05:00
peelAnn (Ann tm ty _) = Just (tm, ty)
peelAnn e = Nothing
2024-05-27 15:28:22 -04:00
letHeader : Qty q -> BindName -> Eff Pretty (Doc opts)
2023-12-21 12:03:57 -05:00
letHeader qty x = do
2024-05-27 15:28:22 -04:00
lett <- [|letD <+> prettyQty qnames qty|]
2023-12-21 12:03:57 -05:00
x <- prettyTBind x
pure $ lett <++> x
letBody : forall n. BContext n ->
2024-05-27 15:28:22 -04:00
Doc opts -> Elim q d n -> Eff Pretty (Doc opts)
letBody tnames hdr e =
let names = MkNameContexts' qnames dnames tnames in
case peelAnn e of
Just (tm, ty) => do
ty <- withPrec Outer $ assert_total prettyTerm names ty
tm <- withPrec Outer $ assert_total prettyTerm names tm
colon <- colonD; eq <- cstD; d <- askAt INDENT
pure $ hangSingle d (hangSingle d hdr (colon <++> ty)) (eq <++> tm)
Nothing => do
e <- withPrec Outer $ assert_total prettyElim names e
eq <- cstD; d <- askAt INDENT
inn <- inD
pure $ ifMultiline
(hsep [hdr, eq, e, inn])
(vsep [hdr, indent d $ hsep [eq, e, inn]])
go : forall b. Telescope (LetBinder q d) a b ->
2023-12-21 12:03:57 -05:00
Eff Pretty (BContext b, SnocList (Doc opts))
2024-05-27 15:28:22 -04:00
go [<] = pure (tnames, [<])
2023-12-21 12:03:57 -05:00
go (lets :< (qty, x, rhs)) = do
(ys, docs) <- go lets
doc <- letBody ys !(letHeader qty x) rhs
pure (ys :< x, docs :< doc)
2023-12-04 16:47:52 -05:00
private
isDefaultDir : Dim d -> Dim d -> Bool
isDefaultDir (K Zero _) (K One _) = True
isDefaultDir _ _ = False
2023-05-14 13:58:46 -04:00
-- [fixme] use telescopes in Scoped
private
toTel : BContext s -> BTelescope n (s + n)
toTel [<] = [<]
toTel (ctx :< x) = toTel ctx :< x
private
2024-05-27 15:28:22 -04:00
prettyTyCasePat : {opts : LayoutOpts} ->
2023-05-14 13:58:46 -04:00
(k : TyConKind) -> BContext (arity k) ->
Eff Pretty (Doc opts)
prettyTyCasePat KTYPE [<] = typeD
prettyTyCasePat KIOState [<] = ioStateD
2023-05-14 13:58:46 -04:00
prettyTyCasePat KPi [< a, b] =
parens . hsep =<< sequence [prettyTBind a, arrowD, prettyTBind b]
prettyTyCasePat KSig [< a, b] =
parens . hsep =<< sequence [prettyTBind a, timesD, prettyTBind b]
prettyTyCasePat KEnum [<] = hl Syntax $ text "{}"
prettyTyCasePat KEq [< a0, a1, a, l, r] =
hsep <$> sequence (eqD :: map prettyTBind [a0, a1, a, l, r])
prettyTyCasePat KNat [<] = natD
prettyTyCasePat KString [<] = stringD
2023-05-14 13:58:46 -04:00
prettyTyCasePat KBOX [< a] = bracks =<< prettyTBind a
2024-05-27 15:28:22 -04:00
prettyLambda : {opts : LayoutOpts} ->
NameContexts q d n -> Term q d n -> Eff Pretty (Doc opts)
prettyLambda names s =
let Val q = names.qtyLen
MkSplitLams {dnames = ds, tnames = ts, chunks, body} = splitLams s
names' = extendDimN' ds $ extendTermN' ts names in
parensIfM Outer =<<
hangDSingle !(header chunks) !(assert_total prettyTerm names' body)
2023-05-14 13:58:46 -04:00
where
introChar : NameSort -> Eff Pretty (Doc opts)
introChar T = lamD
introChar D = dlamD
prettyBind : NameSort -> BindName -> Eff Pretty (Doc opts)
prettyBind T = prettyTBind
prettyBind D = prettyDBind
header1 : NameSort -> List BindName -> Eff Pretty (Doc opts)
header1 s xs = hsep <$> sequence
[introChar s, sep <$> traverse (prettyBind s) xs, darrowD]
header : NameChunks -> Eff Pretty (Doc opts)
header cs = sep <$> traverse (\(s, xs) => header1 s (toList xs)) (toList cs)
2024-05-27 15:28:22 -04:00
prettyDisp : {opts : LayoutOpts} -> Universe -> Eff Pretty (Maybe (Doc opts))
2023-05-21 14:09:34 -04:00
prettyDisp 0 = pure Nothing
prettyDisp u = map Just $ hl Universe =<<
ifUnicode (text $ superscript $ show u) (text $ "^" ++ show u)
2024-05-27 15:28:22 -04:00
prettyTerm names (TYPE l _) = do
2024-04-12 15:52:51 -04:00
type <- hl Syntax . text =<< ifUnicode "" "Type"
level <- prettyDisp l
pure $ maybe type (type <+>) level
2023-05-14 13:58:46 -04:00
2024-05-27 15:28:22 -04:00
prettyTerm _ (IOState _) =
ioStateD
2024-05-27 15:28:22 -04:00
prettyTerm names (Pi qty arg res _) = do
let Val q = names.qtyLen
MkSplitPi {binds, cod} = splitPi [< (qty, res.name, arg)] res.term
arr <- arrowD
lines <- map (<++> arr) <$> prettyPiBinds names binds
let names' = extendTermN' (map pbname binds) names
cod <- withPrec Outer $ prettyTerm names' (assert_smaller res cod)
parensIfM Outer $ sepSingle $ toList $ lines :< cod
prettyTerm names s@(Lam {}) =
prettyLambda names s
prettyTerm names (Sig tfst tsnd _) = do
let Val q = names.qtyLen
MkSplitSig {binds, last} = splitSig [< (tsnd.name, tfst)] tsnd.term
times <- timesD
lines <- map (<++> times) <$> prettySigBinds names binds
let names' = extendTermN' (map fst binds) names
last <- withPrec InTimes $ prettyTerm names' (assert_smaller tsnd last)
parensIfM Times $ sepSingle $ toList $ lines :< last
prettyTerm names p@(Pair s t _) = do
let Val q = names.qtyLen
elems = splitTuple [< s] t
lines <- for elems $ \t =>
withPrec Outer $ prettyTerm names $ assert_smaller p t
parens $ separateTight !commaD lines
prettyTerm _ (Enum cases _) =
2023-05-14 13:58:46 -04:00
prettyEnum $ SortedSet.toList cases
2024-05-27 15:28:22 -04:00
prettyTerm _ (Tag tag _) =
2023-05-14 13:58:46 -04:00
prettyTag tag
2024-05-27 15:28:22 -04:00
prettyTerm names (Eq (S _ (N ty)) l r _) = do
l <- withPrec InEq $ prettyTerm names l
r <- withPrec InEq $ prettyTerm names r
ty <- withPrec InEq $ prettyTerm names ty
parensIfM Eq $ sep [l <++> !eqndD, r <++> !colonD, ty]
prettyTerm names (Eq ty l r _) = do
ty <- prettyTypeLine names ty
l <- withPrec Arg $ prettyTerm names l
r <- withPrec Arg $ prettyTerm names r
prettyAppD !eqD [ty, l, r]
prettyTerm names s@(DLam {}) =
prettyLambda names s
prettyTerm _ (NAT _) = natD
prettyTerm _ (Nat n _) = hl Syntax $ pshow n
prettyTerm names (Succ p _) =
prettyAppD !succD [!(withPrec Arg $ prettyTerm names p)]
prettyTerm _ (STRING _) = stringD
prettyTerm _ (Str s _) = prettyStrLit s
prettyTerm names (BOX qty ty _) =
2023-05-14 13:58:46 -04:00
bracks . hcat =<<
2024-05-27 15:28:22 -04:00
sequence [prettyQty names qty, dotD,
withPrec Outer $ prettyTerm names ty]
2023-05-14 13:58:46 -04:00
2024-05-27 15:28:22 -04:00
prettyTerm names (Box val _) =
bracks =<< withPrec Outer (prettyTerm names val)
2023-05-14 13:58:46 -04:00
2024-05-27 15:28:22 -04:00
prettyTerm names (Let qty rhs body _) = do
2023-12-04 16:47:52 -05:00
let Evidence _ (lets, body) = splitLet [< (qty, body.name, rhs)] body.term
2024-05-27 15:28:22 -04:00
heads <- prettyLets names lets
let names' = extendTermN' (map (fst . snd) lets) names
body <- withPrec Outer $ assert_total prettyTerm names' body
2023-12-04 16:47:52 -05:00
let lines = toList $ heads :< body
pure $ ifMultiline (hsep lines) (vsep lines)
2024-05-27 15:28:22 -04:00
prettyTerm names (E e) = do
-- [fixme] do this stuff somewhere else!!!
let Val q = names.qtyLen
case pushSubsts' e {tm = Elim} of
Ann tm _ _ => assert_total prettyTerm names tm
_ => assert_total prettyElim names e
2023-05-14 13:58:46 -04:00
2024-05-27 15:28:22 -04:00
prettyTerm names t0@(CloT (Sub t ph)) = do
let Val q = names.qtyLen
prettyTerm names $ assert_smaller t0 $ pushSubstsWith' id id ph t
2023-05-14 13:58:46 -04:00
2024-05-27 15:28:22 -04:00
prettyTerm names t0@(DCloT (Sub t ph)) = do
let Val q = names.qtyLen
prettyTerm names $ assert_smaller t0 $ pushSubstsWith' id ph id t
2023-05-14 13:58:46 -04:00
2024-05-27 15:28:22 -04:00
prettyTerm names t0@(QCloT (SubR t ph)) = do
let Val q = names.qtyLen
prettyTerm names $ assert_smaller t0 $ pushSubstsWith' ph id id t
prettyElim names (F x u _) = do
2023-05-21 14:09:34 -04:00
x <- prettyFree x
u <- prettyDisp u
pure $ maybe x (x <+>) u
2023-05-14 13:58:46 -04:00
2024-05-27 15:28:22 -04:00
prettyElim names (B i _) =
prettyTBind $ names.tnames !!! i
2023-05-14 13:58:46 -04:00
2024-05-27 15:28:22 -04:00
prettyElim names e@(App {}) = do
let Val q = names.qtyLen
(f, xs) = splitApps e
prettyDTApps names f xs
2023-05-14 13:58:46 -04:00
2024-05-27 15:28:22 -04:00
prettyElim names (CasePair qty pair ret body _) = do
2023-05-14 13:58:46 -04:00
let [< x, y] = body.names
2024-05-27 15:28:22 -04:00
pat <- parens $ !(prettyTBind x) <+> !commaD <++> !(prettyTBind y)
prettyCase names qty pair ret
2023-05-14 13:58:46 -04:00
[MkCaseArm pat [<] [< x, y] body.term]
2024-05-27 15:28:22 -04:00
prettyElim names (Fst pair _) = do
pair <- prettyTArg names (E pair)
prettyAppD !fstD [pair]
2023-09-18 15:52:51 -04:00
2024-05-27 15:28:22 -04:00
prettyElim names (Snd pair _) = do
pair <- prettyTArg names (E pair)
prettyAppD !sndD [pair]
2023-09-18 15:52:51 -04:00
2024-05-27 15:28:22 -04:00
prettyElim names (CaseEnum qty tag ret arms _) = do
2023-05-14 13:58:46 -04:00
arms <- for (SortedMap.toList arms) $ \(tag, body) =>
pure $ MkCaseArm !(prettyTag tag) [<] [<] body
2024-05-27 15:28:22 -04:00
prettyCase names qty tag ret arms
2023-05-14 13:58:46 -04:00
2024-05-27 15:28:22 -04:00
prettyElim names (CaseNat qty qtyIH nat ret zero succ _) = do
2023-05-14 13:58:46 -04:00
let zarm = MkCaseArm !zeroD [<] [<] zero
[< p, ih] = succ.names
spat0 <- [|succD <++> prettyTBind p|]
2024-05-27 15:28:22 -04:00
ihpat0 <- map hcat $ sequence [prettyQty names qtyIH, dotD, prettyTBind ih]
2023-09-20 15:58:27 -04:00
spat <- if ih.val == Unused
2023-05-14 13:58:46 -04:00
then pure spat0
2024-05-27 15:28:22 -04:00
else pure $ spat0 <+> !commaD <++> ihpat0
2023-05-14 13:58:46 -04:00
let sarm = MkCaseArm spat [<] [< p, ih] succ.term
2024-05-27 15:28:22 -04:00
prettyCase names qty nat ret [zarm, sarm]
2023-05-14 13:58:46 -04:00
2024-05-27 15:28:22 -04:00
prettyElim names (CaseBox qty box ret body _) = do
2023-05-14 13:58:46 -04:00
pat <- bracks =<< prettyTBind body.name
let arm = MkCaseArm pat [<] [< body.name] body.term
2024-05-27 15:28:22 -04:00
prettyCase names qty box ret [arm]
prettyElim names e@(DApp {}) = do
let Val q = names.qtyLen
(f, xs) = splitApps e
prettyDTApps names f xs
prettyElim names (Ann tm ty _) = do
-- [fixme] do this stuff somewhere else!!!
let Val q = names.qtyLen
case pushSubsts' tm {tm = Term} of
E e => assert_total prettyElim names e
2023-12-21 12:03:57 -05:00
_ => do
2024-05-27 15:28:22 -04:00
tm <- withPrec AnnL $ assert_total prettyTerm names tm
ty <- withPrec Outer $ assert_total prettyTerm names ty
2023-12-21 12:03:57 -05:00
parensIfM Outer =<< hangDSingle (tm <++> !annD) ty
2023-05-14 13:58:46 -04:00
2024-05-27 15:28:22 -04:00
prettyElim names (Coe ty p p' val _) = do
ty <- prettyTypeLine names ty
p <- prettyDArg names p
p' <- prettyDArg names p'
val <- prettyTArg names val
prettyAppD !coeD [ty, p <++> p', val]
prettyElim names e@(Comp ty p p' val r zero one _) = do
ty <- assert_total $ prettyTypeLine names $ SN ty
pp <- [|prettyDArg names p <++> prettyDArg names p'|]
val <- prettyTArg names val
r <- prettyDArg names r
arm0 <- [|prettyCompArm names Zero zero <+> semiD|]
arm1 <- prettyCompArm names One one
ind <- askAt INDENT
parensIfM App =<< layoutComp [ty, pp] val r [arm0, arm1]
prettyElim names (TypeCase ty ret arms def _) = do
2023-05-14 13:58:46 -04:00
arms <- for (toList arms) $ \(k ** body) => do
pat <- prettyTyCasePat k body.names
pure $ MkCaseArm pat [<] (toTel body.names) body.term
let darm = MkCaseArm !undD [<] [<] def
2024-05-27 15:28:22 -04:00
prettyCase_ names !typecaseD ty (SN ret) $ arms ++ [darm]
prettyElim names e0@(CloE (Sub e ph)) = do
let Val q = names.qtyLen
prettyElim names $ assert_smaller e0 $ pushSubstsWith' id id ph e
2023-05-14 13:58:46 -04:00
2024-05-27 15:28:22 -04:00
prettyElim names e0@(DCloE (Sub e ph)) = do
let Val q = names.qtyLen
prettyElim names $ assert_smaller e0 $ pushSubstsWith' id ph id e
2023-05-14 13:58:46 -04:00
2024-05-27 15:28:22 -04:00
prettyElim names e0@(QCloE (SubR e ph)) = do
let Val q = names.qtyLen
prettyElim names $ assert_smaller e0 $ pushSubstsWith' ph id id e